From 688cedb54f977f8944b3674c194ac930d8d283b3 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Sat, 16 Jul 2022 12:49:50 -0400 Subject: [PATCH 001/173] linear_combination with exponent --- src/tactic/linear_combination.lean | 32 +++++++++++++++++++++++++++--- test/linear_combination.lean | 25 +++++++++++++++++++++++ 2 files changed, 54 insertions(+), 3 deletions(-) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index 24b435ba60b8f..af7bbf8348889 100644 --- a/src/tactic/linear_combination.lean +++ b/src/tactic/linear_combination.lean @@ -266,6 +266,19 @@ This tactic only should be used when the target's type an equality whose meta def set_goal_to_hleft_eq_tleft (hsum_on_left : expr) : tactic unit := do to_expr ``(replace_eq_expr %%hsum_on_left) >>= apply, skip +/- +If an exponent `n` is provided, changes the goal from `t = 0` to `t^n = 0`. +* Input: + * `exponent : option ℕ`, the power to raise the goal by. If `none`, this tactic is a no-op. + +* Output: N/A +-/ +meta def raise_goal_to_power (exponent : option ℕ) : tactic unit := +match exponent with +| some n := refine ``(@pow_eq_zero _ _ _ _ %%`(n) _) +| none := skip +end + /-- This tactic attempts to prove the goal by normalizing the target if the `normalize` field of the given configuration is true. @@ -307,13 +320,14 @@ Note: The left and right sides of all the equalities should have the same * Output: N/A -/ meta def linear_combination (h_eqs_names : list pexpr) (coeffs : list pexpr) - (config : linear_combination_config := {}) : tactic unit := + (exponent : option ℕ := none) (config : linear_combination_config := {}) : tactic unit := do `(@eq %%ext _ _) ← target | fail "linear_combination can only be used to prove equality goals", h_eqs ← h_eqs_names.mmap to_expr, 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 exponent, set_goal_to_hleft_eq_tleft hsum_on_left, prove_equal_if_desired config @@ -354,6 +368,9 @@ setup_tactic_parser prove their target using the linear combination instead of attempting to finish the proof. +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. @@ -399,6 +416,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 := @@ -408,10 +428,16 @@ by linear_combination 3 * h a b + hqc meta def _root_.tactic.interactive.linear_combination (input : parse (as_linear_combo ff [] <$> texpr)?) (_ : parse (tk "with")?) + (exponent : parse (prod.mk <$> ident <*> small_nat)?) (config : linear_combination_config := {}) - : tactic unit := + : tactic unit := do + exponent ← match exponent with + | none := return none + | some (`exponent, n) := return $ some n + | some (id, _) := fail!"linear_combination does not support the modifier {id}" + end, let (h_eqs_names, coeffs) := list.unzip (input.get_or_else []) in -linear_combination h_eqs_names coeffs config +linear_combination h_eqs_names coeffs exponent config add_tactic_doc { name := "linear_combination", diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 3fa4d8420f0bb..d6509161f711d 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 that should fail -/ -- This should fail because there are no hypotheses given From 23352d9fd2d378dce86692eb119c5d994d1a92e8 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Sat, 16 Jul 2022 17:08:37 -0400 Subject: [PATCH 002/173] fix doc string --- src/tactic/linear_combination.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index af7bbf8348889..00d1f06e5e335 100644 --- a/src/tactic/linear_combination.lean +++ b/src/tactic/linear_combination.lean @@ -266,7 +266,7 @@ This tactic only should be used when the target's type an equality whose meta def set_goal_to_hleft_eq_tleft (hsum_on_left : expr) : tactic unit := do to_expr ``(replace_eq_expr %%hsum_on_left) >>= apply, skip -/- +/-- If an exponent `n` is provided, changes the goal from `t = 0` to `t^n = 0`. * Input: * `exponent : option ℕ`, the power to raise the goal by. If `none`, this tactic is a no-op. From 0cb3798a5b33919cf69962cb209f37e14cc7db1c Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Fri, 22 Jul 2022 14:42:11 -0400 Subject: [PATCH 003/173] change handling of power 1 --- src/tactic/linear_combination.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index b9217f42a843b..41435923f9f52 100644 --- a/src/tactic/linear_combination.lean +++ b/src/tactic/linear_combination.lean @@ -269,14 +269,14 @@ 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 : option ℕ`, the power to raise the goal by. If `none`, this tactic is a no-op. + * `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 (exponent : option ℕ) : tactic unit := +meta def raise_goal_to_power (exponent : ℕ) : tactic unit := match exponent with -| some n := refine ``(@pow_eq_zero _ _ _ _ %%`(n) _) -| none := skip +| 1 := skip +| n := refine ``(@pow_eq_zero _ _ _ _ %%`(n) _) end /-- @@ -327,7 +327,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 exponent, + raise_goal_to_power (exponent.get_or_else 1), set_goal_to_hleft_sub_tleft hsum_on_left, normalize_if_desired config From b7211810961e4632e564d1bca37c7547d58161ef Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 22 Jul 2022 15:53:07 +0000 Subject: [PATCH 004/173] feat(order/boolean_algebra): A bounded generalized boolean algebra is a boolean algebra (#15606) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Abstract the construction of `boolean_algebra (finset α)`. --- src/data/fintype/basic.lean | 9 +-------- src/order/boolean_algebra.lean | 10 ++++++++++ 2 files changed, 11 insertions(+), 8 deletions(-) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 7d5cc04e96cb7..8c92ba37cabd3 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -125,14 +125,7 @@ instance : order_top (finset α) := section boolean_algebra variables [decidable_eq α] {a : α} -instance : boolean_algebra (finset α) := -{ compl := λ s, univ \ s, - inf_compl_le_bot := λ s x hx, by simpa using hx, - top_le_sup_compl := λ s x hx, by simp, - sdiff_eq := λ s t, by simp [generalized_boolean_algebra.sdiff, ext_iff, compl], - ..finset.order_top, - ..finset.order_bot, - ..finset.generalized_boolean_algebra } +instance : boolean_algebra (finset α) := generalized_boolean_algebra.to_boolean_algebra lemma compl_eq_univ_sdiff (s : finset α) : sᶜ = univ \ s := rfl diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 1576e7a036024..cb52595dfcfa7 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -618,6 +618,16 @@ class boolean_algebra (α : Type u) extends distrib_lattice α, has_compl α, ha instance boolean_algebra.to_bounded_order [h : boolean_algebra α] : bounded_order α := { ..h } +/-- A bounded generalized boolean algebra is a boolean algebra. -/ +@[reducible] -- See note [reducible non instances] +def generalized_boolean_algebra.to_boolean_algebra [generalized_boolean_algebra α] [order_top α] : + boolean_algebra α := +{ compl := λ a, ⊤ \ a, + inf_compl_le_bot := λ _, disjoint_sdiff_self_right, + top_le_sup_compl := λ _, le_sup_sdiff, + sdiff_eq := λ _ _, by { rw [←inf_sdiff_assoc, inf_top_eq], refl }, + ..‹generalized_boolean_algebra α›, ..generalized_boolean_algebra.to_order_bot, ..‹order_top α› } + section boolean_algebra variables [boolean_algebra α] From 26c4a7c9dbfe5b021681f40f4ea71c62fd9b8b13 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 22 Jul 2022 15:53:08 +0000 Subject: [PATCH 005/173] =?UTF-8?q?fix(data/json):=20`rbmap=20string=20?= =?UTF-8?q?=CE=B1`=20never=20serializes=20to=20`null`=20(#15622)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This change means that `option (rbmap string α)` can now serialize --- src/data/json.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/json.lean b/src/data/json.lean index 8521c6e8bb7c0..5f13ebb66e7f0 100644 --- a/src/data/json.lean +++ b/src/data/json.lean @@ -98,7 +98,7 @@ meta instance {α} [json_serializable α] : non_null_json_serializable (list α) json.array l ← success j | exception (λ _, format!"array expected, got {j.typename}"), l.mmap (of_json α) } -meta instance {α} [json_serializable α] : json_serializable (rbmap string α) := +meta instance {α} [json_serializable α] : non_null_json_serializable (rbmap string α) := { to_json := λ m, json.object (m.to_list.map $ λ x, (x.1, to_json x.2)), of_json := λ j, do json.object l ← success j | exception (λ _, format!"object expected, got {j.typename}"), From 8618fdb78520830ecb1a09a22331d7e3f3b73e11 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 22 Jul 2022 18:27:52 +0000 Subject: [PATCH 006/173] feat(topology/metric_space/isometry): use namespace, add lemmas (#15591) * Use `namespace isometry`. * Add lemmas like `isometry.preimage_ball`. --- src/analysis/normed/group/add_torsor.lean | 6 +- src/analysis/normed/group/basic.lean | 8 +- src/analysis/normed_space/basic.lean | 2 +- .../function/conditional_expectation.lean | 8 +- src/topology/continuous_function/bounded.lean | 2 +- src/topology/metric_space/completion.lean | 2 +- src/topology/metric_space/gluing.lean | 12 +- .../metric_space/gromov_hausdorff.lean | 4 +- .../gromov_hausdorff_realized.lean | 4 +- src/topology/metric_space/isometry.lean | 165 ++++++++++-------- src/topology/metric_space/kuratowski.lean | 2 +- 11 files changed, 118 insertions(+), 97 deletions(-) diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index 59e792dca1523..ca0d5166bca5e 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -74,7 +74,7 @@ by rw [dist_comm, dist_vadd_left] addition/subtraction of `x : P`. -/ @[simps] def isometric.vadd_const (x : P) : V ≃ᵢ P := { to_equiv := equiv.vadd_const x, - isometry_to_fun := isometry_emetric_iff_metric.2 $ λ _ _, dist_vadd_cancel_right _ _ _ } + isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_right _ _ _ } section @@ -83,7 +83,7 @@ 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 := { to_equiv := equiv.const_vadd P x, - isometry_to_fun := isometry_emetric_iff_metric.2 $ λ _ _, dist_vadd_cancel_left _ _ _ } + isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_left _ _ _ } end @@ -94,7 +94,7 @@ by rw [dist_eq_norm, vsub_sub_vsub_cancel_left, dist_comm, dist_eq_norm_vsub V] subtraction from `x : P`. -/ @[simps] def isometric.const_vsub (x : P) : P ≃ᵢ V := { to_equiv := equiv.const_vsub x, - isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_vsub_cancel_left _ _ _ } + 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 diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 4da7b2422f751..b597da76db301 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -378,7 +378,7 @@ namespace isometric /-- Addition `y ↦ y + x` as an `isometry`. -/ protected def add_right (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_add_right _ _ _, +{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_right _ _ _, .. equiv.add_right x } @[simp] lemma add_right_to_equiv (x : E) : @@ -394,7 +394,7 @@ ext $ λ y, rfl /-- Addition `y ↦ x + y` as an `isometry`. -/ protected def add_left (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_add_left _ _ _, +{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_left _ _ _, to_equiv := equiv.add_left x } @[simp] lemma add_left_to_equiv (x : E) : @@ -410,7 +410,7 @@ variable (E) /-- Negation `x ↦ -x` as an `isometry`. -/ protected def neg : E ≃ᵢ E := -{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ x y, dist_neg_neg _ _, +{ isometry_to_fun := isometry.of_dist_eq $ λ x y, dist_neg_neg _ _, to_equiv := equiv.neg E } variable {E} @@ -511,7 +511,7 @@ end lemma add_monoid_hom_class.isometry_iff_norm {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] (f : 𝓕) : isometry f ↔ ∀ x, ∥f x∥ = ∥x∥ := begin - simp only [isometry_emetric_iff_metric, dist_eq_norm, ←map_sub], + simp only [isometry_iff_dist_eq, dist_eq_norm, ←map_sub], refine ⟨λ h x, _, λ h x y, h _⟩, simpa using h x 0 end diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index b4e1bd06932af..391ff8cbc9974 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -411,7 +411,7 @@ variables (𝕜 𝕜') /-- In a normed algebra, the inclusion of the base field in the extended field is an isometry. -/ lemma algebra_map_isometry [norm_one_class 𝕜'] : isometry (algebra_map 𝕜 𝕜') := begin - refine isometry_emetric_iff_metric.2 (λx y, _), + refine isometry.of_dist_eq (λx y, _), rw [dist_eq_norm, dist_eq_norm, ← ring_hom.map_sub, norm_algebra_map'], end diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index f18b753473b65..a31e6a938552f 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -475,12 +475,8 @@ 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) := -begin - rw isometry_emetric_iff_metric, - intros f g, - rw [dist_eq_norm, ← Lp_meas_subgroup_to_Lp_trim_sub, Lp_meas_subgroup_to_Lp_trim_norm_map, - dist_eq_norm], -end +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. -/ diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 49371a2f33292..282e1a6ce4625 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -392,7 +392,7 @@ end lemma isometry_extend (f : α ↪ δ) (h : δ →ᵇ β) : isometry (λ g : α →ᵇ β, extend f g h) := -isometry_emetric_iff_metric.2 $ λ g₁ g₂, by simp [dist_nonneg] +isometry.of_dist_eq $ λ g₁ g₂, by simp [dist_nonneg] end extend diff --git a/src/topology/metric_space/completion.lean b/src/topology/metric_space/completion.lean index 5f3269b118b4c..29c31422613e8 100644 --- a/src/topology/metric_space/completion.lean +++ b/src/topology/metric_space/completion.lean @@ -173,7 +173,7 @@ instance : metric_space (completion α) := /-- The embedding of a metric space in its completion is an isometry. -/ lemma coe_isometry : isometry (coe : α → completion α) := -isometry_emetric_iff_metric.2 completion.dist_eq +isometry.of_dist_eq completion.dist_eq @[simp] protected lemma edist_eq (x y : α) : edist (x : completion α) y = edist x y := coe_isometry x y diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index e369594eef2fb..a1d45f7937fe2 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -338,11 +338,11 @@ lemma sum.dist_eq {x y : X ⊕ Y} : dist x y = sum.dist x y := rfl /-- The left injection of a space in a disjoint union is an isometry -/ lemma isometry_inl : isometry (sum.inl : X → (X ⊕ Y)) := -isometry_emetric_iff_metric.2 $ λx y, rfl +isometry.of_dist_eq $ λ x y, rfl /-- The right injection of a space in a disjoint union is an isometry -/ lemma isometry_inr : isometry (sum.inr : Y → (X ⊕ Y)) := -isometry_emetric_iff_metric.2 $ λx y, rfl +isometry.of_dist_eq $ λ x y, rfl end sum @@ -487,7 +487,7 @@ open filter /-- The injection of a space in a disjoint union is an isometry -/ lemma isometry_mk (i : ι) : isometry (sigma.mk i : E i → Σ k, E k) := -isometry_emetric_iff_metric.2 (by simp) +isometry.of_dist_eq (λ x y, by simp) /-- A disjoint union of complete metric spaces is complete. -/ protected lemma complete_space [∀ i, complete_space (E i)] : complete_space (Σ i, E i) := @@ -559,10 +559,10 @@ begin end lemma to_glue_l_isometry (hΦ : isometry Φ) (hΨ : isometry Ψ) : isometry (to_glue_l hΦ hΨ) := -isometry_emetric_iff_metric.2 $ λ_ _, rfl +isometry.of_dist_eq $ λ_ _, rfl lemma to_glue_r_isometry (hΦ : isometry Φ) (hΨ : isometry Ψ) : isometry (to_glue_r hΦ hΨ) := -isometry_emetric_iff_metric.2 $ λ_ _, rfl +isometry.of_dist_eq $ λ_ _, rfl end gluing --section @@ -656,7 +656,7 @@ instance (I : ∀ n, isometry (f n)) [inhabited (X 0)] : inhabited (inductive_li /-- The map `to_inductive_limit n` mapping `X n` to the inductive limit is an isometry. -/ lemma to_inductive_limit_isometry (I : ∀ n, isometry (f n)) (n : ℕ) : - isometry (to_inductive_limit I n) := isometry_emetric_iff_metric.2 $ λx y, + isometry (to_inductive_limit I n) := isometry.of_dist_eq $ λ x y, begin change inductive_limit_dist f ⟨n, x⟩ ⟨n, y⟩ = dist x y, rw [inductive_limit_dist_eq_dist I ⟨n, x⟩ ⟨n, y⟩ n (le_refl n) (le_refl n), diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 9033f3d784633..fd9c2b517a402 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -534,8 +534,8 @@ begin glue_metric_approx (λ x:s, (x:X)) (λ x, Φ x) (ε₂/2 + δ) (by linarith) this, let Fl := @sum.inl X Y, let Fr := @sum.inr X Y, - have Il : isometry Fl := isometry_emetric_iff_metric.2 (λ x y, rfl), - have Ir : isometry Fr := isometry_emetric_iff_metric.2 (λ x y, rfl), + have Il : isometry Fl := isometry.of_dist_eq (λ x y, rfl), + have Ir : isometry Fr := isometry.of_dist_eq (λ x y, rfl), /- The proof goes as follows : the `GH_dist` is bounded by the Hausdorff distance of the images in the coupling, which is bounded (using the triangular inequality) by the sum of the Hausdorff distances of `X` and `s` (in the coupling or, equivalently in the original space), of `s` and diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index f82894ad72d45..fc2e8017e8830 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -461,7 +461,7 @@ def optimal_GH_injl (x : X) : optimal_GH_coupling X Y := ⟦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_emetric_iff_metric.2 (λx y, _), + 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 @@ -472,7 +472,7 @@ def optimal_GH_injr (y : Y) : optimal_GH_coupling X Y := ⟦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_emetric_iff_metric.2 (λx y, _), + 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 diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index c7dd5d2fb9c5d..d92a4530380cd 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -30,130 +30,137 @@ between pseudoemetric spaces, or equivalently the distance between pseudometric def isometry [pseudo_emetric_space α] [pseudo_emetric_space β] (f : α → β) : Prop := ∀x1 x2 : α, edist (f x1) (f x2) = edist x1 x2 +/-- On pseudometric spaces, a map is an isometry if and only if it preserves nonnegative +distances. -/ +lemma isometry_iff_nndist_eq [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} : + isometry f ↔ (∀x y, nndist (f x) (f y) = nndist x y) := +by simp only [isometry, edist_nndist, ennreal.coe_eq_coe] + /-- On pseudometric spaces, a map is an isometry if and only if it preserves distances. -/ -lemma isometry_emetric_iff_metric [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} : +lemma isometry_iff_dist_eq [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} : isometry f ↔ (∀x y, dist (f x) (f y) = dist x y) := -⟨assume H x y, by simp [dist_edist, H x y], -assume H x y, by simp [edist_dist, H x y]⟩ - -/-- An isometry preserves edistances. -/ -theorem isometry.edist_eq [pseudo_emetric_space α] [pseudo_emetric_space β] {f : α → β} - (hf : isometry f) (x y : α) : edist (f x) (f y) = edist x y := -hf x y +by simp only [isometry_iff_nndist_eq, ← coe_nndist, nnreal.coe_eq] /-- An isometry preserves distances. -/ -theorem isometry.dist_eq [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} - (hf : isometry f) (x y : α) : dist (f x) (f y) = dist x y := -by rw [dist_edist, dist_edist, hf] +alias isometry_iff_dist_eq ↔ isometry.dist_eq _ + +/-- A map that preserves distances is an isometry -/ +alias isometry_iff_dist_eq ↔ _ isometry.of_dist_eq /-- An isometry preserves non-negative distances. -/ -theorem isometry.nndist_eq [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} - (hf : isometry f) (x y : α) : nndist (f x) (f y) = nndist x y := -subtype.ext $ hf.dist_eq x y +alias isometry_iff_nndist_eq ↔ isometry.nndist_eq _ + +/-- A map that preserves non-negative distances is an isometry. -/ +alias isometry_iff_nndist_eq ↔ _ isometry.of_nndist_eq + +namespace isometry section pseudo_emetric_isometry variables [pseudo_emetric_space α] [pseudo_emetric_space β] [pseudo_emetric_space γ] variables {f : α → β} {x y z : α} {s : set α} -lemma isometry.lipschitz (h : isometry f) : lipschitz_with 1 f := -lipschitz_with.of_edist_le $ λ x y, le_of_eq (h x y) +/-- An isometry preserves edistances. -/ +theorem edist_eq (hf : isometry f) (x y : α) : edist (f x) (f y) = edist x y := hf x y -lemma isometry.antilipschitz (h : isometry f) : antilipschitz_with 1 f := -λ x y, by simp only [h x y, ennreal.coe_one, one_mul, le_refl] +lemma lipschitz (h : isometry f) : lipschitz_with 1 f := +lipschitz_with.of_edist_le $ λ x y, (h x y).le -/-- An isometry from an emetric space is injective -/ -lemma isometry.injective {α : Type u} [emetric_space α] {f : α → β} (h : isometry f) : - injective f := h.antilipschitz.injective +lemma antilipschitz (h : isometry f) : antilipschitz_with 1 f := +λ x y, by simp only [h x y, ennreal.coe_one, one_mul, le_refl] /-- Any map on a subsingleton is an isometry -/ -theorem isometry_subsingleton [subsingleton α] : isometry f := +@[nontriviality] theorem _root_.isometry_subsingleton [subsingleton α] : isometry f := λx y, by rw subsingleton.elim x y; simp /-- The identity is an isometry -/ -lemma isometry_id : isometry (id : α → α) := -λx y, rfl +lemma _root_.isometry_id : isometry (id : α → α) := λ x y, rfl /-- The composition of isometries is an isometry -/ -theorem isometry.comp {g : β → γ} {f : α → β} (hg : isometry g) (hf : isometry f) : - isometry (g ∘ f) := -assume x y, calc - edist ((g ∘ f) x) ((g ∘ f) y) = edist (f x) (f y) : hg _ _ - ... = edist x y : hf _ _ +theorem comp {g : β → γ} {f : α → β} (hg : isometry g) (hf : isometry f) : isometry (g ∘ f) := +λ x y, (hg _ _).trans (hf _ _) + +/-- An isometry from a metric space is a uniform continuous map -/ +protected theorem uniform_continuous (hf : isometry f) : uniform_continuous f := +hf.lipschitz.uniform_continuous /-- An isometry from a metric space is a uniform inducing map -/ -theorem isometry.uniform_inducing (hf : isometry f) : - uniform_inducing f := -hf.antilipschitz.uniform_inducing hf.lipschitz.uniform_continuous +protected theorem uniform_inducing (hf : isometry f) : uniform_inducing f := +hf.antilipschitz.uniform_inducing hf.uniform_continuous -lemma isometry.tendsto_nhds_iff {ι : Type*} {f : α → β} - {g : ι → α} {a : filter ι} {b : α} (hf : isometry f) : +lemma tendsto_nhds_iff {ι : Type*} {f : α → β} {g : ι → α} {a : filter ι} {b : α} + (hf : isometry f) : filter.tendsto g a (𝓝 b) ↔ filter.tendsto (f ∘ g) a (𝓝 (f b)) := hf.uniform_inducing.inducing.tendsto_nhds_iff /-- An isometry is continuous. -/ -lemma isometry.continuous (hf : isometry f) : continuous f := -hf.lipschitz.continuous +protected lemma continuous (hf : isometry f) : continuous f := hf.lipschitz.continuous /-- The right inverse of an isometry is an isometry. -/ -lemma isometry.right_inv {f : α → β} {g : β → α} (h : isometry f) (hg : right_inverse g f) : +lemma right_inv {f : α → β} {g : β → α} (h : isometry f) (hg : right_inverse g f) : isometry g := λ x y, by rw [← h, hg _, hg _] +lemma preimage_emetric_closed_ball (h : isometry f) (x : α) (r : ℝ≥0∞) : + f ⁻¹' (emetric.closed_ball (f x) r) = emetric.closed_ball x r := +by { ext y, simp [h.edist_eq] } + +lemma preimage_emetric_ball (h : isometry f) (x : α) (r : ℝ≥0∞) : + f ⁻¹' (emetric.ball (f x) r) = emetric.ball x r := +by { ext y, simp [h.edist_eq] } + /-- Isometries preserve the diameter in pseudoemetric spaces. -/ -lemma isometry.ediam_image (hf : isometry f) (s : set α) : - emetric.diam (f '' s) = emetric.diam s := +lemma ediam_image (hf : isometry f) (s : set α) : emetric.diam (f '' s) = emetric.diam s := eq_of_forall_ge_iff $ λ d, by simp only [emetric.diam_le_iff, ball_image_iff, hf.edist_eq] -lemma isometry.ediam_range (hf : isometry f) : - emetric.diam (range f) = emetric.diam (univ : set α) := +lemma ediam_range (hf : isometry f) : emetric.diam (range f) = emetric.diam (univ : set α) := by { rw ← image_univ, exact hf.ediam_image univ } -lemma isometry.maps_to_emetric_ball (hf : isometry f) (x : α) (r : ℝ≥0∞) : +lemma maps_to_emetric_ball (hf : isometry f) (x : α) (r : ℝ≥0∞) : maps_to f (emetric.ball x r) (emetric.ball (f x) r) := -λ y hy, by rwa [emetric.mem_ball, hf] +(hf.preimage_emetric_ball x r).ge -lemma isometry.maps_to_emetric_closed_ball (hf : isometry f) (x : α) (r : ℝ≥0∞) : +lemma maps_to_emetric_closed_ball (hf : isometry f) (x : α) (r : ℝ≥0∞) : maps_to f (emetric.closed_ball x r) (emetric.closed_ball (f x) r) := -λ y hy, by rwa [emetric.mem_closed_ball, hf] +(hf.preimage_emetric_closed_ball x r).ge /-- The injection from a subtype is an isometry -/ -lemma isometry_subtype_coe {s : set α} : isometry (coe : s → α) := +lemma _root_.isometry_subtype_coe {s : set α} : isometry (coe : s → α) := λx y, rfl -lemma isometry.comp_continuous_on_iff {γ} [topological_space γ] (hf : isometry f) {g : γ → α} - {s : set γ} : +lemma comp_continuous_on_iff {γ} [topological_space γ] (hf : isometry f) {g : γ → α} {s : set γ} : continuous_on (f ∘ g) s ↔ continuous_on g s := hf.uniform_inducing.inducing.continuous_on_iff.symm -lemma isometry.comp_continuous_iff {γ} [topological_space γ] (hf : isometry f) {g : γ → α} : +lemma comp_continuous_iff {γ} [topological_space γ] (hf : isometry f) {g : γ → α} : continuous (f ∘ g) ↔ continuous g := hf.uniform_inducing.inducing.continuous_iff.symm end pseudo_emetric_isometry --section section emetric_isometry -variables [emetric_space α] +variables [emetric_space α] [pseudo_emetric_space β] {f : α → β} -/-- An isometry from a metric space is a uniform embedding -/ -theorem isometry.uniform_embedding [pseudo_emetric_space β] {f : α → β} (hf : isometry f) : - uniform_embedding f := +/-- An isometry from an emetric space is injective -/ +protected lemma injective (h : isometry f) : injective f := h.antilipschitz.injective + +/-- An isometry from an emetric space is a uniform embedding -/ +protected theorem uniform_embedding (hf : isometry f) : uniform_embedding f := hf.antilipschitz.uniform_embedding hf.lipschitz.uniform_continuous -/-- An isometry from a metric space is an embedding -/ -theorem isometry.embedding [pseudo_emetric_space β] {f : α → β} (hf : isometry f) : - embedding f := +/-- An isometry from an emetric space is an embedding -/ +protected theorem embedding (hf : isometry f) : embedding f := hf.uniform_embedding.embedding /-- An isometry from a complete emetric space is a closed embedding -/ -theorem isometry.closed_embedding [complete_space α] [emetric_space β] - {f : α → β} (hf : isometry f) : closed_embedding f := +theorem closed_embedding [complete_space α] [emetric_space γ] + {f : α → γ} (hf : isometry f) : closed_embedding f := hf.antilipschitz.closed_embedding hf.lipschitz.uniform_continuous end emetric_isometry --section -namespace isometry +section pseudo_metric_isometry variables [pseudo_metric_space α] [pseudo_metric_space β] {f : α → β} @@ -164,19 +171,37 @@ by rw [metric.diam, metric.diam, hf.ediam_image] lemma diam_range (hf : isometry f) : metric.diam (range f) = metric.diam (univ : set α) := by { rw ← image_univ, exact hf.diam_image univ } +lemma preimage_set_of_dist (hf : isometry f) (x : α) (p : ℝ → Prop) : + f ⁻¹' {y | p (dist y (f x))} = {y | p (dist y x)} := +by { ext y, simp [hf.dist_eq] } + +lemma preimage_closed_ball (hf : isometry f) (x : α) (r : ℝ) : + f ⁻¹' (metric.closed_ball (f x) r) = metric.closed_ball x r := +hf.preimage_set_of_dist x (≤ r) + +lemma preimage_ball (hf : isometry f) (x : α) (r : ℝ) : + f ⁻¹' (metric.ball (f x) r) = metric.ball x r := +hf.preimage_set_of_dist x (< r) + +lemma preimage_sphere (hf : isometry f) (x : α) (r : ℝ) : + f ⁻¹' (metric.sphere (f x) r) = metric.sphere x r := +hf.preimage_set_of_dist x (= r) + lemma maps_to_ball (hf : isometry f) (x : α) (r : ℝ) : maps_to f (metric.ball x r) (metric.ball (f x) r) := -λ y hy, by rwa [metric.mem_ball, hf.dist_eq] +(hf.preimage_ball x r).ge lemma maps_to_sphere (hf : isometry f) (x : α) (r : ℝ) : maps_to f (metric.sphere x r) (metric.sphere (f x) r) := -λ y hy, by rwa [metric.mem_sphere, hf.dist_eq] +(hf.preimage_sphere x r).ge lemma maps_to_closed_ball (hf : isometry f) (x : α) (r : ℝ) : maps_to f (metric.closed_ball x r) (metric.closed_ball (f x) r) := -λ y hy, by rwa [metric.mem_closed_ball, hf.dist_eq] +(hf.preimage_closed_ball x r).ge + +end pseudo_metric_isometry -- section -end isometry +end isometry -- namespace /-- A uniform embedding from a uniform space to a metric space is an isometry with respect to the induced metric space structure on the source space. -/ @@ -187,7 +212,7 @@ lemma uniform_embedding.to_isometry {α β} [uniform_space α] [metric_space β] (@metric_space.to_pseudo_metric_space α (h.comap_metric_space f))) (by apply_instance) f := begin - apply isometry_emetric_iff_metric.2, + apply isometry.of_dist_eq, assume x y, refl end @@ -201,7 +226,7 @@ lemma embedding.to_isometry {α β} [topological_space α] [metric_space β] {f (@metric_space.to_pseudo_metric_space α (h.comap_metric_space f))) (by apply_instance) f := begin - apply isometry_emetric_iff_metric.2, + apply isometry.of_dist_eq, assume x y, refl end @@ -330,11 +355,11 @@ by rw [← image_symm, ediam_image] @[simp] lemma preimage_emetric_ball (h : α ≃ᵢ β) (x : β) (r : ℝ≥0∞) : h ⁻¹' (emetric.ball x r) = emetric.ball (h.symm x) r := -by { ext y, simp [← h.edist_eq] } +by rw [← h.isometry.preimage_emetric_ball (h.symm x) r, h.apply_symm_apply] @[simp] lemma preimage_emetric_closed_ball (h : α ≃ᵢ β) (x : β) (r : ℝ≥0∞) : h ⁻¹' (emetric.closed_ball x r) = emetric.closed_ball (h.symm x) r := -by { ext y, simp [← h.edist_eq] } +by rw [← h.isometry.preimage_emetric_closed_ball (h.symm x) r, h.apply_symm_apply] @[simp] lemma image_emetric_ball (h : α ≃ᵢ β) (x : α) (r : ℝ≥0∞) : h '' (emetric.ball x r) = emetric.ball (h x) r := @@ -411,15 +436,15 @@ congr_arg ennreal.to_real h.ediam_univ @[simp] lemma preimage_ball (h : α ≃ᵢ β) (x : β) (r : ℝ) : h ⁻¹' (metric.ball x r) = metric.ball (h.symm x) r := -by { ext y, simp [← h.dist_eq] } +by rw [← h.isometry.preimage_ball (h.symm x) r, h.apply_symm_apply] @[simp] lemma preimage_sphere (h : α ≃ᵢ β) (x : β) (r : ℝ) : h ⁻¹' (metric.sphere x r) = metric.sphere (h.symm x) r := -by { ext y, simp [← h.dist_eq] } +by rw [← h.isometry.preimage_sphere (h.symm x) r, h.apply_symm_apply] @[simp] lemma preimage_closed_ball (h : α ≃ᵢ β) (x : β) (r : ℝ) : h ⁻¹' (metric.closed_ball x r) = metric.closed_ball (h.symm x) r := -by { ext y, simp [← h.dist_eq] } +by rw [← h.isometry.preimage_closed_ball (h.symm x) r, h.apply_symm_apply] @[simp] lemma image_ball (h : α ≃ᵢ β) (x : α) (r : ℝ) : h '' (metric.ball x r) = metric.ball (h x) r := diff --git a/src/topology/metric_space/kuratowski.lean b/src/topology/metric_space/kuratowski.lean index 876a40798238b..b7107bedd1658 100644 --- a/src/topology/metric_space/kuratowski.lean +++ b/src/topology/metric_space/kuratowski.lean @@ -54,7 +54,7 @@ end /-- When the reference set is dense, the embedding map is an isometry on its image. -/ lemma embedding_of_subset_isometry (H : dense_range x) : isometry (embedding_of_subset x) := begin - refine isometry_emetric_iff_metric.2 (λa b, _), + refine isometry.of_dist_eq (λa b, _), refine (embedding_of_subset_dist_le x a b).antisymm (le_of_forall_pos_le_add (λe epos, _)), /- First step: find n with dist a (x n) < e -/ rcases metric.mem_closure_range_iff.1 (H a) (e/2) (half_pos epos) with ⟨n, hn⟩, From 63708058d9f8cfb62adaf15884b6c63a85a79ac6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 22 Jul 2022 18:27:53 +0000 Subject: [PATCH 007/173] =?UTF-8?q?move(order/{boolean=5Falgebra=20?= =?UTF-8?q?=E2=86=92=20basic}):=20move=20`has=5Fcompl`=20and=20trivial=20i?= =?UTF-8?q?nstances=20(#15602)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We move the trivial `pi.has_sdiff`, `pi.has_compl`, and `Prop.has_compl` instances to `order/basic.lean`. The main effect of this is that `rᶜ` for the complement of a relation is now available basically anywhere. --- src/order/basic.lean | 37 +++++++++++++++++++++++++++++ src/order/boolean_algebra.lean | 43 ++-------------------------------- 2 files changed, 39 insertions(+), 41 deletions(-) diff --git a/src/order/basic.lean b/src/order/basic.lean index cd37bd0a2f8a2..99ca82e76f1fe 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -29,6 +29,7 @@ classes and allows to transfer order instances. * `has_sup`: type class for the `⊔` notation * `has_inf`: type class for the `⊓` notation +* `has_compl`: type class for the `ᶜ` notation * `densely_ordered`: An order with no gap, i.e. for any two elements `a < b` there exists `c` such that `a < c < b`. @@ -446,6 +447,31 @@ linear_order.ext $ λ _ _, iff.rfl end order_dual +/-! ### `has_compl` -/ + +/-- Set / lattice complement -/ +@[notation_class] class has_compl (α : Type*) := (compl : α → α) + +export has_compl (compl) + +postfix `ᶜ`:(max+1) := compl + +instance Prop.has_compl : has_compl Prop := ⟨not⟩ + +instance pi.has_compl {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] : + has_compl (Π i, α i) := +⟨λ x i, (x i)ᶜ⟩ + +lemma pi.compl_def {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] (x : Π i, α i) : + xᶜ = λ i, (x i)ᶜ := rfl + +@[simp] +lemma pi.compl_apply {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] (x : Π i, α i) (i : ι) : + xᶜ i = (x i)ᶜ := rfl + +instance is_irrefl.compl (r) [is_irrefl α r] : is_refl α rᶜ := ⟨@irrefl α r _⟩ +instance is_refl.compl (r) [is_refl α r] : is_irrefl α rᶜ := ⟨λ a, not_not_intro (refl a)⟩ + /-! ### Order instances on the function space -/ instance pi.has_le {ι : Type u} {α : ι → Type v} [∀ i, has_le (α i)] : has_le (Π i, α i) := @@ -484,6 +510,17 @@ instance pi.partial_order {ι : Type u} {α : ι → Type v} [∀ i, partial_ord { le_antisymm := λ f g h1 h2, funext (λ b, (h1 b).antisymm (h2 b)), ..pi.preorder } +instance pi.has_sdiff {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] : + has_sdiff (Π i, α i) := +⟨λ x y i, x i \ y i⟩ + +lemma pi.sdiff_def {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] (x y : Π i, α i) : + (x \ y) = λ i, x i \ y i := rfl + +@[simp] +lemma pi.sdiff_apply {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] (x y : Π i, α i) (i : ι) : + (x \ y) i = x i \ y i := rfl + /-! ### `min`/`max` recursors -/ section min_max_rec diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index cb52595dfcfa7..587fe912136f4 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -26,7 +26,6 @@ intervals.) ## Main declarations -* `has_compl`: a type class for the complement operator * `generalized_boolean_algebra`: a type class for generalized Boolean algebras * `boolean_algebra`: a type class for Boolean algebras. * `Prop.boolean_algebra`: the Boolean algebra instance on `Prop` @@ -42,11 +41,6 @@ complement operator `a \ b` for all `a`, `b`. Instead, the postulates there amou that for all `a, b : α` where `a ≤ b`, the equations `x ⊔ a = b` and `x ⊓ a = ⊥` have a solution `x`. `disjoint.sdiff_unique` proves that this `x` is in fact `b \ a`. -## Notations - -* `xᶜ` is notation for `compl x` -* `x \ y` is notation for `sdiff x y`. - ## References * @@ -586,14 +580,6 @@ end generalized_boolean_algebra ### Boolean algebras -/ - -/-- Set / lattice complement -/ -@[notation_class] class has_compl (α : Type*) := (compl : α → α) - -export has_compl (compl) - -postfix `ᶜ`:(max+1) := compl - /-- A Boolean algebra is a bounded distributive lattice with a complement operator `ᶜ` such that `x ⊓ xᶜ = ⊥` and `x ⊔ xᶜ = ⊤`. For convenience, it must also provide a set difference operation `\` satisfying `x \ y = x ⊓ yᶜ`. @@ -773,37 +759,12 @@ by rw [sdiff_eq, compl_inf, compl_compl] end boolean_algebra instance Prop.boolean_algebra : boolean_algebra Prop := -{ compl := not, - inf_compl_le_bot := λ p ⟨Hp, Hpc⟩, Hpc Hp, +{ inf_compl_le_bot := λ p ⟨Hp, Hpc⟩, Hpc Hp, top_le_sup_compl := λ p H, classical.em p, + .. Prop.has_compl, .. Prop.distrib_lattice, .. Prop.bounded_order } -instance pi.has_sdiff {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] : - has_sdiff (Π i, α i) := -⟨λ x y i, x i \ y i⟩ - -lemma pi.sdiff_def {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] (x y : Π i, α i) : - (x \ y) = λ i, x i \ y i := rfl - -@[simp] -lemma pi.sdiff_apply {ι : Type u} {α : ι → Type v} [∀ i, has_sdiff (α i)] (x y : Π i, α i) (i : ι) : - (x \ y) i = x i \ y i := rfl - -instance pi.has_compl {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] : - has_compl (Π i, α i) := -⟨λ x i, (x i)ᶜ⟩ - -lemma pi.compl_def {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] (x : Π i, α i) : - xᶜ = λ i, (x i)ᶜ := rfl - -instance is_irrefl.compl (r) [is_irrefl α r] : is_refl α rᶜ := ⟨@irrefl α r _⟩ -instance is_refl.compl (r) [is_refl α r] : is_irrefl α rᶜ := ⟨λ a, not_not_intro (refl a)⟩ - -@[simp] -lemma pi.compl_apply {ι : Type u} {α : ι → Type v} [∀ i, has_compl (α i)] (x : Π i, α i) (i : ι) : - xᶜ i = (x i)ᶜ := rfl - instance pi.boolean_algebra {ι : Type u} {α : ι → Type v} [∀ i, boolean_algebra (α i)] : boolean_algebra (Π i, α i) := { sdiff_eq := λ x y, funext $ λ i, sdiff_eq, From c751bda87504e38d9cef63c0ede561f91e5edecc Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 22 Jul 2022 22:11:06 +0000 Subject: [PATCH 008/173] feat(analysis/locally_convex/with_seminorms): in a normed space, von Neumann bounded = metric bounded (#15124) --- src/analysis/locally_convex/bounded.lean | 74 ++++++++++++++++++++++++ src/topology/metric_space/basic.lean | 6 ++ 2 files changed, 80 insertions(+) diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index f092e5dfe54b3..9872c2549ca24 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ import analysis.locally_convex.basic +import analysis.seminorm import topology.bornology.basic import topology.algebra.uniform_group import analysis.locally_convex.balanced_core_hull @@ -190,3 +191,76 @@ begin end end uniform_add_group + +section vonN_bornology_eq_metric + +variables (𝕜 E) [nondiscrete_normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] + +namespace normed_space + +lemma is_vonN_bounded_ball (r : ℝ) : + bornology.is_vonN_bounded 𝕜 (metric.ball (0 : E) r) := +begin + rw [metric.nhds_basis_ball.is_vonN_bounded_basis_iff, ← ball_norm_seminorm 𝕜 E], + exact λ ε hε, (norm_seminorm 𝕜 E).ball_zero_absorbs_ball_zero hε +end + +lemma is_vonN_bounded_closed_ball (r : ℝ) : + bornology.is_vonN_bounded 𝕜 (metric.closed_ball (0 : E) r) := +(is_vonN_bounded_ball 𝕜 E (r+1)).subset (metric.closed_ball_subset_ball $ by linarith) + +lemma is_vonN_bounded_iff (s : set E) : + bornology.is_vonN_bounded 𝕜 s ↔ bornology.is_bounded s := +begin + rw [← metric.bounded_iff_is_bounded, metric.bounded_iff_subset_ball (0 : E)], + split, + { intros h, + 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), + 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 } +end + +/-- In a normed space, the von Neumann bornology (`bornology.vonN_bornology`) is equal to the +metric bornology. -/ +lemma vonN_bornology_eq : bornology.vonN_bornology 𝕜 E = pseudo_metric_space.to_bornology := +begin + rw bornology.ext_iff_is_bounded, + intro s, + rw bornology.is_bounded_iff_is_vonN_bounded, + exact is_vonN_bounded_iff 𝕜 E s +end + +variable (𝕜) + +lemma is_bounded_iff_subset_smul_ball {s : set E} : + bornology.is_bounded s ↔ ∃ a : 𝕜, s ⊆ a • metric.ball 0 1 := +begin + rw ← is_vonN_bounded_iff 𝕜, + split, + { intros h, + rcases h (metric.ball_mem_nhds 0 zero_lt_one) with ⟨ρ, hρ, hρball⟩, + rcases normed_field.exists_lt_norm 𝕜 ρ with ⟨a, ha⟩, + exact ⟨a, hρball a ha.le⟩ }, + { rintros ⟨a, ha⟩, + exact ((is_vonN_bounded_ball 𝕜 E 1).image (a • 1 : E →L[𝕜] E)).subset ha } +end + +lemma is_bounded_iff_subset_smul_closed_ball {s : set E} : + bornology.is_bounded s ↔ ∃ a : 𝕜, s ⊆ a • metric.closed_ball 0 1 := +begin + split, + { rw is_bounded_iff_subset_smul_ball 𝕜, + exact exists_imp_exists + (λ a ha, ha.trans $ set.smul_set_mono $ metric.ball_subset_closed_ball) }, + { rw ← is_vonN_bounded_iff 𝕜, + rintros ⟨a, ha⟩, + exact ((is_vonN_bounded_closed_ball 𝕜 E 1).image (a • 1 : E →L[𝕜] E)).subset ha } +end + +end normed_space + +end vonN_bornology_eq_metric diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index cd0cafe93a0c2..e45791b7d41e0 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2040,6 +2040,12 @@ def bounded (s : set α) : Prop := section bounded variables {x : α} {s t : set α} {r : ℝ} +lemma bounded_iff_is_bounded (s : set α) : bounded s ↔ is_bounded s := +begin + change bounded s ↔ sᶜ ∈ (cobounded α).sets, + simp [pseudo_metric_space.cobounded_sets, metric.bounded], +end + @[simp] lemma bounded_empty : bounded (∅ : set α) := ⟨0, by simp⟩ From b8b4346b2c12d5307d70377f29f877f83aedb070 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 22 Jul 2022 22:11:08 +0000 Subject: [PATCH 009/173] feat(data/polynomial/degree/definitions): redefine `polynomial.degree` as `p.support.max` (#15199) This PR redefines `polynomial.degree p`: * old: `p.support.sup coe` * new: `p.support.max`. The two definitions are defeq and relatively few changes are required. Weirdness: `open finset` seems to no longer work consistently. This is the largest source of differences. In particular, the file `ring_theory/polynomial/cyclotomic/basic` only changed because I added `finset.` in several places. --- src/data/finset/lattice.lean | 2 + src/data/mv_polynomial/equiv.lean | 3 +- src/data/polynomial/degree/definitions.lean | 14 ++--- src/data/polynomial/reverse.lean | 4 +- src/ring_theory/polynomial/basic.lean | 2 +- .../polynomial/cyclotomic/basic.lean | 14 ++--- .../polynomial/cyclotomic/eval.lean | 51 ++++++++++--------- 7 files changed, 48 insertions(+), 42 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 4026aa3297a79..a5acdb750bd9e 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -809,6 +809,8 @@ and `⊥` otherwise. It belongs to `with_bot α`. If you want to get an element protected def max (s : finset α) : with_bot α := sup s coe +lemma max_eq_sup_coe {s : finset α} : s.max = s.sup coe := rfl + theorem max_eq_sup_with_bot (s : finset α) : s.max = sup s coe := rfl diff --git a/src/data/mv_polynomial/equiv.lean b/src/data/mv_polynomial/equiv.lean index c52ecf4955f05..bbdf61fa66fa7 100644 --- a/src/data/mv_polynomial/equiv.lean +++ b/src/data/mv_polynomial/equiv.lean @@ -459,7 +459,8 @@ lemma degree_fin_succ_equiv {f : mv_polynomial (fin (n + 1)) R} (h : f ≠ 0) : begin have h' : (fin_succ_equiv R n f).support.sup (λ x , x) = degree_of 0 f, { rw [degree_of_eq_sup, fin_succ_equiv_support f, finset.sup_image] }, - rw [polynomial.degree, ← h', finset.coe_sup_of_nonempty (support_fin_succ_equiv_nonempty h)], + rw [polynomial.degree, ← h', finset.coe_sup_of_nonempty (support_fin_succ_equiv_nonempty h), + finset.max_eq_sup_coe], end lemma nat_degree_fin_succ_equiv (f : mv_polynomial (fin (n + 1)) R) : diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 3c87ca28129dd..f29b94445b659 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -34,7 +34,7 @@ variables [semiring R] {p q r : R[X]} /-- `degree p` is the degree of the polynomial `p`, i.e. the largest `X`-exponent in `p`. `degree p = some n` when `p ≠ 0` and `n` is the highest power of `X` that appears in `p`, otherwise `degree 0 = ⊥`. -/ -def degree (p : R[X]) : with_bot ℕ := p.support.sup coe +def degree (p : R[X]) : with_bot ℕ := p.support.max lemma degree_lt_wf : well_founded (λp q : R[X], degree p < degree q) := inv_image.wf degree (with_bot.well_founded_lt nat.lt_wf) @@ -161,7 +161,8 @@ lemma nat_degree_le_nat_degree [semiring S] {q : S[X]} (hpq : p.degree ≤ q.deg with_bot.gi_get_or_else_bot.gc.monotone_l hpq @[simp] lemma degree_C (ha : a ≠ 0) : degree (C a) = (0 : with_bot ℕ) := -by { rw [degree, ← monomial_zero_left, support_monomial 0 ha, sup_singleton], refl } +by rw [degree, ← monomial_zero_left, support_monomial 0 ha, max_eq_sup_coe, sup_singleton, + with_bot.coe_zero] lemma degree_C_le : degree (C a) ≤ 0 := begin @@ -565,7 +566,7 @@ begin rw [degree, support_update], split_ifs, { exact (finset.max_mono (erase_subset _ _)).trans (le_max_left _ _) }, - { rw [sup_insert, max_comm], + { rw [max_insert, max_comm], exact le_rfl }, end @@ -836,14 +837,15 @@ by rw [← nonpos_iff_eq_zero, nat_degree_le_iff_degree_le, with_bot.coe_zero] theorem degree_le_iff_coeff_zero (f : R[X]) (n : with_bot ℕ) : degree f ≤ n ↔ ∀ m : ℕ, n < m → coeff f m = 0 := -by simp only [degree, finset.sup_le_iff, mem_support_iff, ne.def, ← not_le, not_imp_comm] +by simp only [degree, finset.max, finset.sup_le_iff, mem_support_iff, ne.def, ← not_le, + not_imp_comm] theorem degree_lt_iff_coeff_zero (f : R[X]) (n : ℕ) : degree f < n ↔ ∀ m : ℕ, n ≤ m → coeff f m = 0 := begin refine ⟨λ hf m hm, coeff_eq_zero_of_degree_lt (lt_of_lt_of_le hf (with_bot.coe_le_coe.2 hm)), _⟩, simp only [degree, finset.sup_lt_iff (with_bot.bot_lt_coe n), mem_support_iff, - with_bot.some_eq_coe, with_bot.coe_lt_coe, ← @not_le ℕ], + with_bot.some_eq_coe, with_bot.coe_lt_coe, ← @not_le ℕ, max_eq_sup_coe], exact λ h m, mt (h m), end @@ -875,7 +877,7 @@ lemma eq_C_of_nat_degree_eq_zero (h : nat_degree p = 0) : p = C (coeff p 0) := eq_C_of_nat_degree_le_zero h.le lemma ne_zero_of_coe_le_degree (hdeg : ↑n ≤ p.degree) : p ≠ 0 := -by rw ← degree_nonneg_iff_ne_zero; exact trans (by exact_mod_cast n.zero_le) hdeg +degree_nonneg_iff_ne_zero.mp $ (with_bot.coe_le_coe.mpr n.zero_le).trans hdeg lemma le_nat_degree_of_coe_le_degree (hdeg : ↑n ≤ p.degree) : n ≤ p.nat_degree := diff --git a/src/data/polynomial/reverse.lean b/src/data/polynomial/reverse.lean index 38373d31395af..c74102003ecff 100644 --- a/src/data/polynomial/reverse.lean +++ b/src/data/polynomial/reverse.lean @@ -88,11 +88,11 @@ noncomputable def reflect (N : ℕ) : R[X] → R[X] | ⟨f⟩ := ⟨finsupp.emb_domain (rev_at N) f⟩ lemma reflect_support (N : ℕ) (f : R[X]) : - (reflect N f).support = image (rev_at N) f.support := + (reflect N f).support = finset.image (rev_at N) f.support := begin rcases f, ext1, - rw [reflect, mem_image, support, support, support_emb_domain, mem_map], + simp only [reflect, support_of_finsupp, support_emb_domain, finset.mem_map, finset.mem_image], end @[simp] lemma coeff_reflect (N : ℕ) (f : R[X]) (i : ℕ) : diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index 38015760e7c95..affcf102a72bf 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -80,7 +80,7 @@ end theorem mem_degree_lt {n : ℕ} {f : R[X]} : f ∈ degree_lt R n ↔ degree f < n := -by { simp_rw [degree_lt, submodule.mem_infi, linear_map.mem_ker, degree, +by { simp_rw [degree_lt, submodule.mem_infi, linear_map.mem_ker, degree, finset.max_eq_sup_coe, finset.sup_lt_iff (with_bot.bot_lt_coe n), mem_support_iff, with_bot.coe_lt_coe, lt_iff_not_le, ne, not_imp_not], refl } diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 0fe57a78fdc29..2aba39c26b744 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -423,7 +423,7 @@ open_locale big_operators open finset lemma prod_cyclotomic_eq_geom_sum {n : ℕ} (h : 0 < n) (R) [comm_ring R] [is_domain R] : - ∏ i in n.divisors \ {1}, cyclotomic i R = ∑ i in range n, X ^ i := + ∏ i in n.divisors \ {1}, cyclotomic i R = ∑ i in finset.range n, X ^ i := begin apply_fun (* cyclotomic 1 R) using mul_left_injective₀ (cyclotomic_ne_zero 1 R), have : ∏ i in {1}, cyclotomic i R = cyclotomic 1 R := finset.prod_singleton, @@ -432,10 +432,10 @@ begin end lemma cyclotomic_dvd_geom_sum_of_dvd (R) [comm_ring R] {d n : ℕ} (hdn : d ∣ n) - (hd : d ≠ 1) : cyclotomic d R ∣ ∑ i in range n, X ^ i := + (hd : d ≠ 1) : cyclotomic d R ∣ ∑ i in finset.range n, X ^ i := begin suffices : (cyclotomic d ℤ).map (int.cast_ring_hom R) ∣ - (∑ i in range n, X ^ i).map (int.cast_ring_hom R), + (∑ i in finset.range n, X ^ i).map (int.cast_ring_hom R), { have key := (map_ring_hom (int.cast_ring_hom R)).map_geom_sum X n, simp only [coe_map_ring_hom, map_X] at key, rwa [map_cyclotomic, key] at this }, @@ -707,7 +707,7 @@ end /-- If `p` is prime, then `cyclotomic p R = ∑ i in range p, X ^ i`. -/ lemma cyclotomic_eq_geom_sum {R : Type*} [comm_ring R] {p : ℕ} - (hp : nat.prime p) : cyclotomic p R = ∑ i in range p, X ^ i := + (hp : nat.prime p) : cyclotomic p R = ∑ i in finset.range p, X ^ i := begin refine ((eq_cyclotomic_iff hp.pos _).mpr _).symm, simp only [nat.prime.proper_divisors hp, geom_sum_mul, finset.prod_singleton, cyclotomic_one], @@ -720,10 +720,10 @@ by rw [cyclotomic_eq_geom_sum hn.out, geom_sum_mul] /-- If `p ^ k` is a prime power, then `cyclotomic (p ^ (n + 1)) R = ∑ i in range p, (X ^ (p ^ n)) ^ i`. -/ lemma cyclotomic_prime_pow_eq_geom_sum {R : Type*} [comm_ring R] {p n : ℕ} (hp : nat.prime p) : - cyclotomic (p ^ (n + 1)) R = ∑ i in range p, (X ^ (p ^ n)) ^ i := + cyclotomic (p ^ (n + 1)) R = ∑ i in finset.range p, (X ^ (p ^ n)) ^ i := begin - have : ∀ m, cyclotomic (p ^ (m + 1)) R = ∑ i in range p, (X ^ (p ^ m)) ^ i ↔ - (∑ i in range p, (X ^ (p ^ m)) ^ i) * ∏ (x : ℕ) in finset.range (m + 1), + have : ∀ m, cyclotomic (p ^ (m + 1)) R = ∑ i in finset.range p, (X ^ (p ^ m)) ^ i ↔ + (∑ i in finset.range p, (X ^ (p ^ m)) ^ i) * ∏ (x : ℕ) in finset.range (m + 1), cyclotomic (p ^ x) R = X ^ p ^ (m + 1) - 1, { intro m, have := eq_cyclotomic_iff (pow_pos hp.pos (m + 1)) _, diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 1c1fd60e0196d..06eff72e88654 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -26,8 +26,8 @@ open_locale big_operators @[simp] lemma eval_one_cyclotomic_prime {R : Type*} [comm_ring R] {p : ℕ} [hn : fact p.prime] : eval 1 (cyclotomic p R) = p := -by simp only [cyclotomic_eq_geom_sum hn.out, eval_X, one_pow, sum_const, eval_pow, - eval_finset_sum, card_range, smul_one_eq_coe] +by simp only [cyclotomic_eq_geom_sum hn.out, eval_X, one_pow, finset.sum_const, eval_pow, + eval_finset_sum, finset.card_range, smul_one_eq_coe] @[simp] lemma eval₂_one_cyclotomic_prime {R S : Type*} [comm_ring R] [semiring S] (f : R →+* S) {p : ℕ} [fact p.prime] : eval₂ f 1 (cyclotomic p R) = p := @@ -35,8 +35,8 @@ by simp @[simp] lemma eval_one_cyclotomic_prime_pow {R : Type*} [comm_ring R] {p : ℕ} (k : ℕ) [hn : fact p.prime] : eval 1 (cyclotomic (p ^ (k + 1)) R) = p := -by simp only [cyclotomic_prime_pow_eq_geom_sum hn.out, eval_X, one_pow, sum_const, - eval_pow, eval_finset_sum, card_range, smul_one_eq_coe] +by simp only [cyclotomic_prime_pow_eq_geom_sum hn.out, eval_X, one_pow, finset.sum_const, + eval_pow, eval_finset_sum, finset.card_range, smul_one_eq_coe] @[simp] lemma eval₂_one_cyclotomic_prime_pow {R S : Type*} [comm_ring R] [semiring S] (f : R →+* S) {p : ℕ} (k : ℕ) [fact p.prime] : eval₂ f 1 (cyclotomic (p ^ (k + 1)) R) = p := @@ -71,18 +71,18 @@ 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', insert_sdiff_of_not_mem, - prod_insert, eval_mul, eval_geom_sum] at this, + rw [divisors_eq_proper_divisors_insert_self_of_pos hn', finset.insert_sdiff_of_not_mem, + finset.prod_insert, eval_mul, eval_geom_sum] at this, rotate, - { simp only [lt_self_iff_false, mem_sdiff, not_false_iff, mem_proper_divisors, and_false, + { simp only [lt_self_iff_false, finset.mem_sdiff, not_false_iff, mem_proper_divisors, and_false, false_and]}, - { simpa only [mem_singleton] using hn''.ne' }, - rcases lt_trichotomy 0 (∑ i in range n, x ^ i) with h | h | h, + { simpa only [finset.mem_singleton] using hn''.ne' }, + rcases lt_trichotomy 0 (∑ i in finset.range n, x ^ i) with h | h | h, { apply pos_of_mul_pos_left, { rwa this }, rw eval_prod, - refine prod_nonneg (λ i hi, _), - simp only [mem_sdiff, mem_proper_divisors, mem_singleton] at hi, + refine finset.prod_nonneg (λ i hi, _), + simp only [finset.mem_sdiff, mem_proper_divisors, finset.mem_singleton] at hi, rw geom_sum_pos_iff hn'' at h, cases h with hk hx, { refine (ih _ hi.1.2 (nat.two_lt_of_ne _ hi.2 _)).le; rintro rfl, @@ -99,12 +99,13 @@ begin { rwa this }, rw [geom_sum_neg_iff hn''] at h, have h2 : {2} ⊆ n.proper_divisors \ {1}, - { rw [singleton_subset_iff, mem_sdiff, mem_proper_divisors, not_mem_singleton], + { rw [finset.singleton_subset_iff, finset.mem_sdiff, mem_proper_divisors, + finset.not_mem_singleton], exact ⟨⟨even_iff_two_dvd.mp h.1, hn⟩, (nat.one_lt_bit0 one_ne_zero).ne'⟩ }, - rw [eval_prod, ←prod_sdiff h2, prod_singleton]; try { apply_instance }, + rw [eval_prod, ←finset.prod_sdiff h2, finset.prod_singleton]; try { apply_instance }, apply mul_nonpos_of_nonneg_of_nonpos, - { refine prod_nonneg (λ i hi, le_of_lt _), - simp only [mem_sdiff, mem_proper_divisors, mem_singleton] at hi, + { refine finset.prod_nonneg (λ i hi, le_of_lt _), + simp only [finset.mem_sdiff, mem_proper_divisors, finset.mem_singleton] at hi, refine ih _ hi.1.1.2 (nat.two_lt_of_ne _ hi.1.2 hi.2), rintro rfl, rw zero_dvd_iff at hi, @@ -220,15 +221,15 @@ begin simp [cyclotomic_nonneg n hq'.le], }, simp only [cyclotomic_eq_prod_X_sub_primitive_roots hζ, eval_prod, eval_C, eval_X, eval_sub, nnnorm_prod, units.mk0_prod], - convert prod_lt_prod' _ _, + convert finset.prod_lt_prod' _ _, swap, { exact λ _, units.mk0 (real.to_nnreal (q - 1)) (by simp [hq']) }, { simp [complex.card_primitive_roots] }, - { simp only [subtype.coe_mk, mem_attach, forall_true_left, subtype.forall, ←units.coe_le_coe, - ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, real.coe_to_nnreal', - coe_nnnorm, complex.norm_eq_abs, max_le_iff, tsub_le_iff_right], + { simp only [subtype.coe_mk, finset.mem_attach, forall_true_left, subtype.forall, + ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, + real.coe_to_nnreal', coe_nnnorm, complex.norm_eq_abs, max_le_iff, tsub_le_iff_right], intros x hx, simpa using hfor x hx, }, - { simp only [subtype.coe_mk, mem_attach, exists_true_left, subtype.exists, + { simp only [subtype.coe_mk, finset.mem_attach, exists_true_left, subtype.exists, ← nnreal.coe_lt_coe, ← units.coe_lt_coe, units.coe_mk0 _, coe_nnnorm], simpa [hq'.le] using hex, }, end @@ -283,16 +284,16 @@ begin linarith }, simp only [cyclotomic_eq_prod_X_sub_primitive_roots hζ, eval_prod, eval_C, eval_X, eval_sub, nnnorm_prod, units.mk0_prod], - convert prod_lt_prod' _ _, + convert finset.prod_lt_prod' _ _, swap, { exact λ _, units.mk0 (real.to_nnreal (q + 1)) (by simp; linarith only [hq']) }, { simp [complex.card_primitive_roots], }, - { simp only [subtype.coe_mk, mem_attach, forall_true_left, subtype.forall, ←units.coe_le_coe, - ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, real.coe_to_nnreal, - coe_nnnorm, complex.norm_eq_abs, max_le_iff], + { simp only [subtype.coe_mk, finset.mem_attach, forall_true_left, subtype.forall, + ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, + real.coe_to_nnreal, coe_nnnorm, complex.norm_eq_abs, max_le_iff], intros x hx, have : complex.abs _ ≤ _ := hfor x hx, simp [this], }, - { simp only [subtype.coe_mk, mem_attach, exists_true_left, subtype.exists, + { simp only [subtype.coe_mk, finset.mem_attach, exists_true_left, subtype.exists, ← nnreal.coe_lt_coe, ← units.coe_lt_coe, units.coe_mk0 _, coe_nnnorm], obtain ⟨ζ, hζ, hhζ : complex.abs _ < _⟩ := hex, exact ⟨ζ, hζ, by simp [hhζ]⟩ }, From a80d18edd117abbf4cf2612cfff84d133f61a906 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Jul 2022 01:31:07 +0000 Subject: [PATCH 010/173] feat(order/upper_lower): Upper closure of a set (#15581) Define the upper/lower set generated by a set. Co-authored-by: Peter Nelson --- .../set_family/intersecting.lean | 2 +- src/order/minimal.lean | 22 +- src/order/upper_lower.lean | 198 +++++++++++++++--- 3 files changed, 181 insertions(+), 41 deletions(-) diff --git a/src/combinatorics/set_family/intersecting.lean b/src/combinatorics/set_family/intersecting.lean index 5750e04b45a6d..03aad14f5b461 100644 --- a/src/combinatorics/set_family/intersecting.lean +++ b/src/combinatorics/set_family/intersecting.lean @@ -154,7 +154,7 @@ lemma intersecting.is_max_iff_card_eq (hs : (s : set α).intersecting) : (∀ t : finset α, (t : set α).intersecting → s ⊆ t → s = t) ↔ 2 * s.card = fintype.card α := begin classical, - refine ⟨λ h, _, λ h t ht hst, eq_of_subset_of_card_le hst $ + refine ⟨λ h, _, λ h t ht hst, finset.eq_of_subset_of_card_le hst $ le_of_mul_le_mul_left (ht.card_le.trans_eq h.symm) two_pos⟩, suffices : s ∪ s.map ⟨compl, compl_injective⟩ = finset.univ, { rw [fintype.card, ←this, two_mul, card_union_eq, card_map], diff --git a/src/order/minimal.lean b/src/order/minimal.lean index 60050cdc2e8a2..3fa840b4b353b 100644 --- a/src/order/minimal.lean +++ b/src/order/minimal.lean @@ -4,6 +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 /-! # Minimal/maximal elements of a set @@ -130,17 +131,12 @@ eq_singleton_iff_unique_mem.2 ⟨h.mem_minimals, λ b hb, hb.2 h.1 $ h.2 hb.1⟩ lemma is_greatest.maximals_eq (h : is_greatest s a) : maximals (≤) s = {a} := eq_singleton_iff_unique_mem.2 ⟨h.mem_maximals, λ b hb, hb.2 h.1 $ h.2 hb.1⟩ -lemma is_antichain.max_lower_set_of (hs : is_antichain (≤) s) : - maximals (≤) {x | ∃ y ∈ s, x ≤ y} = s := -begin - ext x, - simp only [maximals, exists_prop, mem_set_of_eq, forall_exists_index, and_imp, sep_set_of], - refine ⟨λ h, exists.elim h.1 (λ y hy, ((h.2 _ hy.1 rfl.le hy.2).symm.subst hy.1)), - λ h, ⟨⟨x,h,rfl.le⟩,λ b y hy hby hxy, _⟩⟩, - have : x = y := by_contra (λ h_eq, (hs h hy h_eq (hxy.trans hby)).elim), - exact hxy.antisymm (this.symm.subst hby), -end +lemma is_antichain.minimals_upper_closure (hs : is_antichain (≤) s) : + minimals (≤) (upper_closure s : set α) = s := +hs.max_minimals (λ a ⟨⟨b, hb, hba⟩, h⟩, by rwa h (subset_upper_closure hb) hba) $ λ a ha, + ⟨a, ⟨subset_upper_closure ha, λ b ⟨c, hc, hcb⟩ hba, hba.antisymm' $ + by rwa hs.eq' ha hc (hcb.trans hba)⟩, le_rfl⟩ -lemma is_antichain.min_upper_set_of (hs : is_antichain (≤) s) : - minimals (≤) {x | ∃ y ∈ s, y ≤ x} = s := -hs.to_dual.max_lower_set_of +lemma is_antichain.maximals_lower_closure (hs : is_antichain (≤) s) : + maximals (≤) (lower_closure s : set α) = s := +hs.to_dual.minimals_upper_closure diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index ae047996f1eb9..abb12e02c4aed 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Sara Rousta -/ import data.set_like.basic +import data.set.intervals.ord_connected import order.hom.complete_lattice /-! @@ -19,6 +20,8 @@ This file defines upper and lower sets in an order. of the set is in the set itself. * `upper_set`: The type of upper sets. * `lower_set`: The type of lower sets. +* `upper_closure`: The greatest upper set containing a set. +* `lower_closure`: The least lower set containing a set. * `upper_set.Ici`: Principal upper set. `set.Ici` as an upper set. * `upper_set.Ioi`: Strict principal upper set. `set.Ioi` as an upper set. * `lower_set.Iic`: Principal lower set. `set.Iic` as an lower set. @@ -133,15 +136,30 @@ alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.to_dual end has_le section preorder -variables [preorder α] (a : α) +variables [preorder α] {s : set α} (a : α) lemma is_upper_set_Ici : is_upper_set (Ici a) := λ _ _, ge_trans lemma is_lower_set_Iic : is_lower_set (Iic a) := λ _ _, le_trans lemma is_upper_set_Ioi : is_upper_set (Ioi a) := λ _ _, flip lt_of_lt_of_le lemma is_lower_set_Iio : is_lower_set (Iio a) := λ _ _, lt_of_le_of_lt +lemma is_upper_set_iff_Ici_subset : is_upper_set s ↔ ∀ ⦃a⦄, a ∈ s → Ici a ⊆ s := +by simp [is_upper_set, subset_def, @forall_swap (_ ∈ s)] + +lemma is_lower_set_iff_Iic_subset : is_lower_set s ↔ ∀ ⦃a⦄, a ∈ s → Iic a ⊆ s := +by simp [is_lower_set, subset_def, @forall_swap (_ ∈ s)] + +alias is_upper_set_iff_Ici_subset ↔ is_upper_set.Ici_subset _ +alias is_lower_set_iff_Iic_subset ↔ is_lower_set.Iic_subset _ + +lemma is_upper_set.ord_connected (h : is_upper_set s) : s.ord_connected := +⟨λ a ha b _, Icc_subset_Ici_self.trans $ h.Ici_subset ha⟩ + +lemma is_lower_set.ord_connected (h : is_lower_set s) : s.ord_connected := +⟨λ a _ b hb, Icc_subset_Iic_self.trans $ h.Iic_subset hb⟩ + section order_top -variables [order_top α] {s : set α} +variables [order_top α] lemma is_lower_set.top_mem (hs : is_lower_set s) : ⊤ ∈ s ↔ s = univ := ⟨λ h, eq_univ_of_forall $ λ a, hs le_top h, λ h, h.symm ▸ mem_univ _⟩ @@ -155,7 +173,7 @@ hs.top_mem.not.trans not_nonempty_iff_eq_empty end order_top section order_bot -variables [order_bot α] {s : set α} +variables [order_bot α] lemma is_upper_set.bot_mem (hs : is_upper_set s) : ⊥ ∈ s ↔ s = univ := ⟨λ h, eq_univ_of_forall $ λ a, hs bot_le h, λ h, h.symm ▸ mem_univ _⟩ @@ -232,18 +250,21 @@ instance : complete_distrib_lattice (upper_set α) := instance : inhabited (upper_set α) := ⟨⊥⟩ -@[simp] lemma coe_top : ((⊤ : upper_set α) : set α) = ∅ := rfl -@[simp] lemma coe_bot : ((⊥ : upper_set α) : set α) = univ := rfl -@[simp] lemma coe_sup (s t : upper_set α) : (↑(s ⊔ t) : set α) = s ∩ t := rfl -@[simp] lemma coe_inf (s t : upper_set α) : (↑(s ⊓ t) : set α) = s ∪ t := rfl -@[simp] lemma coe_Sup (S : set (upper_set α)) : (↑(Sup S) : set α) = ⋂ s ∈ S, ↑s := rfl -@[simp] lemma coe_Inf (S : set (upper_set α)) : (↑(Inf S) : set α) = ⋃ s ∈ S, ↑s := rfl -@[simp] lemma coe_supr (f : ι → upper_set α) : (↑(⨆ i, f i) : set α) = ⋂ i, f i := by simp [supr] -@[simp] lemma coe_infi (f : ι → upper_set α) : (↑(⨅ i, f i) : set α) = ⋃ i, f i := by simp [infi] -@[simp] lemma coe_supr₂ (f : Π i, κ i → upper_set α) : (↑(⨆ i j, f i j) : set α) = ⋂ i j, f i j := -by simp_rw coe_supr -@[simp] lemma coe_infi₂ (f : Π i, κ i → upper_set α) : (↑(⨅ i j, f i j) : set α) = ⋃ i j, f i j := -by simp_rw coe_infi +@[simp, norm_cast] lemma coe_subset_coe : (s : set α) ⊆ t ↔ t ≤ s := iff.rfl +@[simp, norm_cast] lemma coe_top : ((⊤ : upper_set α) : set α) = ∅ := rfl +@[simp, norm_cast] lemma coe_bot : ((⊥ : upper_set α) : set α) = univ := rfl +@[simp, norm_cast] lemma coe_sup (s t : upper_set α) : (↑(s ⊔ t) : set α) = s ∩ t := rfl +@[simp, norm_cast] lemma coe_inf (s t : upper_set α) : (↑(s ⊓ t) : set α) = s ∪ t := rfl +@[simp, norm_cast] lemma coe_Sup (S : set (upper_set α)) : (↑(Sup S) : set α) = ⋂ s ∈ S, ↑s := rfl +@[simp, norm_cast] lemma coe_Inf (S : set (upper_set α)) : (↑(Inf S) : set α) = ⋃ s ∈ S, ↑s := rfl +@[simp, norm_cast] lemma coe_supr (f : ι → upper_set α) : (↑(⨆ i, f i) : set α) = ⋂ i, f i := +by simp [supr] +@[simp, norm_cast] lemma coe_infi (f : ι → upper_set α) : (↑(⨅ i, f i) : set α) = ⋃ i, f i := +by simp [infi] +@[simp, norm_cast] lemma coe_supr₂ (f : Π i, κ i → upper_set α) : + (↑(⨆ i j, f i j) : set α) = ⋂ i j, f i j := by simp_rw coe_supr +@[simp, norm_cast] lemma coe_infi₂ (f : Π i, κ i → upper_set α) : + (↑(⨅ i j, f i j) : set α) = ⋃ i j, f i j := by simp_rw coe_infi @[simp] lemma not_mem_top : a ∉ (⊤ : upper_set α) := id @[simp] lemma mem_bot : a ∈ (⊥ : upper_set α) := trivial @@ -278,20 +299,21 @@ set_like.coe_injective.complete_distrib_lattice _ instance : inhabited (lower_set α) := ⟨⊥⟩ -@[simp] lemma coe_top : ((⊤ : lower_set α) : set α) = univ := rfl -@[simp] lemma coe_bot : ((⊥ : lower_set α) : set α) = ∅ := rfl -@[simp] lemma coe_sup (s t : lower_set α) : (↑(s ⊔ t) : set α) = s ∪ t := rfl -@[simp] lemma coe_inf (s t : lower_set α) : (↑(s ⊓ t) : set α) = s ∩ t := rfl -@[simp] lemma coe_Sup (S : set (lower_set α)) : (↑(Sup S) : set α) = ⋃ s ∈ S, ↑s := rfl -@[simp] lemma coe_Inf (S : set (lower_set α)) : (↑(Inf S) : set α) = ⋂ s ∈ S, ↑s := rfl -@[simp] lemma coe_supr (f : ι → lower_set α) : (↑(⨆ i, f i) : set α) = ⋃ i, f i := +@[simp, norm_cast] lemma coe_subset_coe : (s : set α) ⊆ t ↔ s ≤ t := iff.rfl +@[simp, norm_cast] lemma coe_top : ((⊤ : lower_set α) : set α) = univ := rfl +@[simp, norm_cast] lemma coe_bot : ((⊥ : lower_set α) : set α) = ∅ := rfl +@[simp, norm_cast] lemma coe_sup (s t : lower_set α) : (↑(s ⊔ t) : set α) = s ∪ t := rfl +@[simp, norm_cast] lemma coe_inf (s t : lower_set α) : (↑(s ⊓ t) : set α) = s ∩ t := rfl +@[simp, norm_cast] lemma coe_Sup (S : set (lower_set α)) : (↑(Sup S) : set α) = ⋃ s ∈ S, ↑s := rfl +@[simp, norm_cast] lemma coe_Inf (S : set (lower_set α)) : (↑(Inf S) : set α) = ⋂ s ∈ S, ↑s := rfl +@[simp, norm_cast] lemma coe_supr (f : ι → lower_set α) : (↑(⨆ i, f i) : set α) = ⋃ i, f i := by simp_rw [supr, coe_Sup, mem_range, Union_exists, Union_Union_eq'] -@[simp] lemma coe_infi (f : ι → lower_set α) : (↑(⨅ i, f i) : set α) = ⋂ i, f i := +@[simp, norm_cast] lemma coe_infi (f : ι → lower_set α) : (↑(⨅ i, f i) : set α) = ⋂ i, f i := by simp_rw [infi, coe_Inf, mem_range, Inter_exists, Inter_Inter_eq'] -@[simp] lemma coe_supr₂ (f : Π i, κ i → lower_set α) : (↑(⨆ i j, f i j) : set α) = ⋃ i j, f i j := -by simp_rw coe_supr -@[simp] lemma coe_infi₂ (f : Π i, κ i → lower_set α) : (↑(⨅ i j, f i j) : set α) = ⋂ i j, f i j := -by simp_rw coe_infi +@[simp, norm_cast] lemma coe_supr₂ (f : Π i, κ i → lower_set α) : + (↑(⨆ i j, f i j) : set α) = ⋃ i j, f i j := by simp_rw coe_supr +@[simp, norm_cast] lemma coe_infi₂ (f : Π i, κ i → lower_set α) : + (↑(⨅ i j, f i j) : set α) = ⋂ i j, f i j := by simp_rw coe_infi @[simp] lemma mem_top : a ∈ (⊤ : lower_set α) := trivial @[simp] lemma not_mem_bot : a ∉ (⊥ : lower_set α) := id @@ -515,3 +537,125 @@ def Iic_Inf_hom : Inf_hom α (lower_set α) := ⟨Iic, λ s, (Iic_Inf s).trans I end complete_lattice end lower_set + +section closure +variables [preorder α] {s t : set α} {x : α} + +/-- The greatest upper set containing a given set. -/ +def upper_closure (s : set α) : upper_set α := +⟨{x | ∃ a ∈ s, a ≤ x}, λ x y h, Exists₂.imp $ λ a _, h.trans'⟩ + +/-- The least lower set containing a given set. -/ +def lower_closure (s : set α) : lower_set α := +⟨{x | ∃ a ∈ s, x ≤ a}, λ x y h, Exists₂.imp $ λ a _, h.trans⟩ + +@[simp, norm_cast] lemma coe_upper_closure (s : set α) : + ↑(upper_closure s) = {x | ∃ a ∈ s, a ≤ x} := rfl + +@[simp, norm_cast] lemma coe_lower_closure (s : set α) : + ↑(lower_closure s) = {x | ∃ a ∈ s, x ≤ a} := rfl + +@[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 + +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⟩ + +lemma upper_closure_min (h : s ⊆ t) (ht : is_upper_set t) : ↑(upper_closure s) ⊆ t := +λ a ⟨b, hb, hba⟩, ht hba $ h hb + +lemma lower_closure_min (h : s ⊆ t) (ht : is_lower_set t) : ↑(lower_closure s) ⊆ t := +λ a ⟨b, hb, hab⟩, ht hab $ h hb + +@[simp] lemma upper_set.infi_Ici (s : set α) : (⨅ a ∈ s, upper_set.Ici a) = upper_closure s := +by { ext, simp } + +@[simp] lemma lower_set.supr_Iic (s : set α) : (⨆ a ∈ s, lower_set.Iic a) = lower_closure s := +by { ext, simp } + +lemma gc_upper_closure_coe : + galois_connection (to_dual ∘ upper_closure : set α → (upper_set α)ᵒᵈ) (coe ∘ of_dual) := +λ s t, ⟨λ h, subset_upper_closure.trans $ upper_set.coe_subset_coe.2 h, + λ h, upper_closure_min h t.upper⟩ + +lemma gc_lower_closure_coe : galois_connection (lower_closure : set α → lower_set α) coe := +λ s t, ⟨λ h, subset_lower_closure.trans $ lower_set.coe_subset_coe.2 h, + λ h, lower_closure_min h t.lower⟩ + +/-- `upper_closure` forms a reversed Galois insertion with the coercion from upper sets to sets. -/ +def gi_upper_closure_coe : + galois_insertion (to_dual ∘ upper_closure : set α → (upper_set α)ᵒᵈ) (coe ∘ of_dual) := +{ choice := λ s hs, to_dual (⟨s, λ a b hab ha, hs ⟨a, ha, hab⟩⟩ : upper_set α), + gc := gc_upper_closure_coe, + le_l_u := λ _, subset_upper_closure, + choice_eq := λ s hs, + of_dual.injective $ set_like.coe_injective $ subset_upper_closure.antisymm hs } + +/-- `lower_closure` forms a Galois insertion with the coercion from lower sets to sets. -/ +def gi_lower_closure_coe : galois_insertion (lower_closure : set α → lower_set α) coe := +{ choice := λ s hs, ⟨s, λ a b hba ha, hs ⟨a, ha, hba⟩⟩, + gc := gc_lower_closure_coe, + le_l_u := λ _, subset_lower_closure, + choice_eq := λ s hs, set_like.coe_injective $ subset_lower_closure.antisymm hs } + +lemma upper_closure_anti : antitone (upper_closure : set α → upper_set α) := +gc_upper_closure_coe.monotone_l + +lemma lower_closure_mono : monotone (lower_closure : set α → lower_set α) := +gc_lower_closure_coe.monotone_l + +@[simp] lemma upper_closure_empty : upper_closure (∅ : set α) = ⊤ := by { ext, simp } +@[simp] lemma lower_closure_empty : lower_closure (∅ : set α) = ⊥ := by { ext, simp } + +@[simp] lemma upper_closure_univ : upper_closure (univ : set α) = ⊥ := +le_bot_iff.1 subset_upper_closure + +@[simp] lemma lower_closure_univ : lower_closure (univ : set α) = ⊤ := +top_le_iff.1 subset_lower_closure + +@[simp] lemma upper_closure_eq_top_iff : upper_closure s = ⊤ ↔ s = ∅ := +⟨λ h, subset_empty_iff.1 $ subset_upper_closure.trans (congr_arg coe h).subset, + by { rintro rfl, exact upper_closure_empty }⟩ + +@[simp] lemma lower_closure_eq_bot_iff : lower_closure s = ⊥ ↔ s = ∅ := +⟨λ h, subset_empty_iff.1 $ subset_lower_closure.trans (congr_arg coe h).subset, + by { rintro rfl, exact lower_closure_empty }⟩ + +@[simp] lemma upper_closure_union (s t : set α) : + upper_closure (s ∪ t) = upper_closure s ⊓ upper_closure t := +by { ext, simp [or_and_distrib_right, exists_or_distrib] } + +@[simp] lemma lower_closure_union (s t : set α) : + lower_closure (s ∪ t) = lower_closure s ⊔ lower_closure t := +by { ext, simp [or_and_distrib_right, exists_or_distrib] } + +@[simp] lemma upper_closure_Union (f : ι → set α) : + upper_closure (⋃ i, f i) = ⨅ i, upper_closure (f i) := +by { ext, simp [←exists_and_distrib_right, @exists_comm α] } + +@[simp] lemma lower_closure_Union (f : ι → set α) : + lower_closure (⋃ i, f i) = ⨆ i, lower_closure (f i) := +by { ext, simp [←exists_and_distrib_right, @exists_comm α] } + +@[simp] lemma upper_closure_sUnion (S : set (set α)) : + upper_closure (⋃₀ S) = ⨅ s ∈ S, upper_closure s := +by simp_rw [sUnion_eq_bUnion, upper_closure_Union] + +@[simp] lemma lower_closure_sUnion (S : set (set α)) : + lower_closure (⋃₀ S) = ⨆ s ∈ S, lower_closure s := +by simp_rw [sUnion_eq_bUnion, lower_closure_Union] + +lemma set.ord_connected.upper_closure_inter_lower_closure (h : s.ord_connected) : + ↑(upper_closure s) ∩ ↑(lower_closure s) = s := +(subset_inter subset_upper_closure subset_lower_closure).antisymm' $ λ a ⟨⟨b, hb, hba⟩, c, hc, hac⟩, + h.out hb hc ⟨hba, hac⟩ + +lemma ord_connected_iff_upper_closure_inter_lower_closure : + s.ord_connected ↔ ↑(upper_closure s) ∩ ↑(lower_closure s) = s := +begin + refine ⟨set.ord_connected.upper_closure_inter_lower_closure, λ h, _⟩, + rw ←h, + exact (upper_set.upper _).ord_connected.inter (lower_set.lower _).ord_connected, +end + +end closure From 8e630047b564899b5d8a0641c081ee177ad92031 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 23 Jul 2022 03:35:52 +0000 Subject: [PATCH 011/173] feat(set_theory/game/nim): make the file `noncomputable theory` (#15367) Since we're interfacing with ordinals and since `pgame` holds no data, we simplify `nim` and allow it to be noncomputable. We need to give `nim` the `noncomputable!` attribute to avoid a VM compilation bug. --- src/set_theory/game/nim.lean | 23 ++++++++++++----------- 1 file changed, 12 insertions(+), 11 deletions(-) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index a04a968ea5b8d..58935edcf070d 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -28,18 +28,19 @@ instead use `o.out.α` for the possible moves. You can use `to_left_moves_nim` a vice versa. -/ -universes u +noncomputable theory -/-- `ordinal.out'` has the sole purpose of making `nim` computable. It performs the same job as - `quotient.out` but is specific to ordinals. -/ -def ordinal.out' (o : ordinal) : Well_order := -⟨o.out.α, (<), o.out.wo⟩ +universe u /-- The definition of single-heap nim, which can be viewed as a pile of stones where each player can take a positive number of stones from it on their turn. -/ -def nim : ordinal → pgame -| o₁ := let f := λ o₂, have hwf : ordinal.typein o₁.out'.r o₂ < o₁ := ordinal.typein_lt_self o₂, - nim (ordinal.typein o₁.out'.r o₂) in ⟨o₁.out'.α, o₁.out'.α, f, f⟩ +-- Uses `noncomputable!` to avoid `rec_fn_macro only allowed in meta definitions` VM error +noncomputable! def nim : ordinal.{u} → pgame.{u} +| o₁ := + let f := λ o₂, + have ordinal.typein o₁.out.r o₂ < o₁ := ordinal.typein_lt_self o₂, + nim (ordinal.typein o₁.out.r o₂) + in ⟨o₁.out.α, o₁.out.α, f, f⟩ using_well_founded { dec_tac := tactic.assumption } namespace pgame @@ -106,10 +107,10 @@ by { rw nim_def, exact ordinal.is_empty_out_zero } instance : is_empty (nim 0).right_moves := by { rw nim_def, exact ordinal.is_empty_out_zero } -noncomputable instance : unique (nim 1).left_moves := +instance : unique (nim 1).left_moves := by { rw nim_def, exact ordinal.unique_out_one } -noncomputable instance : unique (nim 1).right_moves := +instance : unique (nim 1).right_moves := by { rw nim_def, exact ordinal.unique_out_one } /-- `nim 0` has exactly the same moves as `0`. -/ @@ -118,7 +119,7 @@ def nim_zero_relabelling : nim 0 ≡r 0 := relabelling.is_empty _ theorem nim_zero_equiv : nim 0 ≈ 0 := equiv.is_empty _ /-- `nim 1` has exactly the same moves as `star`. -/ -noncomputable def nim_one_relabelling : nim 1 ≡r star := +def nim_one_relabelling : nim 1 ≡r star := begin rw nim_def, refine ⟨_, _, λ i, _, λ j, _⟩, From 5a919a9ca259bcf2942bb98ddd8ada5b71274548 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 23 Jul 2022 07:51:52 +0000 Subject: [PATCH 012/173] feat(ring_theory/noetherian): Finitely generated idempotent ideal is principal. (#15561) --- src/algebra/ring/idempotents.lean | 5 +- src/ring_theory/noetherian.lean | 108 +++++++++++++++++++++--------- 2 files changed, 80 insertions(+), 33 deletions(-) diff --git a/src/algebra/ring/idempotents.lean b/src/algebra/ring/idempotents.lean index 683809ccecede..d6d32cfbe379f 100644 --- a/src/algebra/ring/idempotents.lean +++ b/src/algebra/ring/idempotents.lean @@ -27,7 +27,7 @@ projection, idempotent variables {M N S M₀ M₁ R G G₀ : Type*} variables [has_mul M] [monoid N] [semigroup S] [mul_zero_class M₀] [mul_one_class M₁] - [non_assoc_ring R] [group G] [group_with_zero G₀] + [non_assoc_ring R] [group G] [cancel_monoid_with_zero G₀] /-- An element `p` is said to be idempotent if `p * p = p` @@ -70,8 +70,7 @@ begin refine iff.intro (λ h, or_iff_not_imp_left.mpr (λ hp, _)) (λ h, h.elim (λ hp, hp.symm ▸ zero) (λ hp, hp.symm ▸ one)), - lift p to G₀ˣ using is_unit.mk0 _ hp, - exact_mod_cast iff_eq_one.mp (by exact_mod_cast h.eq), + exact mul_left_cancel₀ hp (h.trans (mul_one p).symm) end /-! ### Instances on `subtype is_idempotent_elem` -/ diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index a48ef2c51b355..25594cff21040 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -10,6 +10,7 @@ import order.order_iso_nat import ring_theory.ideal.operations import order.compactly_generated import linear_algebra.linear_independent +import algebra.ring.idempotents /-! # Noetherian rings and modules @@ -135,6 +136,15 @@ begin exact add_mem (smul_mem _ _ hci) (smul_mem _ _ hz) } end +theorem exists_mem_and_smul_eq_self_of_fg_of_le_smul {R : Type*} [comm_ring R] + {M : Type*} [add_comm_group M] [module R M] + (I : ideal R) (N : submodule R M) (hn : N.fg) (hin : N ≤ I • N) : + ∃ r ∈ I, ∀ n ∈ N, r • n = n := +begin + obtain ⟨r, hr, hr'⟩ := N.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul I hn hin, + exact ⟨-(r-1), I.neg_mem hr, λ n hn, by simpa [sub_smul] using hr' n hn⟩, +end + theorem fg_bot : (⊥ : submodule R M).fg := ⟨∅, by rw [finset.coe_empty, span_empty]⟩ @@ -282,23 +292,6 @@ begin { rw [finset.coe_insert, submodule.span_insert], apply h₂; apply_assumption } end -/-- An ideal of `R` is finitely generated if it is the span of a finite subset of `R`. - -This is defeq to `submodule.fg`, but unfolds more nicely. -/ -def _root_.ideal.fg (I : ideal R) : Prop := ∃ S : finset R, ideal.span ↑S = I - -/-- The image of a finitely generated ideal is finitely generated. - -This is the `ideal` version of `submodule.fg.map`. -/ -lemma _root_.ideal.fg.map {R S : Type*} [semiring R] [semiring S] {I : ideal R} (h : I.fg) - (f : R →+* S) : (I.map f).fg := -begin - classical, - obtain ⟨s, hs⟩ := h, - refine ⟨s.image f, _⟩, - rw [finset.coe_image, ←ideal.map_span, hs], -end - /-- The kernel of the composition of two linear maps is finitely generated if both kernels are and the first morphism is surjective. -/ lemma fg_ker_comp {R M N P : Type*} [ring R] [add_comm_group M] [module R M] @@ -321,19 +314,6 @@ begin exact submodule.span_eq_restrict_scalars R S M X h end -lemma _root_.ideal.fg_ker_comp {R S A : Type*} [comm_ring R] [comm_ring S] [comm_ring A] - (f : R →+* S) (g : S →+* A) (hf : f.ker.fg) (hg : g.ker.fg) (hsur : function.surjective f) : - (g.comp f).ker.fg := -begin - letI : algebra R S := ring_hom.to_algebra f, - letI : algebra R A := ring_hom.to_algebra (g.comp f), - letI : algebra S A := ring_hom.to_algebra g, - letI : is_scalar_tower R S A := is_scalar_tower.of_algebra_map_eq (λ _, rfl), - let f₁ := algebra.linear_map R S, - let g₁ := (is_scalar_tower.to_alg_hom R S A).to_linear_map, - exact fg_ker_comp f₁ g₁ hf (fg_restrict_scalars g.ker hg hsur) hsur -end - /-- Finitely generated submodules are precisely compact elements in the submodule lattice. -/ theorem fg_iff_compact (s : submodule R M) : s.fg ↔ complete_lattice.is_compact_element s := begin @@ -364,6 +344,74 @@ end end submodule +namespace ideal + +variables {R : Type*} {M : Type*} [semiring R] [add_comm_monoid M] [module R M] + +/-- An ideal of `R` is finitely generated if it is the span of a finite subset of `R`. + +This is defeq to `submodule.fg`, but unfolds more nicely. -/ +def fg (I : ideal R) : Prop := ∃ S : finset R, ideal.span ↑S = I + +/-- The image of a finitely generated ideal is finitely generated. + +This is the `ideal` version of `submodule.fg.map`. -/ +lemma fg.map {R S : Type*} [semiring R] [semiring S] {I : ideal R} (h : I.fg) + (f : R →+* S) : (I.map f).fg := +begin + classical, + obtain ⟨s, hs⟩ := h, + refine ⟨s.image f, _⟩, + rw [finset.coe_image, ←ideal.map_span, hs], +end + +lemma fg_ker_comp {R S A : Type*} [comm_ring R] [comm_ring S] [comm_ring A] + (f : R →+* S) (g : S →+* A) (hf : f.ker.fg) (hg : g.ker.fg) (hsur : function.surjective f) : + (g.comp f).ker.fg := +begin + letI : algebra R S := ring_hom.to_algebra f, + letI : algebra R A := ring_hom.to_algebra (g.comp f), + letI : algebra S A := ring_hom.to_algebra g, + letI : is_scalar_tower R S A := is_scalar_tower.of_algebra_map_eq (λ _, rfl), + let f₁ := algebra.linear_map R S, + let g₁ := (is_scalar_tower.to_alg_hom R S A).to_linear_map, + exact submodule.fg_ker_comp f₁ g₁ hf (submodule.fg_restrict_scalars g.ker hg hsur) hsur +end + +/-- A finitely generated idempotent ideal is generated by an idempotent element -/ +lemma is_idempotent_elem_iff_of_fg {R : Type*} [comm_ring R] (I : ideal R) + (h : I.fg) : + is_idempotent_elem I ↔ ∃ e : R, is_idempotent_elem e ∧ I = R ∙ e := +begin + split, + { intro e, + obtain ⟨r, hr, hr'⟩ := submodule.exists_mem_and_smul_eq_self_of_fg_of_le_smul I I h + (by { rw [smul_eq_mul], exact e.ge }), + simp_rw smul_eq_mul at hr', + refine ⟨r, hr' r hr, antisymm _ ((submodule.span_singleton_le_iff_mem _ _).mpr hr)⟩, + intros x hx, + rw ← hr' x hx, + exact ideal.mem_span_singleton'.mpr ⟨_, mul_comm _ _⟩ }, + { rintros ⟨e, he, rfl⟩, + simp [is_idempotent_elem, ideal.span_singleton_mul_span_singleton, he.eq] } +end + +lemma is_idempotent_elem_iff_eq_bot_or_top {R : Type*} [comm_ring R] [is_domain R] + (I : ideal R) (h : I.fg) : + is_idempotent_elem I ↔ I = ⊥ ∨ I = ⊤ := +begin + split, + { intro H, + obtain ⟨e, he, rfl⟩ := (I.is_idempotent_elem_iff_of_fg h).mp H, + simp only [ideal.submodule_span_eq, ideal.span_singleton_eq_bot], + apply or_of_or_of_imp_of_imp (is_idempotent_elem.iff_eq_zero_or_one.mp he) id, + rintro rfl, + simp }, + { rintro (rfl|rfl); simp [is_idempotent_elem] } +end + +end ideal + /-- `is_noetherian R M` is the proposition that `M` is a Noetherian `R`-module, implemented as the predicate that all `R`-submodules of `M` are finitely generated. From 8f02755bf65a7fe2e54e0710e29e04f617182857 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Jul 2022 13:51:43 +0000 Subject: [PATCH 013/173] feat(combinatorics/simple_graph/density): Bound on the difference between edge densities (#15353) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Auxiliary lemma for Szemerédi Regularity Lemma. Co-authored-by: Bhavik Mehta --- src/combinatorics/simple_graph/density.lean | 103 +++++++++++++++++++- 1 file changed, 102 insertions(+), 1 deletion(-) diff --git a/src/combinatorics/simple_graph/density.lean b/src/combinatorics/simple_graph/density.lean index 16999a6634677..5931847b14f16 100644 --- a/src/combinatorics/simple_graph/density.lean +++ b/src/combinatorics/simple_graph/density.lean @@ -21,6 +21,7 @@ Between two finsets of vertices, -/ open finset +open_locale big_operators variables {ι κ α β : Type*} @@ -29,7 +30,7 @@ variables {ι κ α β : Type*} namespace rel section asymmetric variables (r : α → β → Prop) [Π a, decidable_pred (r a)] {s s₁ s₂ : finset α} {t t₁ t₂ : finset β} - {a : α} {b : β} + {a : α} {b : β} {δ : ℚ} /-- Finset of edges of a relation between two finsets of vertices. -/ def interedges (s : finset α) (t : finset β) : finset (α × β) := @@ -122,6 +123,106 @@ by rw [edge_density, finset.card_empty, nat.cast_zero, zero_mul, div_zero] @[simp] lemma edge_density_empty_right (s : finset α) : edge_density r s ∅ = 0 := by rw [edge_density, finset.card_empty, nat.cast_zero, mul_zero, div_zero] +lemma card_interedges_finpartition_left [decidable_eq α] (P : finpartition s) (t : finset β) : + (interedges r s t).card = ∑ a in P.parts, (interedges r a t).card := +begin + classical, + simp_rw [←P.bUnion_parts, interedges_bUnion_left, id.def], + rw card_bUnion, + exact λ x hx y hy h, interedges_disjoint_left r (P.disjoint hx hy h) _, +end + +lemma card_interedges_finpartition_right [decidable_eq β] (s : finset α) (P : finpartition t) : + (interedges r s t).card = ∑ b in P.parts, (interedges r s b).card := +begin + classical, + simp_rw [←P.bUnion_parts, interedges_bUnion_right, id], + rw card_bUnion, + exact λ x hx y hy h, interedges_disjoint_right r _ (P.disjoint hx hy h), +end + +lemma card_interedges_finpartition [decidable_eq α] [decidable_eq β] (P : finpartition s) + (Q : finpartition t) : + (interedges r s t).card = ∑ ab in P.parts.product Q.parts, (interedges r ab.1 ab.2).card := +by simp_rw [card_interedges_finpartition_left _ P, card_interedges_finpartition_right _ _ Q, + sum_product] + +lemma mul_edge_density_le_edge_density (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hs₂ : s₂.nonempty) + (ht₂ : t₂.nonempty) : + (s₂.card : ℚ)/s₁.card * (t₂.card/t₁.card) * edge_density r s₂ t₂ ≤ edge_density r s₁ t₁ := +begin + have hst : (s₂.card : ℚ) * t₂.card ≠ 0 := by simp [hs₂.ne_empty, ht₂.ne_empty], + rw [edge_density, edge_density, div_mul_div_comm, mul_comm, div_mul_div_cancel _ hst], + refine div_le_div_of_le (by exact_mod_cast (s₁.card * t₁.card).zero_le) _, + exact_mod_cast card_le_of_subset (interedges_mono hs ht), +end + +lemma edge_density_sub_edge_density_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hs₂ : s₂.nonempty) + (ht₂ : t₂.nonempty) : + edge_density r s₂ t₂ - edge_density r s₁ t₁ ≤ 1 - (s₂.card)/s₁.card * (t₂.card/t₁.card) := +begin + refine (sub_le_sub_left (mul_edge_density_le_edge_density r hs ht hs₂ ht₂) _).trans _, + refine le_trans _ (mul_le_of_le_one_right _ (edge_density_le_one r s₂ t₂)), + { rw [sub_mul, one_mul] }, + refine sub_nonneg_of_le (mul_le_one _ (div_nonneg (nat.cast_nonneg _) (nat.cast_nonneg _)) _); + exact div_le_one_of_le (nat.cast_le.2 (card_le_of_subset ‹_›)) (nat.cast_nonneg _), +end + +lemma abs_edge_density_sub_edge_density_le_one_sub_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) + (hs₂ : s₂.nonempty) (ht₂ : t₂.nonempty) : + |edge_density r s₂ t₂ - edge_density r s₁ t₁| ≤ 1 - s₂.card/s₁.card * (t₂.card/t₁.card) := +begin + have habs : abs (edge_density r s₂ t₂ - edge_density r s₁ t₁) ≤ 1, + { rw [abs_sub_le_iff, ←sub_zero (1 : ℚ)], + split; exact sub_le_sub (edge_density_le_one r _ _) (edge_density_nonneg r _ _) }, + refine abs_sub_le_iff.2 ⟨edge_density_sub_edge_density_le_one_sub_mul r hs ht hs₂ ht₂, _⟩, + rw [←add_sub_cancel (edge_density r s₁ t₁) (edge_density (λ x y, ¬r x y) s₁ t₁), + ←add_sub_cancel (edge_density r s₂ t₂) (edge_density (λ x y, ¬r x y) s₂ t₂), + edge_density_add_edge_density_compl _ (hs₂.mono hs) (ht₂.mono ht), + edge_density_add_edge_density_compl _ hs₂ ht₂, sub_sub_sub_cancel_left], + exact edge_density_sub_edge_density_le_one_sub_mul _ hs ht hs₂ ht₂, +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 := +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δ' }, + 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δ' }, + 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 }, + { exact div_nonneg (nat.cast_nonneg _) (nat.cast_nonneg _) } +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 * δ := +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 + ((sub_le_self_iff _).2 $ sq_nonneg δ) }, + rw two_mul, + refine (abs_sub _ _).trans (add_le_add (le_trans _ h) (le_trans _ h)); + { rw abs_of_nonneg, + exact_mod_cast edge_density_le_one r _ _, + exact_mod_cast edge_density_nonneg r _ _ } +end + end asymmetric section symmetric From d22b20c69cf3fc5827dc2aedb37b74ea9d905fc7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Jul 2022 14:34:52 +0000 Subject: [PATCH 014/173] feat(order/partition/finpartition): Bound size of a bit of `finpartition.atomise` (#15350) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Auxiliary lemmas for Szemerédi Regularity Lemma. Co-authored-by: Bhavik Mehta --- src/order/partition/finpartition.lean | 28 ++++++++++++++++++++------- 1 file changed, 21 insertions(+), 7 deletions(-) diff --git a/src/order/partition/finpartition.lean b/src/order/partition/finpartition.lean index e2446127518e8..af8bf00615867 100644 --- a/src/order/partition/finpartition.lean +++ b/src/order/partition/finpartition.lean @@ -475,7 +475,7 @@ of_erase variables {F : finset (finset α)} -lemma mem_atomise {t : finset α} : +lemma mem_atomise : t ∈ (atomise s F).parts ↔ t.nonempty ∧ ∃ (Q ⊆ F), s.filter (λ i, ∀ u ∈ F, u ∈ Q ↔ i ∈ u) = t := by simp only [atomise, of_erase, bot_eq_empty, mem_erase, mem_image, nonempty_iff_ne_empty, mem_singleton, and_comm, mem_powerset, exists_prop] @@ -490,17 +490,31 @@ end lemma card_atomise_le : (atomise s F).parts.card ≤ 2^F.card := (card_le_of_subset $ erase_subset _ _).trans $ finset.card_image_le.trans (card_powerset _).le -lemma bUnion_filter_atomise (t : finset α) (ht : t ∈ F) (hts : t ⊆ s) : - ((atomise s F).parts.filter $ λ u, u ⊆ t).bUnion id = t := +lemma bUnion_filter_atomise (ht : t ∈ F) (hts : t ⊆ s) : + ((atomise s F).parts.filter $ λ u, u ⊆ t ∧ u.nonempty).bUnion id = t := begin ext a, - rw mem_bUnion, - refine ⟨λ ⟨u, hu, ha⟩, (mem_filter.1 hu).2 ha, λ ha, _⟩, + refine mem_bUnion.trans ⟨λ ⟨u, hu, ha⟩, (mem_filter.1 hu).2.1 ha, λ ha, _⟩, obtain ⟨u, hu, hau⟩ := (atomise s F).exists_mem (hts ha), - refine ⟨u, mem_filter.2 ⟨hu, λ b hb, _⟩, hau⟩, + refine ⟨u, mem_filter.2 ⟨hu, λ b hb, _, _, hau⟩, hau⟩, obtain ⟨Q, hQ, rfl⟩ := (mem_atomise.1 hu).2, rw mem_filter at hau hb, - rwa [←hb.2 _ ht, hau.2 _ ht] + rwa [←hb.2 _ ht, hau.2 _ ht], +end + +lemma card_filter_atomise_le_two_pow (ht : t ∈ F) : + ((atomise s F).parts.filter $ λ u, u ⊆ t ∧ u.nonempty).card ≤ 2 ^ (F.card - 1) := +begin + suffices h : (atomise s F).parts.filter (λ u, u ⊆ t ∧ u.nonempty) + ⊆ (F.erase t).powerset.image (λ P, s.filter $ λ i, ∀ x ∈ F, x ∈ insert t P ↔ i ∈ x), + { refine (card_le_of_subset h).trans (card_image_le.trans _), + rw [card_powerset, card_erase_of_mem ht] }, + rw subset_iff, + simp only [mem_erase, mem_sdiff, mem_powerset, mem_image, exists_prop, mem_filter, and_assoc, + finset.nonempty, exists_imp_distrib, and_imp, mem_atomise, forall_apply_eq_imp_iff₂], + rintro P' i hi P PQ rfl hy₂ j hj, + refine ⟨P.erase t, erase_subset_erase _ PQ, _⟩, + simp only [insert_erase (((mem_filter.1 hi).2 _ ht).2 $ hy₂ hi), filter_congr_decidable], end end atomise From 5dee850c2a60750075a00ebb21ed52a148fc9ff3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Jul 2022 14:34:53 +0000 Subject: [PATCH 015/173] feat(data/finset/basic): There is exactly one set in empty types (#15607) and transfer a few boolean algebra lemmas to `finset`. --- src/data/finset/basic.lean | 10 ++++++++++ src/data/fintype/basic.lean | 2 ++ src/data/set/basic.lean | 2 ++ 3 files changed, 14 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 9f0547f0454f1..3c9773f8ccc69 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -494,6 +494,13 @@ 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 +instance [nonempty α] : nontrivial (finset α) := +‹nonempty α›.elim $ λ a, ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ + +instance [is_empty α] : unique (finset α) := +{ default := ∅, + uniq := λ s, eq_empty_of_forall_not_mem is_empty_elim } + /-! ### cons -/ section cons @@ -1234,6 +1241,9 @@ set.ext $ λ _, mem_sdiff @[simp] theorem sdiff_union_self_eq_union : (s \ t) ∪ t = s ∪ t := sup_sdiff_self_left +lemma union_sdiff_left (s t : finset α) : (s ∪ t) \ s = t \ s := sup_sdiff_left_self +lemma union_sdiff_right (s t : finset α) : (s ∪ t) \ t = s \ t := sup_sdiff_right_self + lemma union_sdiff_symm : s ∪ (t \ s) = t ∪ (s \ t) := sup_sdiff_symm lemma sdiff_union_inter (s t : finset α) : (s \ t) ∪ (s ∩ t) = s := sup_sdiff_inf _ _ diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 8c92ba37cabd3..70db1e66434e9 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -127,6 +127,8 @@ variables [decidable_eq α] {a : α} instance : boolean_algebra (finset α) := generalized_boolean_algebra.to_boolean_algebra +lemma sdiff_eq_inter_compl (s t : finset α) : s \ t = s ∩ tᶜ := sdiff_eq + lemma compl_eq_univ_sdiff (s : finset α) : sᶜ = univ \ s := rfl @[simp] lemma mem_compl : a ∈ sᶜ ↔ a ∉ s := by simp [compl_eq_univ_sdiff] diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index b965f547a1c6f..5eb4f220426b2 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -451,6 +451,8 @@ by simp [subset_def] lemma univ_unique [unique α] : @set.univ α = {default} := set.ext $ λ x, iff_of_true trivial $ subsingleton.elim x default +instance [nonempty α] : nontrivial (set α) := ⟨⟨∅, univ, empty_ne_univ⟩⟩ + /-! ### Lemmas about union -/ theorem union_def {s₁ s₂ : set α} : s₁ ∪ s₂ = {a | a ∈ s₁ ∨ a ∈ s₂} := rfl From d0bf3e4d80c6e37c84b31364d5ad9c39f7aeeef2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 23 Jul 2022 14:34:54 +0000 Subject: [PATCH 016/173] chore(algebra/lie/direct_sum): remove `direct_sum.lie_algebra_is_internal` (#15631) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This meant the same thing as the unprefixed version, and wasn't used anywhere: ```lean example [decidable_eq ι] : direct_sum.lie_algebra_is_internal I ↔ direct_sum.is_internal I := iff.rfl ``` I think it was added before `direct_sum.is_internal` generalized to arbitrary additive subobjects. --- src/algebra/lie/direct_sum.lean | 10 +--------- 1 file changed, 1 insertion(+), 9 deletions(-) diff --git a/src/algebra/lie/direct_sum.lean b/src/algebra/lie/direct_sum.lean index 257b61b88aab9..1b8a6dd8e8ed7 100644 --- a/src/algebra/lie/direct_sum.lean +++ b/src/algebra/lie/direct_sum.lean @@ -11,7 +11,7 @@ import algebra.lie.basic /-! # Direct sums of Lie algebras and Lie modules -Direct sums of Lie algebras and Lie modules carry natural algbebra and module structures. +Direct sums of Lie algebras and Lie modules carry natural algebra and module structures. ## Tags @@ -194,14 +194,6 @@ section ideals variables {L : Type w} [lie_ring L] [lie_algebra R L] (I : ι → lie_ideal R L) -/-- Given a Lie algebra `L` and a family of ideals `I i ⊆ L`, informally this definition is the -statement that `L = ⨁ i, I i`. - -More formally, the inclusions give a natural map from the (external) direct sum to the enclosing Lie -algebra: `(⨁ i, I i) → L`, and this definition is the proposition that this map is bijective. -/ -def lie_algebra_is_internal [decidable_eq ι] : Prop := -function.bijective $ to_module R ι L $ λ i, ((I i).incl : I i →ₗ[R] L) - /-- The fact that this instance is necessary seems to be a bug in typeclass inference. See [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/ Typeclass.20resolution.20under.20binders/near/245151099). -/ From af2cbcd8669f8fdcc7c734c3dee90f39d779e01a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 23 Jul 2022 17:07:52 +0000 Subject: [PATCH 017/173] refactor(data/set/countable): golf using `_root_.countable`, review API (#15624) * add `set.countable_coe_iff`, `set.countable.to_subtype`, `set.to_countable`, `set.countable_univ`, `countable.to_set`, `set.countable.exists_surjective`; * generalize some lemmas from `[encodable]` to `[countable]`; * move `section enumerate` up to use it in `countable_iff_exists_subset_range`; * golf some proofs using facts about `_root_.countable`; * drop `countable_encodable` and `countable_encodable'`, use `set.to_countable` instead. --- src/data/set/countable.lean | 149 ++++++++---------- src/dynamics/ergodic/conservative.lean | 6 +- .../constructions/borel_space.lean | 8 +- .../decomposition/unsigned_hahn.lean | 2 +- .../integral/integral_eq_improper.lean | 4 +- src/measure_theory/integral/lebesgue.lean | 4 +- src/measure_theory/measurable_space.lean | 4 +- src/measure_theory/measure/ae_disjoint.lean | 4 +- src/measure_theory/measure/measure_space.lean | 2 +- src/model_theory/fraisse.lean | 2 +- src/number_theory/liouville/measure.lean | 2 +- src/order/filter/archimedean.lean | 4 +- src/order/filter/at_top_bot.lean | 4 +- src/order/filter/bases.lean | 6 +- src/probability/stopping.lean | 14 +- src/topology/G_delta.lean | 4 +- src/topology/bases.lean | 2 +- 17 files changed, 105 insertions(+), 116 deletions(-) diff --git a/src/data/set/countable.lean b/src/data/set/countable.lean index eeeb3a6553d24..fd84f0ff87cc4 100644 --- a/src/data/set/countable.lean +++ b/src/data/set/countable.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import data.set.finite +import data.countable.basic import logic.equiv.list /-! @@ -15,8 +16,8 @@ open function set encodable open classical (hiding some) open_locale classical -universes u v w -variables {α : Type u} {β : Type v} {γ : Type w} +universes u v w x +variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} namespace set @@ -27,10 +28,17 @@ constructive analogue of countability. (For the most part, theorems about -/ protected def countable (s : set α) : Prop := nonempty (encodable s) +@[simp] lemma countable_coe_iff {s : set α} : countable s ↔ s.countable := nonempty_encodable.symm + +/-- Prove `set.countable` from a `countable` instance on the subtype. -/ +lemma to_countable (s : set α) [countable s] : s.countable := countable_coe_iff.mp ‹_› + +/-- Restate `set.countable` as a `countable` instance. -/ +alias countable_coe_iff ↔ _root_.countable.to_set countable.to_subtype + protected lemma countable_iff_exists_injective {s : set α} : - s.countable ↔ ∃f:s → ℕ, injective f := -⟨λ ⟨h⟩, by exactI ⟨encode, encode_injective⟩, - λ ⟨f, h⟩, ⟨⟨f, partial_inv f, partial_inv_left h⟩⟩⟩ + s.countable ↔ ∃ f : s → ℕ, injective f := +countable_coe_iff.symm.trans (countable_iff_exists_injective s) /-- A set `s : set α` is countable if and only if there exists a function `α → ℕ` injective on `s`. -/ @@ -42,17 +50,39 @@ set.countable_iff_exists_injective.trans hf $ by simpa [as, bs] using h⟩, λ ⟨f, hf⟩, ⟨_, inj_on_iff_injective.1 hf⟩⟩ -lemma countable_iff_exists_subset_range [ne : nonempty α] {s : set α} : - s.countable ↔ ∃f:ℕ → α, s ⊆ range f := -⟨λ ⟨h⟩, by inhabit α; exactI ⟨λ n, ((decode s n).map subtype.val).iget, - λ a as, ⟨encode (⟨a, as⟩ : s), by simp [encodek]⟩⟩, - λ ⟨f, hf⟩, ⟨⟨ - λ x, inv_fun f x.1, - λ n, if h : f n ∈ s then some ⟨f n, h⟩ else none, - λ ⟨x, hx⟩, begin - have := inv_fun_eq (hf hx), dsimp at this ⊢, - simp [this, hx] - end⟩⟩⟩ +/-- Convert `set.countable s` to `encodable s` (noncomputable). -/ +protected def countable.to_encodable {s : set α} : s.countable → encodable s := +classical.choice + +section enumerate + +/-- Noncomputably enumerate elements in a set. The `default` value is used to extend the domain to +all of `ℕ`. -/ + +def enumerate_countable {s : set α} (h : s.countable) (default : α) : ℕ → α := +assume n, match @encodable.decode s h.to_encodable n with + | (some y) := y + | (none) := default + end + +lemma subset_range_enumerate {s : set α} (h : s.countable) (default : α) : + s ⊆ range (enumerate_countable h default) := +assume x hx, +⟨@encodable.encode s h.to_encodable ⟨x, hx⟩, +by simp [enumerate_countable, encodable.encodek]⟩ + +end enumerate + +lemma countable.mono {s₁ s₂ : set α} (h : s₁ ⊆ s₂) : s₂.countable → s₁.countable +| ⟨H⟩ := ⟨@of_inj _ _ H _ (embedding_of_subset _ _ h).2⟩ + +lemma countable_range [countable ι] (f : ι → β) : (range f).countable := +surjective_onto_range.countable.to_set + +lemma countable_iff_exists_subset_range [nonempty α] {s : set α} : + s.countable ↔ ∃ f : ℕ → α, s ⊆ range f := +⟨λ h, by { inhabit α, exact ⟨enumerate_countable h default, subset_range_enumerate _ _⟩ }, + λ ⟨f, hsf⟩, (countable_range f).mono hsf⟩ /-- A non-empty set is countable iff there exists a surjection from the @@ -60,53 +90,29 @@ natural numbers onto the subtype induced by the set. -/ protected lemma countable_iff_exists_surjective {s : set α} (hs : s.nonempty) : s.countable ↔ ∃ f : ℕ → s, surjective f := -have inhabited s, from ⟨classical.choice hs.to_subtype⟩, -have s.countable → ∃ f : ℕ → s, surjective f, from assume ⟨h⟩, - by exactI ⟨λ n, (decode s n).iget, λ a, ⟨encode a, by simp [encodek]⟩⟩, -have (∃ f : ℕ → s, surjective f) → s.countable, from assume ⟨f, fsurj⟩, - ⟨⟨inv_fun f, option.some ∘ f, - by intro h; simp [(inv_fun_eq (fsurj h) : f (inv_fun f h) = h)]⟩⟩, -by split; assumption +countable_coe_iff.symm.trans $ @countable_iff_exists_surjective s hs.to_subtype -/-- Convert `set.countable s` to `encodable s` (noncomputable). -/ -protected def countable.to_encodable {s : set α} : s.countable → encodable s := -classical.choice - -lemma countable_encodable' (s : set α) [H : encodable s] : s.countable := -⟨H⟩ +alias set.countable_iff_exists_surjective ↔ countable.exists_surjective _ -lemma countable_encodable [encodable α] (s : set α) : s.countable := -⟨by apply_instance⟩ +lemma countable_univ [countable α] : (univ : set α).countable := to_countable univ /-- If `s : set α` is a nonempty countable set, then there exists a map `f : ℕ → α` such that `s = range f`. -/ lemma countable.exists_eq_range {s : set α} (hc : s.countable) (hs : s.nonempty) : ∃ f : ℕ → α, s = range f := begin - letI : encodable s := countable.to_encodable hc, - letI : nonempty s := hs.to_subtype, - have : (univ : set s).countable := countable_encodable _, - rcases set.countable_iff_exists_subset_range.1 this with ⟨g, hg⟩, - have : range g = univ := univ_subset_iff.1 hg, - use coe ∘ g, - simp only [range_comp, this, image_univ, subtype.range_coe] + rcases hc.exists_surjective hs with ⟨f, hf⟩, + refine ⟨coe ∘ f, _⟩, + rw [hf.range_comp, subtype.range_coe] end -@[simp] lemma countable_empty : (∅ : set α).countable := -⟨⟨λ x, x.2.elim, λ n, none, λ x, x.2.elim⟩⟩ +@[simp] lemma countable_empty : (∅ : set α).countable := to_countable _ @[simp] lemma countable_singleton (a : α) : ({a} : set α).countable := ⟨of_equiv _ (equiv.set.singleton a)⟩ -lemma countable.mono {s₁ s₂ : set α} (h : s₁ ⊆ s₂) : s₂.countable → s₁.countable -| ⟨H⟩ := ⟨@of_inj _ _ H _ (embedding_of_subset _ _ h).2⟩ - lemma countable.image {s : set α} (hs : s.countable) (f : α → β) : (f '' s).countable := -have surjective ((maps_to_image f s).restrict _ _ _), from surjective_maps_to_image_restrict f s, -⟨@encodable.of_inj _ _ hs.to_encodable (surj_inv this) (injective_surj_inv this)⟩ - -lemma countable_range [encodable α] (f : α → β) : (range f).countable := -by rw ← image_univ; exact (countable_encodable _).image _ +by { rw [image_eq_range], haveI := hs.to_subtype, apply countable_range } lemma maps_to.countable_of_inj_on {s : set α} {t : set β} {f : α → β} (hf : maps_to f s t) (hf' : inj_on f s) (ht : t.countable) : @@ -148,17 +154,16 @@ lemma countable_of_injective_of_countable_image {s : set α} {f : α → β} let ⟨g, hg⟩ := countable_iff_exists_inj_on.1 hs in countable_iff_exists_inj_on.2 ⟨g ∘ f, hg.comp hf (maps_to_image _ _)⟩ -lemma countable_Union {t : α → set β} [encodable α] (ht : ∀a, (t a).countable) : +lemma countable_Union {t : ι → set α} [countable ι] (ht : ∀ i, (t i).countable) : (⋃a, t a).countable := -by haveI := (λ a, (ht a).to_encodable); - rw Union_eq_range_sigma; apply countable_range +by { haveI := λ a, (ht a).to_subtype, rw Union_eq_range_psigma, apply countable_range } lemma countable.bUnion {s : set α} {t : Π x ∈ s, set β} (hs : s.countable) (ht : ∀a∈s, (t a ‹_›).countable) : (⋃a∈s, t a ‹_›).countable := begin rw bUnion_eq_Union, - haveI := hs.to_encodable, + haveI := hs.to_subtype, exact countable_Union (by simpa using ht) end @@ -212,45 +217,29 @@ begin simpa using @ts a end -lemma countable_pi {π : α → Type*} [fintype α] {s : Πa, set (π a)} (hs : ∀a, (s a).countable) : +lemma countable_univ_pi {π : α → Type*} [finite α] {s : Π a, set (π a)} + (hs : ∀ a, (s a).countable) : (pi univ s).countable := +begin + haveI := λ a, (hs a).to_subtype, + exact (countable.of_equiv _ (equiv.set.univ_pi s).symm).to_set +end + +lemma countable_pi {π : α → Type*} [finite α] {s : Πa, set (π a)} (hs : ∀a, (s a).countable) : {f : Πa, π a | ∀a, f a ∈ s a}.countable := -countable.mono - (show {f : Πa, π a | ∀a, f a ∈ s a} ⊆ range (λf : Πa, s a, λa, (f a).1), from - assume f hf, ⟨λa, ⟨f a, hf a⟩, funext $ assume a, rfl⟩) $ -have trunc (encodable (Π (a : α), s a)), from - @encodable.fintype_pi α _ _ _ (assume a, (hs a).to_encodable), -trunc.induction_on this $ assume h, -@countable_range _ _ h _ +by simpa only [← mem_univ_pi] using countable_univ_pi hs protected lemma countable.prod {s : set α} {t : set β} (hs : s.countable) (ht : t.countable) : set.countable (s ×ˢ t) := begin - haveI : encodable s := hs.to_encodable, - haveI : encodable t := ht.to_encodable, - exact ⟨of_equiv (s × t) (equiv.set.prod _ _)⟩ + haveI : countable s := hs.to_subtype, + haveI : countable t := ht.to_subtype, + exact (countable.of_equiv _ $ (equiv.set.prod _ _).symm).to_set end lemma countable.image2 {s : set α} {t : set β} (hs : s.countable) (ht : t.countable) (f : α → β → γ) : (image2 f s t).countable := by { rw ← image_prod, exact (hs.prod ht).image _ } -section enumerate - -/-- Enumerate elements in a countable set.-/ -def enumerate_countable {s : set α} (h : s.countable) (default : α) : ℕ → α := -assume n, match @encodable.decode s h.to_encodable n with - | (some y) := y - | (none) := default - end - -lemma subset_range_enumerate {s : set α} (h : s.countable) (default : α) : - s ⊆ range (enumerate_countable h default) := -assume x hx, -⟨@encodable.encode s h.to_encodable ⟨x, hx⟩, -by simp [enumerate_countable, encodable.encodek]⟩ - -end enumerate - end set lemma finset.countable_to_set (s : finset α) : set.countable (↑s : set α) := diff --git a/src/dynamics/ergodic/conservative.lean b/src/dynamics/ergodic/conservative.lean index 58899311fd3c6..0835c3ff44a00 100644 --- a/src/dynamics/ergodic/conservative.lean +++ b/src/dynamics/ergodic/conservative.lean @@ -82,10 +82,10 @@ begin rcases ihN with ⟨n, hn, hμn⟩, set T := s ∩ ⋃ n ≥ N + 1, (f^[n]) ⁻¹' s, have hT : measurable_set T, - from hs.inter (measurable_set.bUnion (countable_encodable _) + from hs.inter (measurable_set.bUnion (to_countable _) (λ _ _, hf.measurable.iterate _ hs)), have hμT : μ T = 0, - { convert (measure_bUnion_null_iff $ countable_encodable _).2 hN, + { convert (measure_bUnion_null_iff $ to_countable _).2 hN, rw ←inter_Union₂, refl }, have : μ ((s ∩ (f^[n]) ⁻¹' s) \ T) ≠ 0, by rwa [measure_diff_null hμT], rcases hf.exists_mem_image_mem ((hs.inter (hf.measurable.iterate n hs)).diff hT) this @@ -114,7 +114,7 @@ begin by_contradiction H, have : measurable_set (s ∩ {x | ∀ m ≥ n, f^[m] x ∉ s}), { simp only [set_of_forall, ← compl_set_of], - exact hs.inter (measurable_set.bInter (countable_encodable _) + exact hs.inter (measurable_set.bInter (to_countable _) (λ m _, hf.measurable.iterate m hs.compl)) }, rcases (hf.exists_gt_measure_inter_ne_zero this H) n with ⟨m, hmn, hm⟩, rcases nonempty_of_measure_ne_zero hm with ⟨x, ⟨hxs, hxn⟩, hxm, -⟩, diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 53284c77942ba..55162094112d9 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -64,7 +64,7 @@ lemma borel_eq_top_of_encodable [topological_space α] [t1_space α] [encodable borel α = ⊤ := begin refine (top_le_iff.1 $ λ s hs, bUnion_of_singleton s ▸ _), - apply measurable_set.bUnion s.countable_encodable, + apply measurable_set.bUnion s.to_countable, intros x hx, apply measurable_set.of_compl, apply generate_measurable.basic, @@ -1281,14 +1281,14 @@ end @[measurability] lemma measurable_liminf {f : ℕ → δ → α} (hf : ∀ i, measurable (f i)) : measurable (λ x, liminf at_top (λ i, f i x)) := -measurable_liminf' hf at_top_countable_basis (λ i, countable_encodable _) +measurable_liminf' hf at_top_countable_basis (λ i, to_countable _) /-- `limsup` over `ℕ` is measurable. See `measurable_limsup'` for a version with a general filter. -/ @[measurability] lemma measurable_limsup {f : ℕ → δ → α} (hf : ∀ i, measurable (f i)) : measurable (λ x, limsup at_top (λ i, f i x)) := -measurable_limsup' hf at_top_countable_basis (λ i, countable_encodable _) +measurable_limsup' hf at_top_countable_basis (λ i, to_countable _) end complete_linear_order @@ -1600,7 +1600,7 @@ begin { have hg : ∀ q : ℚ, measurable_set[g] (Iio q) := λ q, generate_measurable.basic (Iio q) (by simp), refine @measurable_set.inter _ g _ _ _ (hg _), - refine @measurable_set.bUnion _ _ g _ _ (countable_encodable _) (λ c h, _), + refine @measurable_set.bUnion _ _ g _ _ (to_countable _) (λ c h, _), exact @measurable_set.compl _ _ g (hg _) }, { suffices : x < ↑b → (↑a < x ↔ ∃ (i : ℚ), a < i ∧ ↑i ≤ x), by simpa, refine λ _, ⟨λ h, _, λ ⟨i, hai, hix⟩, (rat.cast_lt.2 hai).trans_le hix⟩, diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 7c41df895fa79..3dc48900f73ed 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -106,7 +106,7 @@ begin have hf : ∀n m, measurable_set (f n m), { assume n m, simp only [f, finset.inf_eq_infi], - exact measurable_set.bInter (countable_encodable _) (assume i _, he₁ _) }, + exact measurable_set.bInter (to_countable _) (assume i _, he₁ _) }, have f_subset_f : ∀{a b c d}, a ≤ b → c ≤ d → f a d ⊆ f b c, { assume a b c d hab hcd, diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 18f1050d4bf0a..ed7b5adcea130 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -312,7 +312,7 @@ lemma ae_cover.bUnion_Iic_ae_cover [preorder ι] {φ : ι → set α} (hφ : ae_ ae_cover μ at_top (λ (n : ι), ⋃ k (h : k ∈ Iic n), φ k) := { ae_eventually_mem := hφ.ae_eventually_mem.mono (λ x h, h.mono (λ i hi, mem_bUnion right_mem_Iic hi)), - measurable := λ i, measurable_set.bUnion (countable_encodable _) (λ n _, hφ.measurable n) } + measurable := λ i, measurable_set.bUnion (to_countable _) (λ n _, hφ.measurable n) } lemma ae_cover.bInter_Ici_ae_cover [semilattice_sup ι] [nonempty ι] {φ : ι → set α} (hφ : ae_cover μ at_top φ) : ae_cover μ at_top (λ (n : ι), ⋂ k (h : k ∈ Ici n), φ k) := @@ -325,7 +325,7 @@ lemma ae_cover.bInter_Ici_ae_cover [semilattice_sup ι] [nonempty ι] {φ : ι intros j hj, exact mem_bInter (λ k hk, hi k (le_trans hj hk)), end, - measurable := λ i, measurable_set.bInter (countable_encodable _) (λ n _, hφ.measurable n) } + measurable := λ i, measurable_set.bInter (to_countable _) (λ n _, hφ.measurable n) } end ae_cover_Union_Inter_encodable diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 10c2a1f86d3e7..5779306ddc4d6 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1824,7 +1824,7 @@ calc by simp only [liminf_eq_supr_infi_of_nat] ... = ⨆n:ℕ, ∫⁻ a, ⨅i≥n, f i a ∂μ : lintegral_supr' - (assume n, ae_measurable_binfi _ (countable_encodable _) h_meas) + (assume n, ae_measurable_binfi _ (to_countable _) h_meas) (ae_of_all μ (assume a n m hnm, infi_le_infi_of_subset $ λ i hi, le_trans hnm hi)) ... ≤ ⨆n:ℕ, ⨅i≥n, ∫⁻ a, f i a ∂μ : supr_mono $ λ n, le_infi₂_lintegral _ @@ -1846,7 +1846,7 @@ calc ... = ∫⁻ a, ⨅n:ℕ, ⨆i≥n, f i a ∂μ : begin refine (lintegral_infi _ _ _).symm, - { assume n, exact measurable_bsupr _ (countable_encodable _) hf_meas }, + { assume n, exact measurable_bsupr _ (to_countable _) hf_meas }, { assume n m hnm a, exact (supr_le_supr_of_subset $ λ i hi, le_trans hnm hi) }, { refine ne_top_of_le_ne_top h_fin (lintegral_mono_ae _), refine (ae_all_iff.2 h_bound).mono (λ n hn, _), diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 39d00b6d71f93..214a70853781f 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -679,7 +679,7 @@ by { rw [pi_def], exact measurable_set.bInter hs (λ i hi, measurable_pi_apply _ lemma measurable_set.univ_pi [encodable δ] {t : Π i : δ, set (π i)} (ht : ∀ i, measurable_set (t i)) : measurable_set (pi univ t) := -measurable_set.pi (countable_encodable _) (λ i _, ht i) +measurable_set.pi (to_countable _) (λ i _, ht i) lemma measurable_set_pi_of_nonempty {s : set δ} {t : Π i, set (π i)} (hs : s.countable) @@ -735,7 +735,7 @@ local attribute [instance] fintype.to_encodable lemma measurable_set.pi_fintype [fintype δ] {s : set δ} {t : Π i, set (π i)} (ht : ∀ i ∈ s, measurable_set (t i)) : measurable_set (pi s t) := -measurable_set.pi (countable_encodable _) ht +measurable_set.pi (to_countable _) ht lemma measurable_set.univ_pi_fintype [fintype δ] {t : Π i, set (π i)} (ht : ∀ i, measurable_set (t i)) : measurable_set (pi univ t) := diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index 9ebd4c2d60837..53efebd9be758 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -33,8 +33,8 @@ lemma exists_null_pairwise_disjoint_diff [encodable ι] {s : ι → set α} begin refine ⟨λ i, to_measurable μ (s i ∩ ⋃ j ∈ ({i}ᶜ : set ι), s j), λ i, measurable_set_to_measurable _ _, λ i, _, _⟩, - { simp only [measure_to_measurable, inter_Union, measure_bUnion_null_iff (countable_encodable _)], - exact λ j hj, hd _ _ (ne.symm hj) }, + { simp only [measure_to_measurable, inter_Union], + exact (measure_bUnion_null_iff $ to_countable _).2 (λ j hj, hd _ _ (ne.symm hj)) }, { simp only [pairwise, disjoint_left, on_fun, mem_diff, not_and, and_imp, not_not], intros i j hne x hi hU hj, replace hU : x ∉ s i ∩ ⋃ j ≠ i, s j := λ h, hU (subset_to_measurable _ _ h), diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 548a46c09e353..431bbcde4d5b0 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1630,7 +1630,7 @@ lemma map_eq_sum [encodable β] [measurable_singleton_class β] begin ext1 s hs, have : ∀ y ∈ s, measurable_set (f ⁻¹' {y}), from λ y _, hf (measurable_set_singleton _), - simp [← tsum_measure_preimage_singleton (countable_encodable s) this, *, + simp [← tsum_measure_preimage_singleton (to_countable s) this, *, tsum_subtype s (λ b, μ (f ⁻¹' {b})), ← indicator_mul_right s (λ b, μ (f ⁻¹' {b}))] end diff --git a/src/model_theory/fraisse.lean b/src/model_theory/fraisse.lean index 7f923196040ad..4b8b79d05201f 100644 --- a/src/model_theory/fraisse.lean +++ b/src/model_theory/fraisse.lean @@ -238,7 +238,7 @@ begin { rintros ⟨Kn, eqinv, cq, hfg, hp, jep⟩, obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn eqinv cq hfg hp jep, haveI := ((Structure.cg_iff_countable).1 hM).some, - refine ⟨M, countable_encodable _, rfl⟩, } + refine ⟨M, to_countable _, rfl⟩, } end variables {K} (L) (M) diff --git a/src/number_theory/liouville/measure.lean b/src/number_theory/liouville/measure.lean index c97376817ebf9..1b8f44c5241fc 100644 --- a/src/number_theory/liouville/measure.lean +++ b/src/number_theory/liouville/measure.lean @@ -75,7 +75,7 @@ begin simp only [← set_of_exists], refine measure_mono_null set_of_liouville_with_subset_aux _, rw measure_Union_null_iff, intro m, rw measure_preimage_add_right, clear m, - refine (measure_bUnion_null_iff $ countable_encodable _).2 (λ n (hn : 1 ≤ n), _), + refine (measure_bUnion_null_iff $ to_countable _).2 (λ n (hn : 1 ≤ n), _), generalize hr : (2 + 1 / n : ℝ) = r, replace hr : 2 < r, by simp [← hr, zero_lt_one.trans_le hn], clear hn n, refine measure_set_of_frequently_eq_zero _, diff --git a/src/order/filter/archimedean.lean b/src/order/filter/archimedean.lean index 396a2852d4ea1..11cb64425b1f3 100644 --- a/src/order/filter/archimedean.lean +++ b/src/order/filter/archimedean.lean @@ -78,14 +78,14 @@ by rw [← tendsto_comap_iff, rat.comap_coe_at_bot] lemma at_top_countable_basis_of_archimedean [linear_ordered_semiring R] [archimedean R] : (at_top : filter R).has_countable_basis (λ n : ℕ, true) (λ n, Ici n) := -{ countable := countable_encodable _, +{ countable := to_countable _, to_has_basis := at_top_basis.to_has_basis (λ x hx, let ⟨n, hn⟩ := exists_nat_ge x in ⟨n, trivial, Ici_subset_Ici.2 hn⟩) (λ n hn, ⟨n, trivial, subset.rfl⟩) } lemma at_bot_countable_basis_of_archimedean [linear_ordered_ring R] [archimedean R] : (at_bot : filter R).has_countable_basis (λ m : ℤ, true) (λ m, Iic m) := -{ countable := countable_encodable _, +{ countable := to_countable _, to_has_basis := at_bot_basis.to_has_basis (λ x hx, let ⟨m, hm⟩ := exists_int_lt x in ⟨m, trivial, Iic_subset_Iic.2 hm.le⟩) (λ m hm, ⟨m, trivial, subset.rfl⟩) } diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 38ef5625e9c47..f5a72c2c263a1 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -191,12 +191,12 @@ at_top_basis.to_has_basis (λ a ha, ⟨a, ha, Ioi_subset_Ici_self⟩) $ lemma at_top_countable_basis [nonempty α] [semilattice_sup α] [encodable α] : has_countable_basis (at_top : filter α) (λ _, true) Ici := -{ countable := countable_encodable _, +{ countable := to_countable _, .. at_top_basis } lemma at_bot_countable_basis [nonempty α] [semilattice_inf α] [encodable α] : has_countable_basis (at_bot : filter α) (λ _, true) Iic := -{ countable := countable_encodable _, +{ countable := to_countable _, .. at_bot_basis } @[priority 200] diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index b9e61e19539c1..cc646bcbb2572 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -914,13 +914,13 @@ begin rcases f.exists_antitone_basis with ⟨s, hs⟩, rcases g.exists_antitone_basis with ⟨t, ht⟩, exact has_countable_basis.is_countably_generated - ⟨hs.to_has_basis.inf ht.to_has_basis, set.countable_encodable _⟩ + ⟨hs.to_has_basis.inf ht.to_has_basis, set.to_countable _⟩ end instance comap.is_countably_generated (l : filter β) [l.is_countably_generated] (f : α → β) : (comap f l).is_countably_generated := let ⟨x, hxl⟩ := l.exists_antitone_basis in -has_countable_basis.is_countably_generated ⟨hxl.to_has_basis.comap _, countable_encodable _⟩ +has_countable_basis.is_countably_generated ⟨hxl.to_has_basis.comap _, to_countable _⟩ instance sup.is_countably_generated (f g : filter α) [is_countably_generated f] [is_countably_generated g] : @@ -929,7 +929,7 @@ begin rcases f.exists_antitone_basis with ⟨s, hs⟩, rcases g.exists_antitone_basis with ⟨t, ht⟩, exact has_countable_basis.is_countably_generated - ⟨hs.to_has_basis.sup ht.to_has_basis, set.countable_encodable _⟩ + ⟨hs.to_has_basis.sup ht.to_has_basis, set.to_countable _⟩ end end is_countably_generated diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index 24c4b321cb627..ae804da5a3940 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -462,7 +462,7 @@ end protected lemma measurable_set_eq_of_encodable [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {a | τ a = i} := -hτ.measurable_set_eq_of_countable (set.countable_encodable _) i +hτ.measurable_set_eq_of_countable (set.to_countable _) i protected lemma measurable_set_lt_of_countable (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : @@ -476,7 +476,7 @@ end protected lemma measurable_set_lt_of_encodable [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {a | τ a < i} := -hτ.measurable_set_lt_of_countable (set.countable_encodable _) i +hτ.measurable_set_lt_of_countable (set.to_countable _) i protected lemma measurable_set_ge_of_countable {ι} [linear_order ι] {τ : α → ι} {f : filtration ι m} (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : @@ -491,7 +491,7 @@ end protected lemma measurable_set_ge_of_encodable {ι} [linear_order ι] {τ : α → ι} {f : filtration ι m} [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {a | i ≤ τ a} := -hτ.measurable_set_ge_of_countable (set.countable_encodable _) i +hτ.measurable_set_ge_of_countable (set.to_countable _) i end is_stopping_time @@ -595,7 +595,7 @@ lemma is_stopping_time_of_measurable_set_eq [preorder ι] [encodable ι] begin intro i, rw show {x | τ x ≤ i} = ⋃ k ≤ i, {x | τ x = k}, by { ext, simp }, - refine measurable_set.bUnion (set.countable_encodable _) (λ k hk, _), + refine measurable_set.bUnion (set.to_countable _) (λ k hk, _), exact f.mono hk _ (hτ k), end @@ -897,7 +897,7 @@ end protected lemma measurable_set_eq_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {a | τ a = i} := -hτ.measurable_set_eq_of_countable' (set.countable_encodable _) i +hτ.measurable_set_eq_of_countable' (set.to_countable _) i protected lemma measurable_set_ge_of_countable' (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : @@ -913,7 +913,7 @@ end protected lemma measurable_set_ge_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {a | i ≤ τ a} := -hτ.measurable_set_ge_of_countable' (set.countable_encodable _) i +hτ.measurable_set_ge_of_countable' (set.to_countable _) i protected lemma measurable_set_lt_of_countable' (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : @@ -928,7 +928,7 @@ end protected lemma measurable_set_lt_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {a | τ a < i} := -hτ.measurable_set_lt_of_countable' (set.countable_encodable _) i +hτ.measurable_set_lt_of_countable' (set.to_countable _) i protected lemma measurable_space_le_of_countable (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) : diff --git a/src/topology/G_delta.lean b/src/topology/G_delta.lean index a33ee4ce59df6..71ae4d3de353e 100644 --- a/src/topology/G_delta.lean +++ b/src/topology/G_delta.lean @@ -111,7 +111,7 @@ lemma is_closed.is_Gδ {α} [uniform_space α] [is_countably_generated (𝓤 α) begin rcases (@uniformity_has_basis_open α _).exists_antitone_subbasis with ⟨U, hUo, hU, -⟩, rw [← hs.closure_eq, ← hU.bInter_bUnion_ball], - refine is_Gδ_bInter (countable_encodable _) (λ n hn, is_open.is_Gδ _), + refine is_Gδ_bInter (to_countable _) (λ n hn, is_open.is_Gδ _), exact is_open_bUnion (λ x hx, uniform_space.is_open_ball _ (hUo _).2) end @@ -145,7 +145,7 @@ lemma is_Gδ_singleton (a : α) : is_Gδ ({a} : set α) := begin rcases (nhds_basis_opens a).exists_antitone_subbasis with ⟨U, hU, h_basis⟩, rw [← bInter_basis_nhds h_basis.to_has_basis], - exact is_Gδ_bInter (countable_encodable _) (λ n hn, (hU n).2.is_Gδ), + exact is_Gδ_bInter (to_countable _) (λ n hn, (hU n).2.is_Gδ), end lemma set.finite.is_Gδ {s : set α} (hs : s.finite) : is_Gδ s := diff --git a/src/topology/bases.lean b/src/topology/bases.lean index d1ac7452e75ba..c57ca30354f10 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -292,7 +292,7 @@ variable {α} @[priority 100] instance encodable.to_separable_space [encodable α] : separable_space α := -{ exists_countable_dense := ⟨set.univ, set.countable_encodable set.univ, dense_univ⟩ } +{ exists_countable_dense := ⟨set.univ, set.countable_univ, dense_univ⟩ } lemma separable_space_of_dense_range {ι : Type*} [encodable ι] (u : ι → α) (hu : dense_range u) : separable_space α := From f538c91b94330ae96a2c3d59dea6feb8675dff8f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 23 Jul 2022 17:07:53 +0000 Subject: [PATCH 018/173] refactor(*): rename `nondiscrete_normed_field` (#15625) This renames `nondiscrete_normed_field` to `nontrivially_normed_field` in accordance with this [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/hypotheses.20for.20a.20field.20property/near/290408091) --- src/analysis/analytic/basic.lean | 2 +- src/analysis/analytic/composition.lean | 2 +- src/analysis/analytic/inverse.lean | 2 +- src/analysis/analytic/linear.lean | 2 +- src/analysis/analytic/radius_liminf.lean | 2 +- src/analysis/calculus/affine_map.lean | 2 +- src/analysis/calculus/cont_diff.lean | 14 +++--- src/analysis/calculus/deriv.lean | 8 ++-- src/analysis/calculus/diff_on_int_cont.lean | 6 +-- src/analysis/calculus/dslope.lean | 4 +- src/analysis/calculus/fderiv.lean | 12 ++--- src/analysis/calculus/fderiv_analytic.lean | 2 +- src/analysis/calculus/fderiv_measurable.lean | 4 +- .../calculus/formal_multilinear_series.lean | 2 +- src/analysis/calculus/implicit.lean | 8 ++-- src/analysis/calculus/inverse.lean | 2 +- src/analysis/calculus/iterated_deriv.lean | 4 +- src/analysis/calculus/tangent_cone.lean | 2 +- src/analysis/complex/basic.lean | 2 +- src/analysis/convolution.lean | 6 +-- .../locally_convex/balanced_core_hull.lean | 2 +- src/analysis/locally_convex/basic.lean | 6 +-- src/analysis/locally_convex/bounded.lean | 4 +- src/analysis/locally_convex/polar.lean | 6 +-- .../locally_convex/with_seminorms.lean | 6 +-- src/analysis/normed/field/basic.lean | 16 +++---- .../normed_space/add_torsor_bases.lean | 2 +- src/analysis/normed_space/algebra.lean | 2 +- src/analysis/normed_space/banach.lean | 2 +- .../normed_space/banach_steinhaus.lean | 2 +- src/analysis/normed_space/basic.lean | 12 ++--- .../normed_space/bounded_linear_maps.lean | 4 +- src/analysis/normed_space/complemented.lean | 2 +- src/analysis/normed_space/completion.lean | 2 +- .../normed_space/continuous_affine_map.lean | 2 +- src/analysis/normed_space/dual.lean | 6 +-- src/analysis/normed_space/exponential.lean | 6 +-- .../normed_space/finite_dimension.lean | 8 ++-- src/analysis/normed_space/is_R_or_C.lean | 2 +- src/analysis/normed_space/multilinear.lean | 6 +-- src/analysis/normed_space/operator_norm.lean | 45 +++++++++-------- src/analysis/normed_space/riesz_lemma.lean | 4 +- src/analysis/normed_space/spectrum.lean | 12 ++--- src/analysis/normed_space/weak_dual.lean | 4 +- src/analysis/special_functions/exp_deriv.lean | 6 +-- .../special_functions/exponential.lean | 6 +-- src/analysis/specific_limits/normed.lean | 4 +- src/data/complex/is_R_or_C.lean | 4 +- .../algebra/left_invariant_derivation.lean | 2 +- src/geometry/manifold/algebra/lie_group.lean | 12 ++--- src/geometry/manifold/algebra/monoid.lean | 12 ++--- .../manifold/algebra/smooth_functions.lean | 2 +- src/geometry/manifold/algebra/structures.lean | 6 +-- src/geometry/manifold/cont_mdiff.lean | 2 +- src/geometry/manifold/cont_mdiff_map.lean | 2 +- src/geometry/manifold/derivation_bundle.lean | 2 +- src/geometry/manifold/diffeomorph.lean | 2 +- .../instances/units_of_normed_algebra.lean | 2 +- src/geometry/manifold/mfderiv.lean | 14 +++--- .../smooth_manifold_with_corners.lean | 48 +++++++++---------- src/geometry/manifold/tangent_bundle.lean | 8 ++-- .../constructions/borel_space.lean | 10 ++-- src/measure_theory/function/l1_space.lean | 4 +- src/measure_theory/function/lp_space.lean | 6 +-- .../function/strongly_measurable.lean | 8 ++-- src/measure_theory/integral/bochner.lean | 4 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/integral/set_to_l1.lean | 4 +- .../measure/with_density_vector_measure.lean | 4 +- src/number_theory/padics/padic_numbers.lean | 2 +- src/topology/algebra/module/basic.lean | 4 +- .../algebra/module/finite_dimension.lean | 14 +++--- src/topology/continuous_function/bounded.lean | 2 +- src/topology/continuous_function/compact.lean | 2 +- .../metric_space/hausdorff_dimension.lean | 2 +- src/topology/vector_bundle/basic.lean | 4 +- src/topology/vector_bundle/hom.lean | 2 +- src/topology/vector_bundle/prod.lean | 2 +- src/topology/vector_bundle/pullback.lean | 2 +- 79 files changed, 233 insertions(+), 230 deletions(-) diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index f620f42501c54..6b18d9f34ad18 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -100,7 +100,7 @@ end formal_multilinear_series /-! ### The radius of a formal multilinear series -/ -variables [nondiscrete_normed_field 𝕜] +variables [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] [normed_group G] [normed_space 𝕜 G] diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index e7d04555e7627..6a38ef670628b 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -288,7 +288,7 @@ end formal_multilinear_series end topological -variables [nondiscrete_normed_field 𝕜] +variables [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] [normed_group G] [normed_space 𝕜 G] diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 7b8f71305b137..ee7453d36ddbb 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -31,7 +31,7 @@ open finset filter namespace formal_multilinear_series -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] diff --git a/src/analysis/analytic/linear.lean b/src/analysis/analytic/linear.lean index 7ce5a69a2d8f0..009ea64a77746 100644 --- a/src/analysis/analytic/linear.lean +++ b/src/analysis/analytic/linear.lean @@ -12,7 +12,7 @@ In this file we prove that a `continuous_linear_map` defines an analytic functio the formal power series `f x = f a + f (x - a)`. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index ecb1d750d27da..0220c26796720 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -13,7 +13,7 @@ In this file we prove that the radius of convergence of a `formal_multilinear_se $\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`. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] diff --git a/src/analysis/calculus/affine_map.lean b/src/analysis/calculus/affine_map.lean index 8aee84800f4ce..331bccf98693f 100644 --- a/src/analysis/calculus/affine_map.lean +++ b/src/analysis/calculus/affine_map.lean @@ -19,7 +19,7 @@ This file contains results about smoothness of affine maps. namespace continuous_affine_map -variables {𝕜 V W : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 V W : Type*} [nontrivially_normed_field 𝕜] variables [normed_group V] [normed_space 𝕜 V] variables [normed_group W] [normed_space 𝕜 W] diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index bb3bd18bde29a..c98832dce5957 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -33,7 +33,7 @@ existence of a nice sequence of derivatives, expressed with a predicate We prove basic properties of these notions. ## Main definitions and results -Let `f : E → F` be a map between normed vector spaces over a nondiscrete normed field `𝕜`. +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 `∞`). @@ -170,7 +170,7 @@ normed_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_com open set fin filter open_locale topological_space -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] @@ -2877,7 +2877,7 @@ theorem homeomorph.cont_diff_symm [complete_space E] (f : E ≃ₜ F) {f₀' : E cont_diff_iff_cont_diff_at.2 $ λ x, f.to_local_homeomorph.cont_diff_at_symm (mem_univ x) (hf₀' _) hf.cont_diff_at -/-- Let `f` be a local homeomorphism of a nondiscrete normed field, let `a` be a point in its +/-- Let `f` be a local homeomorphism of a nontrivially normed field, let `a` be a point in its target. if `f` is `n` times continuously differentiable at `f.symm a`, and if the derivative at `f.symm a` is nonzero, then `f.symm` is `n` times continuously differentiable at the point `a`. @@ -2889,9 +2889,9 @@ theorem local_homeomorph.cont_diff_at_symm_deriv [complete_space 𝕜] cont_diff_at 𝕜 n f.symm a := f.cont_diff_at_symm ha (hf₀'.has_fderiv_at_equiv h₀) hf -/-- Let `f` be an `n` times continuously differentiable homeomorphism of a nondiscrete normed field. -Suppose that the derivative of `f` is never equal to zero. Then `f.symm` is `n` times continuously -differentiable. +/-- Let `f` be an `n` times continuously differentiable homeomorphism of a nontrivially normed +field. Suppose that the derivative of `f` is never equal to zero. Then `f.symm` is `n` times +continuously differentiable. This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ @@ -3233,7 +3233,7 @@ situation where `ℂ` and `ℝ` are replaced respectively by `𝕜'` and `𝕜` over `𝕜`. -/ -variables (𝕜) {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables (𝕜) {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] variables [normed_space 𝕜' E] [is_scalar_tower 𝕜 𝕜' E] variables [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] variables {p' : E → formal_multilinear_series 𝕜' E F} diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 803f3f1801988..81e08abcf102b 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -88,7 +88,7 @@ open filter asymptotics set open continuous_linear_map (smul_right smul_right_one_eq_iff) -variables {𝕜 : Type u} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] section variables {F : Type v} [normed_group F] [normed_space 𝕜 F] @@ -790,7 +790,7 @@ section smul /-! ### Derivative of the multiplication of a scalar function and a vector function -/ -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {c : 𝕜 → 𝕜'} {c' : 𝕜'} theorem has_deriv_within_at.smul @@ -1112,7 +1112,7 @@ 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*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +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) @@ -1566,7 +1566,7 @@ end inverse section division /-! ### Derivative of `x ↦ c x / d x` -/ -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] {c d : 𝕜 → 𝕜'} {c' d' : 𝕜'} lemma has_deriv_within_at.div diff --git a/src/analysis/calculus/diff_on_int_cont.lean b/src/analysis/calculus/diff_on_int_cont.lean index 46ab34dde4d79..465b80892621c 100644 --- a/src/analysis/calculus/diff_on_int_cont.lean +++ b/src/analysis/calculus/diff_on_int_cont.lean @@ -16,7 +16,7 @@ this property and prove basic facts about this predicate. open set filter metric open_locale topological_space -variables (𝕜 : Type*) {E F G : Type*} [nondiscrete_normed_field 𝕜] [normed_group E] +variables (𝕜 : Type*) {E F G : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_group F] [normed_space 𝕜 E] [normed_space 𝕜 F] [normed_group G] [normed_space 𝕜 G] {f g : E → F} {s t : set E} {x : E} @@ -107,13 +107,13 @@ lemma const_smul {R : Type*} [semiring R] [module R F] [smul_comm_class 𝕜 R F diff_cont_on_cl 𝕜 (c • f) s := ⟨hf.1.const_smul c, hf.2.const_smul c⟩ -lemma smul {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +lemma smul {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {c : E → 𝕜'} {f : E → F} {s : set E} (hc : diff_cont_on_cl 𝕜 c s) (hf : diff_cont_on_cl 𝕜 f s) : diff_cont_on_cl 𝕜 (λ x, c x • f x) s := ⟨hc.1.smul hf.1, hc.2.smul hf.2⟩ -lemma smul_const {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +lemma smul_const {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {c : E → 𝕜'} {s : set E} (hc : diff_cont_on_cl 𝕜 c s) (y : F) : diff_cont_on_cl 𝕜 (λ x, c x • y) s := diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 418605813896e..0d4dbf15b1e2c 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -9,7 +9,7 @@ import linear_algebra.affine_space.slope /-! # Slope of a differentiable function -Given a function `f : 𝕜 → E` from a nondiscrete normed field to a normed space over this field, +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`. @@ -20,7 +20,7 @@ differentiability. open_locale classical topological_space filter open function set filter -variables {𝕜 E : Type*} [nondiscrete_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] +variables {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] /-- `dslope f a b` is defined as `slope f a b = (b - a)⁻¹ • (f b - f a)` for `a ≠ b` and `deriv f a` for `a = b`. -/ diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index c586578826a02..39b841409d071 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -121,7 +121,7 @@ noncomputable theory section -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] variables {G : Type*} [normed_group G] [normed_space 𝕜 G] @@ -2267,7 +2267,7 @@ field: e.g., they work for `c : E → ℂ` and `f : E → F` provided that `F` i normed vector space. -/ -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] variables {c : E → 𝕜'} {c' : E →L[𝕜] 𝕜'} @@ -2913,7 +2913,7 @@ end section tangent_cone -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {f : E → F} {s : set E} {f' : E →L[𝕜] F} @@ -2978,8 +2978,8 @@ we give variants of this statement, in the general situation where `ℂ` and ` respectively by `𝕜'` and `𝕜` where `𝕜'` is a normed algebra over `𝕜`. -/ -variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] variables [is_scalar_tower 𝕜 𝕜' E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] [normed_space 𝕜' F] @@ -3054,7 +3054,7 @@ end restrict_scalars section support open function -variables (𝕜 : Type*) {E F : Type*} [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type*) {E F : Type*} [nontrivially_normed_field 𝕜] variables [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] {f : E → F} lemma support_fderiv_subset : support (fderiv 𝕜 f) ⊆ tsupport f := diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index b18d50ad74cf4..c960d0386e654 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -17,7 +17,7 @@ Also the special case in terms of `deriv` when the domain is 1-dimensional. open filter asymptotics open_locale ennreal -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 047eeaa37efb2..ad43c5975cc4d 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -80,7 +80,7 @@ open_locale topological_space namespace continuous_linear_map -variables {𝕜 E F : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 E F : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] lemma measurable_apply₂ [measurable_space E] [opens_measurable_space E] @@ -93,7 +93,7 @@ end continuous_linear_map section fderiv -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] variables {f : E → F} (K : set (E →L[𝕜] F)) diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index 0bd5f0927c5e0..66f0ee7693468 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -114,7 +114,7 @@ end namespace formal_multilinear_series -variables [nondiscrete_normed_field 𝕜] +variables [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] [normed_group G] [normed_space 𝕜 G] diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index ee25f945f903b..68558d8d73bb9 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -88,7 +88,7 @@ such that * the derivatives are surjective; * the kernels of the derivatives are complementary subspaces of `E`. -/ @[nolint has_inhabited_instance] -structure implicit_function_data (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +structure implicit_function_data (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] [complete_space E] (F : Type*) [normed_group F] [normed_space 𝕜 F] [complete_space F] (G : Type*) [normed_group G] [normed_space 𝕜 G] [complete_space G] := @@ -105,7 +105,7 @@ structure implicit_function_data (𝕜 : Type*) [nondiscrete_normed_field 𝕜] namespace implicit_function_data -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] {F : Type*} [normed_group F] [normed_space 𝕜 F] [complete_space F] {G : Type*} [normed_group G] [normed_space 𝕜 G] [complete_space G] @@ -203,7 +203,7 @@ Note that a map with these properties is not unique. E.g., different choices of complementary to `ker f'` lead to different maps `φ`. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] {F : Type*} [normed_group F] [normed_space 𝕜 F] [complete_space F] {f : E → F} {f' : E →L[𝕜] F} {a : E} @@ -335,7 +335,7 @@ complementary to `ker f'` lead to different maps `φ`. section finite_dimensional -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [complete_space 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] {F : Type*} [normed_group F] [normed_space 𝕜 F] [finite_dimensional 𝕜 F] (f : E → F) (f' : E →L[𝕜] F) {a : E} diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 5f832f1054b31..247b5082505bf 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -61,7 +61,7 @@ open_locale topological_space classical nnreal noncomputable theory -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] variables {G : Type*} [normed_group G] [normed_space 𝕜 G] diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 2b5968487d568..51be7f72aa90a 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -15,7 +15,7 @@ and prove their basic properties. ## Main definitions and results -Let `𝕜` be a nondiscrete normed field, and `F` a normed vector space over `𝕜`. Let `f : 𝕜 → F`. +Let `𝕜` be a nontrivially normed field, and `F` a normed vector space over `𝕜`. Let `f : 𝕜 → F`. * `iterated_deriv n f` is the `n`-th derivative of `f`, seen as a function from `𝕜` to `F`. It is defined as the `n`-th Fréchet derivative (which is a multilinear map) applied to the @@ -45,7 +45,7 @@ open_locale classical topological_space big_operators open filter asymptotics set -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index 600f7a9a7b723..7be5a67081f7c 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -30,7 +30,7 @@ property of uniqueness of the derivative is therefore proved in `fderiv.lean`, b properties of the tangent cone we prove here. -/ -variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] open filter set open_locale topological_space diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index a3adb83bb01bc..871f8a70623dd 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -51,7 +51,7 @@ instance : normed_field ℂ := norm_mul' := abs_mul, .. complex.field, .. complex.normed_group } -instance : nondiscrete_normed_field ℂ := +instance : nontrivially_normed_field ℂ := { non_trivial := ⟨2, by simp; norm_num⟩ } instance {R : Type*} [normed_field R] [normed_algebra R ℝ] : normed_algebra R ℂ := diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 4c10e72586f73..84fd5935c9834 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -84,9 +84,9 @@ variables {𝕜 G E E' E'' F F' F'' : Type*} variables [normed_group E] [normed_group E'] [normed_group E''] [normed_group F] variables {f f' : G → E} {g g' : G → E'} {x x' : G} {y y' : E} -section nondiscrete_normed_field +section nontrivially_normed_field -variables [nondiscrete_normed_field 𝕜] +variables [nontrivially_normed_field 𝕜] variables [normed_space 𝕜 E] [normed_space 𝕜 E'] [normed_space 𝕜 E''] [normed_space 𝕜 F] variables (L : E →L[𝕜] E' →L[𝕜] F) @@ -746,7 +746,7 @@ end cont_diff_bump_of_inner end measurability -end nondiscrete_normed_field +end nontrivially_normed_field open_locale convolution diff --git a/src/analysis/locally_convex/balanced_core_hull.lean b/src/analysis/locally_convex/balanced_core_hull.lean index 3db2d610bd7d7..2ba576738c4ce 100644 --- a/src/analysis/locally_convex/balanced_core_hull.lean +++ b/src/analysis/locally_convex/balanced_core_hull.lean @@ -198,7 +198,7 @@ end balanced_hull section topology -variables [nondiscrete_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [topological_space E] +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [topological_space E] [has_continuous_smul 𝕜 E] {U : set E} protected lemma is_closed.balanced_core (hU : is_closed U) : is_closed (balanced_core 𝕜 U) := diff --git a/src/analysis/locally_convex/basic.lean b/src/analysis/locally_convex/basic.lean index 1932ee782c066..e892e70092d59 100644 --- a/src/analysis/locally_convex/basic.lean +++ b/src/analysis/locally_convex/basic.lean @@ -324,8 +324,8 @@ lemma balanced.closure (hA : balanced 𝕜 A) : balanced 𝕜 (closure A) := end normed_field -section nondiscrete_normed_field -variables [nondiscrete_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} +section nontrivially_normed_field +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} lemma absorbs_zero_iff : absorbs 𝕜 s 0 ↔ (0 : E) ∈ s := begin @@ -340,4 +340,4 @@ end lemma absorbent.zero_mem (hs : absorbent 𝕜 s) : (0 : E) ∈ s := absorbs_zero_iff.1 $ absorbent_iff_forall_absorbs_singleton.1 hs _ -end nondiscrete_normed_field +end nontrivially_normed_field diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index 9872c2549ca24..c47f7268e7070 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -163,7 +163,7 @@ end bornology section uniform_add_group -variables (𝕜) [nondiscrete_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] +variables (𝕜) [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] variables [uniform_space E] [uniform_add_group E] [has_continuous_smul 𝕜 E] lemma totally_bounded.is_vonN_bounded {s : set E} (hs : totally_bounded s) : @@ -194,7 +194,7 @@ end uniform_add_group section vonN_bornology_eq_metric -variables (𝕜 E) [nondiscrete_normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] +variables (𝕜 E) [nontrivially_normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] namespace normed_space diff --git a/src/analysis/locally_convex/polar.lean b/src/analysis/locally_convex/polar.lean index fbabe99e3fd6a..aa00e7a61709f 100644 --- a/src/analysis/locally_convex/polar.lean +++ b/src/analysis/locally_convex/polar.lean @@ -111,9 +111,9 @@ end end normed_ring -section nondiscrete_normed_field +section nontrivially_normed_field -variables [nondiscrete_normed_field 𝕜] [add_comm_monoid E] [add_comm_monoid F] +variables [nontrivially_normed_field 𝕜] [add_comm_monoid E] [add_comm_monoid F] variables [module 𝕜 E] [module 𝕜 F] variables (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) @@ -131,6 +131,6 @@ begin ... = ε : mul_one _ end -end nondiscrete_normed_field +end nontrivially_normed_field end linear_map diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index e9b95ea6348a4..ed3ad0f744078 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -336,9 +336,9 @@ end end normed_space -section nondiscrete_normed_field +section nontrivially_normed_field -variables [nondiscrete_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι] +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι] variables {p : seminorm_family 𝕜 E ι} variables [topological_space E] @@ -390,7 +390,7 @@ begin exact ⟨1, zero_lt_one, λ _ _, zero_lt_one⟩, end -end nondiscrete_normed_field +end nontrivially_normed_field section continuous_bounded namespace seminorm diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 8787a8bbf2131..38d1510fc40f0 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -485,10 +485,10 @@ class normed_field (α : Type*) extends has_norm α, field α, metric_space α : (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul' : ∀ a b, norm (a * b) = norm a * norm b) -/-- A nondiscrete normed field is a normed field in which there is an element of norm different from -`0` and `1`. This makes it possible to bring any element arbitrarily close to `0` by multiplication -by the powers of any element, and thus to relate algebra and topology. -/ -class nondiscrete_normed_field (α : Type*) extends normed_field α := +/-- A nontrivially normed field is a normed field in which there is an element of norm different +from `0` and `1`. This makes it possible to bring any element arbitrarily close to `0` by +multiplication by the powers of any element, and thus to relate algebra and topology. -/ +class nontrivially_normed_field (α : Type*) extends normed_field α := (non_trivial : ∃ x : α, 1 < ∥x∥) section normed_field @@ -515,9 +515,9 @@ end normed_field namespace normed_field -variables (α) [nondiscrete_normed_field α] +variables (α) [nontrivially_normed_field α] -lemma exists_one_lt_norm : ∃x : α, 1 < ∥x∥ := ‹nondiscrete_normed_field α›.non_trivial +lemma exists_one_lt_norm : ∃x : α, 1 < ∥x∥ := ‹nontrivially_normed_field α›.non_trivial lemma exists_lt_norm (r : ℝ) : ∃ x : α, r < ∥x∥ := let ⟨w, hw⟩ := exists_one_lt_norm α in @@ -553,7 +553,7 @@ instance : normed_field ℝ := { norm_mul' := abs_mul, .. real.normed_group } -instance : nondiscrete_normed_field ℝ := +instance : nontrivially_normed_field ℝ := { non_trivial := ⟨2, by { unfold norm, rw abs_of_nonneg; norm_num }⟩ } namespace real @@ -661,7 +661,7 @@ instance : normed_field ℚ := norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } -instance : nondiscrete_normed_field ℚ := +instance : nontrivially_normed_field ℚ := { non_trivial := ⟨2, by { unfold norm, rw abs_of_nonneg; norm_num }⟩ } @[norm_cast, simp] lemma rat.norm_cast_real (r : ℚ) : ∥(r : ℝ)∥ = ∥r∥ := rfl diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index 80e0844abce21..bd01b1fc1a814 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -26,7 +26,7 @@ This file contains results about bases in normed affine spaces. section barycentric -variables {ι 𝕜 E P : Type*} [nondiscrete_normed_field 𝕜] [complete_space 𝕜] +variables {ι 𝕜 E P : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] variables [normed_group E] [normed_space 𝕜 E] [finite_dimensional 𝕜 E] variables [metric_space P] [normed_add_torsor E P] variables (b : affine_basis ι 𝕜 P) diff --git a/src/analysis/normed_space/algebra.lean b/src/analysis/normed_space/algebra.lean index fe60d558c6e46..dde5957d77518 100644 --- a/src/analysis/normed_space/algebra.lean +++ b/src/analysis/normed_space/algebra.lean @@ -32,7 +32,7 @@ variables {𝕜 : Type*} {A : Type*} namespace weak_dual namespace character_space -variables [nondiscrete_normed_field 𝕜] [normed_ring A] +variables [nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] [norm_one_class A] lemma norm_one (φ : character_space 𝕜 A) : ∥to_normed_dual (φ : weak_dual 𝕜 A)∥ = 1 := diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index fb75643187b77..0a2b5bb45e8ef 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -17,7 +17,7 @@ bounded linear map between Banach spaces has a bounded inverse. open function metric set filter finset open_locale classical topological_space big_operators nnreal -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] (f : E →L[𝕜] F) diff --git a/src/analysis/normed_space/banach_steinhaus.lean b/src/analysis/normed_space/banach_steinhaus.lean index e51b2f2e97a42..ea1ac98c6d01a 100644 --- a/src/analysis/normed_space/banach_steinhaus.lean +++ b/src/analysis/normed_space/banach_steinhaus.lean @@ -24,7 +24,7 @@ open set variables {E F 𝕜 𝕜₂ : Type*} [semi_normed_group E] [semi_normed_group F] -[nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] +[nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 391ff8cbc9974..c03a6a5809524 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -314,14 +314,14 @@ rescale_to_shell_semi_normed hc εpos (ne_of_lt (norm_pos_iff.2 hx)).symm end normed_group -section normed_space_nondiscrete +section nontrivially_normed_space -variables (𝕜 E : Type*) [nondiscrete_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] +variables (𝕜 E : Type*) [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [nontrivial E] include 𝕜 -/-- If `E` is a nontrivial normed space over a nondiscrete normed field `𝕜`, then `E` is unbounded: +/-- If `E` is a nontrivial normed space over a nontrivially normed field `𝕜`, then `E` is unbounded: for any `c : ℝ`, there exists a vector `x : E` with norm strictly greater than `c`. -/ lemma normed_space.exists_lt_norm (c : ℝ) : ∃ x : E, c < ∥x∥ := begin @@ -336,14 +336,14 @@ protected lemma normed_space.unbounded_univ : ¬bounded (univ : set E) := λ h, let ⟨R, hR⟩ := bounded_iff_forall_norm_le.1 h, ⟨x, hx⟩ := normed_space.exists_lt_norm 𝕜 E R in hx.not_le (hR x trivial) -/-- A normed vector space over a nondiscrete normed field is a noncompact space. This cannot be +/-- A normed vector space over a nontrivially normed field is a noncompact space. This cannot be an instance because in order to apply it, Lean would have to search for `normed_space 𝕜 E` with unknown `𝕜`. We register this as an instance in two cases: `𝕜 = E` and `𝕜 = ℝ`. -/ protected lemma normed_space.noncompact_space : noncompact_space E := ⟨λ h, normed_space.unbounded_univ 𝕜 _ h.bounded⟩ @[priority 100] -instance nondiscrete_normed_field.noncompact_space : noncompact_space 𝕜 := +instance nontrivially_normed_field.noncompact_space : noncompact_space 𝕜 := normed_space.noncompact_space 𝕜 𝕜 omit 𝕜 @@ -352,7 +352,7 @@ omit 𝕜 instance real_normed_space.noncompact_space [normed_space ℝ E] : noncompact_space E := normed_space.noncompact_space ℝ E -end normed_space_nondiscrete +end nontrivially_normed_space section normed_algebra diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index f105ef1de57db..505fd5d8af797 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -53,7 +53,7 @@ open_locale classical big_operators topological_space open filter (tendsto) metric continuous_linear_map -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] @@ -240,7 +240,7 @@ namespace continuous_linear_map -/ variables {R : Type*} -variables {𝕜₂ 𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [nondiscrete_normed_field 𝕜₂] +variables {𝕜₂ 𝕜' : Type*} [nontrivially_normed_field 𝕜'] [nontrivially_normed_field 𝕜₂] variables {M : Type*} [topological_space M] variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] variables {G' : Type*} [normed_group G'] [normed_space 𝕜₂ G'] [normed_space 𝕜' G'] diff --git a/src/analysis/normed_space/complemented.lean b/src/analysis/normed_space/complemented.lean index 331fcda1459c8..893dd4efe9ea2 100644 --- a/src/analysis/normed_space/complemented.lean +++ b/src/analysis/normed_space/complemented.lean @@ -20,7 +20,7 @@ is always a complemented subspace. complemented subspace, normed vector space -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] noncomputable theory diff --git a/src/analysis/normed_space/completion.lean b/src/analysis/normed_space/completion.lean index 8a7821cbcd9d4..a4fa6337279de 100644 --- a/src/analysis/normed_space/completion.lean +++ b/src/analysis/normed_space/completion.lean @@ -51,7 +51,7 @@ to_complₗᵢ.to_continuous_linear_map @[simp] lemma coe_to_complL : ⇑(to_complL : E →L[𝕜] completion E) = coe := rfl -@[simp] lemma norm_to_complL {𝕜 E : Type*} [nondiscrete_normed_field 𝕜] +@[simp] lemma norm_to_complL {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [nontrivial E] : ∥(to_complL : E →L[𝕜] completion E)∥ = 1 := (to_complₗᵢ : E →ₗᵢ[𝕜] completion E).norm_to_continuous_linear_map diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 1ef5420663c8a..837958ee17c8c 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -46,7 +46,7 @@ variables [normed_group V] [metric_space P] [normed_add_torsor V P] variables [normed_group W] [metric_space Q] [normed_add_torsor W Q] variables [normed_group W₂] [metric_space Q₂] [normed_add_torsor W₂ Q₂] variables [normed_field R] [normed_space R V] [normed_space R W] [normed_space R W₂] -variables [nondiscrete_normed_field 𝕜] [normed_space 𝕜 V] [normed_space 𝕜 W] [normed_space 𝕜 W₂] +variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 V] [normed_space 𝕜 W] [normed_space 𝕜 W₂] include V W diff --git a/src/analysis/normed_space/dual.lean b/src/analysis/normed_space/dual.lean index 50e91be276719..a5afb7dbfd9b2 100644 --- a/src/analysis/normed_space/dual.lean +++ b/src/analysis/normed_space/dual.lean @@ -40,7 +40,7 @@ universes u v namespace normed_space section general -variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] variables (E : Type*) [semi_normed_group E] [normed_space 𝕜 E] variables (F : Type*) [normed_group F] [normed_space 𝕜 F] @@ -143,11 +143,11 @@ open metric set normed_space /-- Given a subset `s` in a normed space `E` (over a field `𝕜`), the polar `polar 𝕜 s` is the subset of `dual 𝕜 E` consisting of those functionals which evaluate to something of norm at most one at all points `z ∈ s`. -/ -def polar (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +def polar (𝕜 : Type*) [nontrivially_normed_field 𝕜] {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] : set E → set (dual 𝕜 E) := (dual_pairing 𝕜 E).flip.polar -variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] lemma mem_polar_iff {x' : dual 𝕜 E} (s : set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ∥x' z∥ ≤ 1 := iff.rfl diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index dfd830406afcd..3135feba55d1b 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -161,7 +161,7 @@ section normed section any_field_any_algebra -variables {𝕂 𝔸 𝔹 : Type*} [nondiscrete_normed_field 𝕂] +variables {𝕂 𝔸 𝔹 : Type*} [nontrivially_normed_field 𝕂] variables [normed_ring 𝔸] [normed_ring 𝔹] [normed_algebra 𝕂 𝔸] [normed_algebra 𝕂 𝔹] lemma norm_exp_series_summable_of_mem_ball (x : 𝔸) @@ -295,7 +295,7 @@ end any_field_any_algebra section any_field_division_algebra -variables {𝕂 𝔸 : Type*} [nondiscrete_normed_field 𝕂] [normed_division_ring 𝔸] [normed_algebra 𝕂 𝔸] +variables {𝕂 𝔸 : Type*} [nontrivially_normed_field 𝕂] [normed_division_ring 𝔸] [normed_algebra 𝕂 𝔸] variables (𝕂) @@ -334,7 +334,7 @@ end any_field_division_algebra section any_field_comm_algebra -variables {𝕂 𝔸 : Type*} [nondiscrete_normed_field 𝕂] [normed_comm_ring 𝔸] [normed_algebra 𝕂 𝔸] +variables {𝕂 𝔸 : Type*} [nontrivially_normed_field 𝕂] [normed_comm_ring 𝔸] [normed_algebra 𝕂 𝔸] [complete_space 𝔸] /-- In a commutative Banach-algebra `𝔸` over a normed field `𝕂` of characteristic zero, diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 2b5c239e30de0..a3a8704e38761 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -14,8 +14,8 @@ import topology.instances.matrix /-! # Finite dimensional normed spaces over complete fields -Over a complete nondiscrete field, in finite dimension, all norms are equivalent and all linear maps -are continuous. Moreover, a finite-dimensional subspace is always complete and closed. +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. ## Main results: @@ -110,7 +110,7 @@ end affine_isometry section complete_field -variables {𝕜 : Type u} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] {E : Type v} [normed_group E] [normed_space 𝕜 E] {F : Type w} [normed_group F] [normed_space 𝕜 F] {F' : Type x} [add_comm_group F'] [module 𝕜 F'] [topological_space F'] @@ -592,7 +592,7 @@ by simp_rw [continuous_iff_continuous_on_univ, continuous_on_clm_apply] end complete_field section proper_field -variables (𝕜 : Type u) [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type u) [nontrivially_normed_field 𝕜] (E : Type v) [normed_group E] [normed_space 𝕜 E] [proper_space 𝕜] /-- Any finite-dimensional vector space over a proper field is proper. diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index 4143c1789b63c..ee947e962d0be 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -73,7 +73,7 @@ begin end /-- -`linear_map.bound_of_ball_bound` is a version of this over arbitrary nondiscrete normed fields. +`linear_map.bound_of_ball_bound` is a version of this over arbitrary nontrivially normed fields. It produces a less precise bound so we keep both versions. -/ lemma linear_map.bound_of_ball_bound' {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →ₗ[𝕜] 𝕜) (h : ∀ z ∈ closed_ball (0 : E) r, ∥f z∥ ≤ c) (z : E) : diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 6c75986eeba0e..09bbbddfcf5c2 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -68,7 +68,7 @@ local attribute [-instance] unique.subsingleton pi.subsingleton We use the following type variables in this file: -* `𝕜` : a `nondiscrete_normed_field`; +* `𝕜` : a `nontrivially_normed_field`; * `ι`, `ι'` : finite index types with decidable equality; * `E`, `E₁` : families of normed vector spaces over `𝕜` indexed by `i : ι`; * `E'` : a family of normed vector spaces over `𝕜` indexed by `i' : ι'`; @@ -80,7 +80,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 ι'] [nondiscrete_normed_field 𝕜] + [decidable_eq ι] [fintype ι] [decidable_eq ι'] [fintype ι'] [nontrivially_normed_field 𝕜] [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] [Π i, normed_group (E₁ i)] [Π i, normed_space 𝕜 (E₁ i)] [Π i, normed_group (E' i)] [Π i, normed_space 𝕜 (E' i)] @@ -478,7 +478,7 @@ end section restrict_scalars -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜' 𝕜] +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)] diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index d51dc37e88f30..a1a2339a4ec29 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -99,9 +99,9 @@ rfl end normed_field -variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜 Eₗ] [normed_space 𝕜₂ F] [normed_space 𝕜 Fₗ] - [normed_space 𝕜₃ G] [normed_space 𝕜 Gₗ] +variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] + [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜 Eₗ] [normed_space 𝕜₂ F] + [normed_space 𝕜 Fₗ] [normed_space 𝕜₃ G] [normed_space 𝕜 Gₗ] {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] @@ -133,10 +133,10 @@ begin ring_hom_isometric.is_iso] using hf (δ • x) leδx δxle end -/-- A continuous linear map between seminormed spaces is bounded when the field is nondiscrete. The -continuity ensures boundedness on a ball of some radius `ε`. The nondiscreteness is then used to -rescale any element into an element of norm in `[ε/C, ε]`, whose image has a controlled norm. The -norm control for the original element follows by rescaling. -/ +/-- A continuous linear map between seminormed spaces is bounded when the field is nontrivially +normed. The continuity ensures boundedness on a ball of some radius `ε`. The nontriviality of the +norm is then used to rescale any element into an element of norm in `[ε/C, ε]`, whose image has a +controlled norm. The norm control for the original element follows by rescaling. -/ lemma semilinear_map_class.bound_of_continuous [semilinear_map_class 𝓕 σ₁₂ E F] (f : 𝓕) (hf : continuous f) : ∃ C, 0 < C ∧ (∀ x : E, ∥f x∥ ≤ C * ∥x∥) := begin @@ -930,7 +930,7 @@ end smul_linear section restrict_scalars -variables {𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [normed_algebra 𝕜' 𝕜] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜' 𝕜] variables [normed_space 𝕜' E] [is_scalar_tower 𝕜' 𝕜 E] variables [normed_space 𝕜' Fₗ] [is_scalar_tower 𝕜' 𝕜 Fₗ] @@ -1119,8 +1119,8 @@ omit σ₂₁ namespace continuous_linear_map variables {E' F' : Type*} [semi_normed_group E'] [semi_normed_group F'] -variables {𝕜₁' : Type*} {𝕜₂' : Type*} [nondiscrete_normed_field 𝕜₁'] [nondiscrete_normed_field 𝕜₂'] - [normed_space 𝕜₁' E'] [normed_space 𝕜₂' F'] +variables {𝕜₁' : Type*} {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₁'] + [nontrivially_normed_field 𝕜₂'] [normed_space 𝕜₁' E'] [normed_space 𝕜₂' F'] {σ₁' : 𝕜₁' →+* 𝕜} {σ₁₃' : 𝕜₁' →+* 𝕜₃} {σ₂' : 𝕜₂' →+* 𝕜₂} {σ₂₃' : 𝕜₂' →+* 𝕜₃} [ring_hom_comp_triple σ₁' σ₁₃ σ₁₃'] [ring_hom_comp_triple σ₂' σ₂₃ σ₂₃'] [ring_hom_isometric σ₂₃] [ring_hom_isometric σ₁₃'] [ring_hom_isometric σ₂₃'] @@ -1164,8 +1164,9 @@ variables [normed_group E] [normed_group F] [normed_group G] [normed_group Fₗ] open metric continuous_linear_map section -variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜₃ G] [normed_space 𝕜 Fₗ] (c : 𝕜) +variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] + [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜₃ G] + [normed_space 𝕜 Fₗ] (c : 𝕜) {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} (f g : E →SL[σ₁₂] F) (x y z : E) @@ -1187,7 +1188,7 @@ lemma linear_map.bound_of_ball_bound {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E (h : ∀ z ∈ metric.ball (0 : E) r, ∥f z∥ ≤ c) : ∃ C, ∀ (z : E), ∥f z∥ ≤ C * ∥z∥ := begin - cases @nondiscrete_normed_field.non_trivial 𝕜 _ with k hk, + cases @nontrivially_normed_field.non_trivial 𝕜 _ with k hk, use c * (∥k∥ / r), intro z, refine linear_map.bound_of_shell _ r_pos hk (λ x hko hxo, _) _, @@ -1552,11 +1553,12 @@ end namespace continuous_linear_map -variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜₃ G] [normed_space 𝕜 Fₗ] (c : 𝕜) +variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] + [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜₃ G] + [normed_space 𝕜 Fₗ] (c : 𝕜) {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} -variables {𝕜₂' : Type*} [nondiscrete_normed_field 𝕜₂'] {F' : Type*} [normed_group F'] +variables {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₂'] {F' : Type*} [normed_group F'] [normed_space 𝕜₂' F'] {σ₂' : 𝕜₂' →+* 𝕜₂} {σ₂'' : 𝕜₂ →+* 𝕜₂'} {σ₂₃' : 𝕜₂' →+* 𝕜₃} [ring_hom_inv_pair σ₂' σ₂''] [ring_hom_inv_pair σ₂'' σ₂'] @@ -1669,8 +1671,8 @@ end end continuous_linear_map namespace submodule -variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} +variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] + [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} lemma norm_subtypeL (K : submodule 𝕜 E) [nontrivial K] : ∥K.subtypeL∥ = 1 := K.subtypeₗᵢ.norm_to_continuous_linear_map @@ -1678,8 +1680,9 @@ K.subtypeₗᵢ.norm_to_continuous_linear_map end submodule namespace continuous_linear_equiv -variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₁ : 𝕜₂ →+* 𝕜} +variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] + [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜₂ F] + {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] section @@ -1761,7 +1764,7 @@ end (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := linear_equiv.coord_self 𝕜 E x h -variables {𝕜} {𝕜₄ : Type*} [nondiscrete_normed_field 𝕜₄] +variables {𝕜} {𝕜₄ : Type*} [nontrivially_normed_field 𝕜₄] variables {H : Type*} [normed_group H] [normed_space 𝕜₄ H] [normed_space 𝕜₃ G] variables {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} variables {σ₃₄ : 𝕜₃ →+* 𝕜₄} {σ₄₃ : 𝕜₄ →+* 𝕜₃} diff --git a/src/analysis/normed_space/riesz_lemma.lean b/src/analysis/normed_space/riesz_lemma.lean index a6fd90c3a5b2a..a38eb7fa8e3d5 100644 --- a/src/analysis/normed_space/riesz_lemma.lean +++ b/src/analysis/normed_space/riesz_lemma.lean @@ -13,7 +13,7 @@ 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`. -In a nondiscrete normed field (with an element `c` of norm `> 1`) and any `R > ∥c∥`, one can +In a nontrivially normed field (with an element `c` of norm `> 1`) and any `R > ∥c∥`, one can guarantee `∥x∥ ≤ R` and `∥x - y∥ ≥ 1` for any `y` in `F`. This is `riesz_lemma_of_norm_lt`. A further lemma, `metric.closed_ball_inf_dist_compl_subset_closure`, finds a *closed* ball within @@ -69,7 +69,7 @@ which is at distance at least `1` of every element of `F`. Here, `R` is any giv strictly larger than the norm of an element of norm `> 1`. For a version without an `R`, see `riesz_lemma`. -Since we are considering a general nondiscrete normed field, there may be a gap in possible norms +Since we are considering a general nontrivially normed field, there may be a gap in possible norms (for instance no element of norm in `(1,2)`). Hence, we can not allow `R` arbitrarily close to `1`, and require `R > ∥c∥` for some `c : 𝕜` with norm `> 1`. -/ diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 18220933269a5..18f9ec5c54eaf 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -130,7 +130,7 @@ section resolvent open filter asymptotics -variables [nondiscrete_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] +variables [nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] local notation `ρ` := resolvent_set 𝕜 local notation `↑ₐ` := algebra_map 𝕜 A @@ -185,10 +185,10 @@ open continuous_multilinear_map ennreal formal_multilinear_series open_locale nnreal ennreal variables -[nondiscrete_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] +[nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] variable (𝕜) -/-- In a Banach algebra `A` over a nondiscrete normed field `𝕜`, for any `a : A` the +/-- In a Banach algebra `A` over a nontrivially normed field `𝕜`, for any `a : A` the power series with coefficients `a ^ n` represents the function `(1 - z • a)⁻¹` in a disk of radius `∥a∥₊⁻¹`. -/ lemma has_fpower_series_on_ball_inverse_one_sub_smul [complete_space A] (a : A) : @@ -424,8 +424,8 @@ lemma continuous [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : continuous φ : end normed_field -section nondiscrete_normed_field -variables [nondiscrete_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] +section nontrivially_normed_field +variables [nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] local notation `↑ₐ` := algebra_map 𝕜 A @[simp] lemma to_continuous_linear_map_norm [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : @@ -434,6 +434,6 @@ continuous_linear_map.op_norm_eq_of_bounds zero_le_one (λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (φ.apply_mem_spectrum _)) (λ _ _ h, by simpa only [to_continuous_linear_map_apply, mul_one, map_one, norm_one] using h 1) -end nondiscrete_normed_field +end nontrivially_normed_field end alg_hom diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index 36408415a2d67..4f45144673891 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -41,7 +41,7 @@ weak-* topology on (its type synonym) `weak_dual 𝕜 E`: coarser (not necessarily strictly) than the operator norm topology. * `weak_dual.is_compact_polar` (a version of the Banach-Alaoglu theorem): The polar set of a neighborhood of the origin in a normed space `E` over `𝕜` is compact in `weak_dual _ E`, if the - nondiscrete normed field `𝕜` is proper as a topological space. + nontrivially normed field `𝕜` is proper as a topological space. * `weak_dual.is_compact_closed_ball` (the most common special case of the Banach-Alaoglu theorem): Closed balls in the dual of a normed space `E` over `ℝ` or `ℂ` are compact in the weak-star topology. @@ -98,7 +98,7 @@ i.e., that the weak-* topology is coarser (not necessarily strictly) than the to by the dual-norm (i.e. the operator-norm). -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] namespace normed_space diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 8751ef9f6b49a..136d5a6db6d90 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -23,7 +23,7 @@ open filter asymptotics set function open_locale classical topological_space namespace complex -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_algebra 𝕜 ℂ] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 ℂ] /-- The complex exponential is everywhere differentiable, with the derivative `exp x`. -/ lemma has_deriv_at_exp (x : ℂ) : has_deriv_at exp (exp x) x := @@ -74,7 +74,7 @@ open_map_of_strict_deriv has_strict_deriv_at_exp exp_ne_zero end complex section -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_algebra 𝕜 ℂ] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 ℂ] {f : 𝕜 → ℂ} {f' : ℂ} {x : 𝕜} {s : set 𝕜} lemma has_strict_deriv_at.cexp (hf : has_strict_deriv_at f f' x) : @@ -101,7 +101,7 @@ hc.has_deriv_at.cexp.deriv end section -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_algebra 𝕜 ℂ] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 ℂ] {E : Type*} [normed_group E] [normed_space 𝕜 E] {f : E → ℂ} {f' : E →L[𝕜] ℂ} {x : E} {s : set E} diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index 7d4ffb38ec4fd..1fdf68e1062d6 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -49,7 +49,7 @@ open_locale nat topological_space big_operators ennreal section any_field_any_algebra -variables {𝕂 𝔸 : Type*} [nondiscrete_normed_field 𝕂] [normed_ring 𝔸] [normed_algebra 𝕂 𝔸] +variables {𝕂 𝔸 : Type*} [nontrivially_normed_field 𝕂] [normed_ring 𝔸] [normed_algebra 𝕂 𝔸] [complete_space 𝔸] /-- The exponential in a Banach-algebra `𝔸` over a normed field `𝕂` has strict Fréchet-derivative @@ -73,7 +73,7 @@ end any_field_any_algebra section any_field_comm_algebra -variables {𝕂 𝔸 : Type*} [nondiscrete_normed_field 𝕂] [normed_comm_ring 𝔸] [normed_algebra 𝕂 𝔸] +variables {𝕂 𝔸 : Type*} [nontrivially_normed_field 𝕂] [normed_comm_ring 𝔸] [normed_algebra 𝕂 𝔸] [complete_space 𝔸] /-- The exponential map in a commutative Banach-algebra `𝔸` over a normed field `𝕂` of @@ -110,7 +110,7 @@ end any_field_comm_algebra section deriv -variables {𝕂 : Type*} [nondiscrete_normed_field 𝕂] [complete_space 𝕂] +variables {𝕂 : Type*} [nontrivially_normed_field 𝕂] [complete_space 𝕂] /-- The exponential map in a complete normed field `𝕂` of characteristic zero has strict derivative `exp 𝕂 x` at any point `x` in the disk of convergence. -/ diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index e43964392760d..61de52b86a593 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -67,7 +67,7 @@ begin simpa using is_o.smul_is_O hε (hf.is_O_const (one_ne_zero : (1 : 𝕜) ≠ 0)) end -@[simp] lemma continuous_at_zpow {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {m : ℤ} {x : 𝕜} : +@[simp] lemma continuous_at_zpow {𝕜 : Type*} [nontrivially_normed_field 𝕜] {m : ℤ} {x : 𝕜} : continuous_at (λ x, x ^ m) x ↔ x ≠ 0 ∨ 0 ≤ m := begin refine ⟨_, continuous_at_zpow₀ _ _⟩, @@ -76,7 +76,7 @@ begin (tendsto_norm_zpow_nhds_within_0_at_top hm) end -@[simp] lemma continuous_at_inv {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {x : 𝕜} : +@[simp] lemma continuous_at_inv {𝕜 : Type*} [nontrivially_normed_field 𝕜] {x : 𝕜} : continuous_at has_inv.inv x ↔ x ≠ 0 := by simpa [(@zero_lt_one ℤ _ _).not_le] using @continuous_at_zpow _ _ (-1) x diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index 98b22ac51b55f..914b4a9ec76d2 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -47,7 +47,7 @@ open_locale complex_conjugate This typeclass captures properties shared by ℝ and ℂ, with an API that closely matches that of ℂ. -/ class is_R_or_C (K : Type*) - extends nondiscrete_normed_field K, star_ring K, normed_algebra ℝ K, complete_space K := + extends nontrivially_normed_field K, star_ring K, normed_algebra ℝ K, complete_space K := (re : K →+ ℝ) (im : K →+ ℝ) (I : K) -- Meant to be set to 0 for K=ℝ @@ -745,7 +745,7 @@ noncomputable instance real.is_R_or_C : is_R_or_C ℝ := 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.nondiscrete_normed_field, .. real.metric_space } + .. real.nontrivially_normed_field, .. real.metric_space } end instances diff --git a/src/geometry/manifold/algebra/left_invariant_derivation.lean b/src/geometry/manifold/algebra/left_invariant_derivation.lean index 69567e95ed6eb..869534128aedf 100644 --- a/src/geometry/manifold/algebra/left_invariant_derivation.lean +++ b/src/geometry/manifold/algebra/left_invariant_derivation.lean @@ -23,7 +23,7 @@ noncomputable theory open_locale lie_group manifold derivation -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) (G : Type*) [topological_space G] [charted_space H G] [monoid G] [has_smooth_mul I G] (g h : G) diff --git a/src/geometry/manifold/algebra/lie_group.lean b/src/geometry/manifold/algebra/lie_group.lean index f9a3408b9f97e..81d4dde0809c6 100644 --- a/src/geometry/manifold/algebra/lie_group.lean +++ b/src/geometry/manifold/algebra/lie_group.lean @@ -22,7 +22,7 @@ groups here are not necessarily finite dimensional. * `lie_add_group I G` : a Lie additive group where `G` is a manifold on the model with corners `I`. * `lie_group I G` : a Lie multiplicative group where `G` is a manifold on the model with corners `I`. -* `normed_space_lie_add_group` : a normed vector space over a nondiscrete normed field +* `normed_space_lie_add_group` : a normed vector space over a nontrivially normed field is an additive Lie group. ## Implementation notes @@ -44,7 +44,7 @@ open_locale manifold the addition and negation operations are smooth. -/ -- See note [Design choices about smooth algebraic structures] @[ancestor has_smooth_add] -class lie_add_group {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class lie_add_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [add_group G] [topological_space G] [charted_space H G] @@ -55,7 +55,7 @@ class lie_add_group {𝕜 : Type*} [nondiscrete_normed_field 𝕜] the multiplication and inverse operations are smooth. -/ -- See note [Design choices about smooth algebraic structures] @[ancestor has_smooth_mul, to_additive] -class lie_group {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class lie_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [group G] [topological_space G] [charted_space H G] @@ -64,7 +64,7 @@ class lie_group {𝕜 : Type*} [nondiscrete_normed_field 𝕜] section lie_group -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {F : Type*} [normed_group F] [normed_space 𝕜 F] {J : model_with_corners 𝕜 F F} @@ -121,7 +121,7 @@ section prod_lie_group /- Instance of product group -/ @[to_additive] -instance {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {H : Type*} [topological_space H] +instance {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [topological_space G] [charted_space H G] [group G] [lie_group I G] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] @@ -136,7 +136,7 @@ end prod_lie_group /-! ### Normed spaces are Lie groups -/ -instance normed_space_lie_add_group {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +instance normed_space_lie_add_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] : lie_add_group (𝓘(𝕜, E)) E := { smooth_add := smooth_iff.2 ⟨continuous_add, λ x y, cont_diff_add.cont_diff_on⟩, diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index 9681792fc7a08..3c0f0e31719ac 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -42,7 +42,7 @@ semigroup. A smooth additive monoid over `α`, for example, is obtained by requi instances `add_monoid α` and `has_smooth_add α`. -/ -- See note [Design choices about smooth algebraic structures] @[ancestor smooth_manifold_with_corners] -class has_smooth_add {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class has_smooth_add {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [has_add G] [topological_space G] [charted_space H G] @@ -54,7 +54,7 @@ A smooth monoid over `G`, for example, is obtained by requiring both the instanc and `has_smooth_mul I G`. -/ -- See note [Design choices about smooth algebraic structures] @[ancestor smooth_manifold_with_corners, to_additive] -class has_smooth_mul {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class has_smooth_mul {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [has_mul G] [topological_space G] [charted_space H G] @@ -63,7 +63,7 @@ class has_smooth_mul {𝕜 : Type*} [nondiscrete_normed_field 𝕜] section has_smooth_mul -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [has_mul G] [topological_space G] [charted_space H G] [has_smooth_mul I G] @@ -186,7 +186,7 @@ end /- Instance of product -/ @[to_additive] -instance has_smooth_mul.prod {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +instance has_smooth_mul.prod {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) (G : Type*) [topological_space G] [charted_space H G] @@ -204,7 +204,7 @@ end has_smooth_mul section monoid -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [monoid G] [topological_space G] [charted_space H G] [has_smooth_mul I G] @@ -248,7 +248,7 @@ section comm_monoid open_locale big_operators -variables {ι 𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {ι 𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [comm_monoid G] [topological_space G] [charted_space H G] [has_smooth_mul I G] diff --git a/src/geometry/manifold/algebra/smooth_functions.lean b/src/geometry/manifold/algebra/smooth_functions.lean index b4cdbb28b8341..8e5669a3fe367 100644 --- a/src/geometry/manifold/algebra/smooth_functions.lean +++ b/src/geometry/manifold/algebra/smooth_functions.lean @@ -16,7 +16,7 @@ noncomputable theory open_locale manifold -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} diff --git a/src/geometry/manifold/algebra/structures.lean b/src/geometry/manifold/algebra/structures.lean index 63fa459321fbc..2ab748bf19d44 100644 --- a/src/geometry/manifold/algebra/structures.lean +++ b/src/geometry/manifold/algebra/structures.lean @@ -15,7 +15,7 @@ instead of Lie mainly because Lie ring has currently another use in mathematics. open_locale manifold section smooth_ring -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] {E : Type*} [normed_group E] [normed_space 𝕜 E] @@ -42,7 +42,7 @@ instance smooth_ring.to_lie_add_group (I : model_with_corners 𝕜 E H) end smooth_ring -instance field_smooth_ring {𝕜 : Type*} [nondiscrete_normed_field 𝕜] : +instance field_smooth_ring {𝕜 : Type*} [nontrivially_normed_field 𝕜] : smooth_ring 𝓘(𝕜) 𝕜 := { smooth_mul := begin @@ -55,7 +55,7 @@ instance field_smooth_ring {𝕜 : Type*} [nondiscrete_normed_field 𝕜] : ..normed_space_lie_add_group } variables {𝕜 R E H : Type*} [topological_space R] [topological_space H] - [nondiscrete_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] + [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [charted_space H R] (I : model_with_corners 𝕜 E H) /-- A smooth (semi)ring is a topological (semi)ring. This is not an instance for technical reasons, diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 32daeb4ff4cdd..92239f4f4db43 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -52,7 +52,7 @@ open_locale topological_space manifold /-! ### Definition of smooth functions between manifolds -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -- declare a smooth manifold `M` over the pair `(E, H)`. {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 8fa29bf4479e5..925680dc95af7 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -14,7 +14,7 @@ In this file we define the type `cont_mdiff_map` of `n` times continuously diffe bundled maps. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {H : Type*} [topological_space H] diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 373ca53b09187..c0ac8c6f4b208 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -20,7 +20,7 @@ of the Lie algebra for a Lie group. -/ -variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) (M : Type*) [topological_space M] [charted_space H M] (n : with_top ℕ) diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 4427d841552c6..1ed56603bd220 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -46,7 +46,7 @@ diffeomorphism, manifold open_locale manifold topological_space open function set -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {F : Type*} [normed_group F] [normed_space 𝕜 F] diff --git a/src/geometry/manifold/instances/units_of_normed_algebra.lean b/src/geometry/manifold/instances/units_of_normed_algebra.lean index f80a4a3fe5c6b..77bfb586b575e 100644 --- a/src/geometry/manifold/instances/units_of_normed_algebra.lean +++ b/src/geometry/manifold/instances/units_of_normed_algebra.lean @@ -57,7 +57,7 @@ instance : charted_space R Rˣ := open_embedding_coe.singleton_charted_space lemma chart_at_apply {a : Rˣ} {b : Rˣ} : chart_at R a b = b := rfl lemma chart_at_source {a : Rˣ} : (chart_at R a).source = set.univ := rfl -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_algebra 𝕜 R] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 R] instance : smooth_manifold_with_corners 𝓘(𝕜, R) Rˣ := open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(𝕜, R) diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 6435272cf0b0c..02735f0b81395 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -110,7 +110,7 @@ this specific chart. We use the names `mdifferentiable` and `mfderiv`, where the prefix letter `m` means "manifold". -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -239,7 +239,7 @@ end derivatives_definitions section derivatives_properties /-! ### Unique differentiability sets in manifolds -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] -- @@ -853,7 +853,7 @@ manifold structure, coincides with the usual Frechet derivative `fderiv`. In thi this and related statements. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {f : E → E'} {s : set E} {x : E} @@ -959,7 +959,7 @@ section specific_functions /-! ### Differentiability of specific functions -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -1263,7 +1263,7 @@ end specific_functions /-! ### Differentiable local homeomorphisms -/ namespace local_homeomorph.mdifferentiable -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -1375,7 +1375,7 @@ end local_homeomorph.mdifferentiable section ext_chart_at -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -1402,7 +1402,7 @@ end ext_chart_at /-! ### Unique derivative sets in manifolds -/ section unique_mdiff -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 8c6bf28da0a37..4106ac16e6aa0 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -126,7 +126,7 @@ model vector space `E` over the field `𝕜`. This is all what is needed to define a smooth manifold with model space `H`, and model vector space `E`. -/ @[nolint has_inhabited_instance] -structure model_with_corners (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +structure model_with_corners (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] extends local_equiv H E := (source_eq : source = univ) @@ -137,7 +137,7 @@ structure model_with_corners (𝕜 : Type*) [nondiscrete_normed_field 𝕜] attribute [simp, mfld_simps] model_with_corners.source_eq /-- A vector space is a model with corners. -/ -def model_with_corners_self (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +def model_with_corners_self (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] : model_with_corners 𝕜 E E := { to_local_equiv := local_equiv.refl E, source_eq := rfl, @@ -150,7 +150,7 @@ localized "notation `𝓘(` 𝕜 `, ` E `)` := model_with_corners_self 𝕜 E" i localized "notation `𝓘(` 𝕜 `)` := model_with_corners_self 𝕜 𝕜" in manifold section -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) @@ -163,12 +163,12 @@ protected def symm : local_equiv E H := I.to_local_equiv.symm /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ -def simps.apply (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +def simps.apply (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] (I : model_with_corners 𝕜 E H) : H → E := I /-- See Note [custom simps projection] -/ -def simps.symm_apply (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +def simps.symm_apply (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] (I : model_with_corners 𝕜 E H) : E → H := I.symm @@ -305,7 +305,7 @@ structure on the tangent bundle to a manifold modelled on `(E, H)`: it will be m `(E × E, H × E)`. See note [Manifold type tags] for explanation about `model_prod H H'` vs `H × H'`. -/ @[simps (lemmas_only)] def model_with_corners.prod - {𝕜 : Type u} [nondiscrete_normed_field 𝕜] + {𝕜 : Type u} [nontrivially_normed_field 𝕜] {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] (I : model_with_corners 𝕜 E H) {E' : Type v'} [normed_group E'] [normed_space 𝕜 E'] {H' : Type w'} [topological_space H'] @@ -323,7 +323,7 @@ vs `H × H'`. -/ corners `pi I` on `(Π i, E i, model_pi H)`. See note [Manifold type tags] for explanation about `model_pi H`. -/ def model_with_corners.pi - {𝕜 : Type u} [nondiscrete_normed_field 𝕜] {ι : Type v} [fintype ι] + {𝕜 : Type u} [nontrivially_normed_field 𝕜] {ι : Type v} [fintype ι] {E : ι → Type w} [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] {H : ι → Type u'} [Π i, topological_space (H i)] (I : Π i, model_with_corners 𝕜 (E i) (H i)) : model_with_corners 𝕜 (Π i, E i) (model_pi H) := @@ -336,12 +336,12 @@ def model_with_corners.pi /-- Special case of product model with corners, which is trivial on the second factor. This shows up as the model to tangent bundles. -/ @[reducible] def model_with_corners.tangent - {𝕜 : Type u} [nondiscrete_normed_field 𝕜] + {𝕜 : Type u} [nontrivially_normed_field 𝕜] {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] (I : model_with_corners 𝕜 E H) : model_with_corners 𝕜 (E × E) (model_prod H E) := I.prod (𝓘(𝕜, E)) -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {F : Type*} [normed_group F] [normed_space 𝕜 F] {F' : Type*} [normed_group F'] [normed_space 𝕜 F'] {H : Type*} [topological_space H] {H' : Type*} [topological_space H'] @@ -365,18 +365,18 @@ end model_with_corners_prod section boundaryless /-- Property ensuring that the model with corners `I` defines manifolds without boundary. -/ -class model_with_corners.boundaryless {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class model_with_corners.boundaryless {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) : Prop := (range_eq_univ : range I = univ) /-- The trivial model with corners has no boundary -/ -instance model_with_corners_self_boundaryless (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +instance model_with_corners_self_boundaryless (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_group E] [normed_space 𝕜 E] : (model_with_corners_self 𝕜 E).boundaryless := ⟨by simp⟩ /-- If two model with corners are boundaryless, their product also is -/ -instance model_with_corners.range_eq_univ_prod {𝕜 : Type u} [nondiscrete_normed_field 𝕜] +instance model_with_corners.range_eq_univ_prod {𝕜 : Type u} [nontrivially_normed_field 𝕜] {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] (I : model_with_corners 𝕜 E H) [I.boundaryless] {E' : Type v'} [normed_group E'] [normed_space 𝕜 E'] {H' : Type w'} [topological_space H'] @@ -394,7 +394,7 @@ end boundaryless section cont_diff_groupoid /-! ### Smooth functions on models with corners -/ -variables {m n : with_top ℕ} {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {m n : with_top ℕ} {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) @@ -549,13 +549,13 @@ section smooth_manifold_with_corners /-- Typeclass defining smooth manifolds with corners with respect to a model with corners, over a field `𝕜` and with infinite smoothness to simplify typeclass search and statements later on. -/ @[ancestor has_groupoid] -class smooth_manifold_with_corners {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +class smooth_manifold_with_corners {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] extends has_groupoid M (cont_diff_groupoid ∞ I) : Prop -lemma smooth_manifold_with_corners.mk' {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +lemma smooth_manifold_with_corners.mk' {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -563,7 +563,7 @@ lemma smooth_manifold_with_corners.mk' {𝕜 : Type*} [nondiscrete_normed_field smooth_manifold_with_corners I M := { ..gr } lemma smooth_manifold_with_corners_of_cont_diff_on - {𝕜 : Type*} [nondiscrete_normed_field 𝕜] + {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -578,7 +578,7 @@ lemma smooth_manifold_with_corners_of_cont_diff_on end } /-- For any model with corners, the model space is a smooth manifold -/ -instance model_space_smooth {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +instance model_space_smooth {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} : smooth_manifold_with_corners I H := { .. has_groupoid_model_space _ _ } @@ -590,7 +590,7 @@ namespace smooth_manifold_with_corners charted space with a structure groupoid, avoiding the need to specify the groupoid `cont_diff_groupoid ∞ I` explicitly. -/ -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -617,7 +617,7 @@ lemma compatible_of_mem_maximal_atlas structure_groupoid.compatible_of_mem_maximal_atlas he he' /-- The product of two smooth manifolds with corners is naturally a smooth manifold with corners. -/ -instance prod {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +instance prod {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} @@ -637,7 +637,7 @@ instance prod {𝕜 : Type*} [nondiscrete_normed_field 𝕜] end smooth_manifold_with_corners lemma local_homeomorph.singleton_smooth_manifold_with_corners - {𝕜 : Type*} [nondiscrete_normed_field 𝕜] + {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) {M : Type*} [topological_space M] @@ -647,7 +647,7 @@ lemma local_homeomorph.singleton_smooth_manifold_with_corners e.singleton_has_groupoid h (cont_diff_groupoid ∞ I) lemma open_embedding.singleton_smooth_manifold_with_corners - {𝕜 : Type*} [nondiscrete_normed_field 𝕜] + {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) {M : Type*} [topological_space M] @@ -659,7 +659,7 @@ namespace topological_space.opens open topological_space -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -672,7 +672,7 @@ end topological_space.opens section extended_charts open_locale topological_space -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -891,7 +891,7 @@ end extended_charts /-- In the case of the manifold structure on a vector space, the extended charts are just the identity.-/ -lemma ext_chart_model_space_eq_id (𝕜 : Type*) [nondiscrete_normed_field 𝕜] +lemma ext_chart_model_space_eq_id (𝕜 : Type*) [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] (x : E) : ext_chart_at (model_with_corners_self 𝕜 E) x = local_equiv.refl E := by simp only with mfld_simps diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean index 17af8285ad06a..c2fca4e33bfe7 100644 --- a/src/geometry/manifold/tangent_bundle.lean +++ b/src/geometry/manifold/tangent_bundle.lean @@ -83,7 +83,7 @@ corner `I`) with fiber the normed vector space `F` over `𝕜`, which is trivial 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*} [nondiscrete_normed_field 𝕜] +structure basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -98,7 +98,7 @@ structure basic_smooth_vector_bundle_core {𝕜 : Type*} [nondiscrete_normed_fie /-- The trivial basic smooth bundle core, in which all the changes of coordinates are the identity. -/ -def trivial_basic_smooth_vector_bundle_core {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +def trivial_basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -110,7 +110,7 @@ def trivial_basic_smooth_vector_bundle_core {𝕜 : Type*} [nondiscrete_normed_f namespace basic_smooth_vector_bundle_core -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] @@ -289,7 +289,7 @@ end basic_smooth_vector_bundle_core section tangent_bundle -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_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] diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 55162094112d9..50cde7b97745a 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -2068,7 +2068,7 @@ end continuous_linear_map namespace continuous_linear_map -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] @@ -2094,9 +2094,9 @@ measurable_pi_lambda _ measurable_apply end continuous_linear_map -section continuous_linear_map_nondiscrete_normed_field +section continuous_linear_map_nontrivially_normed_field -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E] [borel_space E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] @@ -2110,10 +2110,10 @@ lemma ae_measurable.apply_continuous_linear_map {φ : α → F →L[𝕜] E} {μ (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_nondiscrete_normed_field +end continuous_linear_map_nontrivially_normed_field section normed_space -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [complete_space 𝕜] [measurable_space 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] [measurable_space 𝕜] variables [borel_space 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E] [borel_space E] diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 8b6c841d5e12a..bd0ba3038840b 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -906,7 +906,7 @@ by simp_rw [div_eq_mul_inv, h.mul_const] end normed_space section normed_space_over_complete_field -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [complete_space 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] lemma integrable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : @@ -1187,7 +1187,7 @@ end measure_theory open measure_theory variables {E : Type*} [normed_group E] - {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] + {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] {H : Type*} [normed_group H] [normed_space 𝕜 H] lemma measure_theory.integrable.apply_continuous_linear_map {φ : α → H →L[𝕜] E} diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 1835606eb0307..f567e377dd856 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1981,7 +1981,7 @@ lemma continuous_comp_Lp [fact (1 ≤ p)] (hg : lipschitz_with c g) (g0 : g 0 = end lipschitz_with namespace continuous_linear_map -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] /-- Composing `f : Lp ` with `L : E →L[𝕜] F`. -/ def comp_Lp (L : E →L[𝕜] F) (f : Lp E p μ) : Lp F p μ := @@ -2686,7 +2686,7 @@ lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] (f to_Lp p μ 𝕜 f =ᵐ[μ] f := ae_eq_fun.coe_fn_mk f _ -lemma to_Lp_norm_le [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : +lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : ∥(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 _ @@ -2739,7 +2739,7 @@ rfl (to_Lp p μ 𝕜 f : α →ₘ[μ] E) = f.to_ae_eq_fun μ := rfl -variables [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +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/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index c5ef42c810a2e..df026a9299c01 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -1511,7 +1511,7 @@ lemma smul_measure {R : Type*} [monoid R] [distrib_mul_action R ℝ≥0∞] ⟨h.mk f, h.strongly_measurable_mk, ae_smul_measure h.ae_eq_mk c⟩ section normed_space -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [complete_space 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] lemma _root_.ae_strongly_measurable_smul_const_iff {f : α → 𝕜} {c : E} (hc : c ≠ 0) : @@ -1542,9 +1542,9 @@ end end mul_action -section continuous_linear_map_nondiscrete_normed_field +section continuous_linear_map_nontrivially_normed_field -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] variables {G : Type*} [normed_group G] [normed_space 𝕜 G] @@ -1565,7 +1565,7 @@ lemma _root_.continuous_linear_map.ae_strongly_measurable_comp₂ (L : E →L[ ae_strongly_measurable (λ x, L (f x) (g x)) μ := L.continuous₂.comp_ae_strongly_measurable $ hf.prod_mk hg -end continuous_linear_map_nondiscrete_normed_field +end continuous_linear_map_nontrivially_normed_field lemma _root_.ae_strongly_measurable_with_density_iff {E : Type*} [normed_group E] [normed_space ℝ E] {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 504473b3158e3..8cde6c4a3b66d 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -587,7 +587,7 @@ open simple_func local notation `Integral` := @integral_clm α E _ _ _ _ _ μ _ -variables [normed_space ℝ E] [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +variables [normed_space ℝ E] [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] [normed_space ℝ F] [complete_space E] section integration_in_L1 @@ -687,7 +687,7 @@ functions, and 0 otherwise; prove its basic properties. -/ variables [normed_group E] [normed_space ℝ E] [complete_space E] - [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] + [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] [normed_group F] [normed_space ℝ F] [complete_space F] section diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 16beeba903c47..657ea7c50ef5e 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -628,7 +628,7 @@ by { simp only [interval_integral, integral_neg], abel } ∫ x in a..b, f x - g x ∂μ = ∫ x in a..b, f x ∂μ - ∫ x in a..b, g x ∂μ := by simpa only [sub_eq_add_neg] using (integral_add hf hg.neg).trans (congr_arg _ integral_neg) -@[simp] lemma integral_smul {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +@[simp] lemma integral_smul {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] (r : 𝕜) (f : ℝ → E) : ∫ x in a..b, r • f x ∂μ = r • ∫ x in a..b, f x ∂μ := by simp only [interval_integral, integral_smul, smul_sub] diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index bf85324211a74..3987e9b4686c5 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -1032,7 +1032,7 @@ section set_to_L1 local attribute [instance] Lp.simple_func.module local attribute [instance] Lp.simple_func.normed_space -variables (𝕜) [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +variables (𝕜) [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] [complete_space F] {T T' T'' : set α → E →L[ℝ] F} {C C' C'' : ℝ} @@ -1457,7 +1457,7 @@ lemma set_to_fun_sub (hT : dominated_fin_meas_additive μ T C) set_to_fun μ T hT (f - g) = set_to_fun μ T hT f - set_to_fun μ T hT g := by rw [sub_eq_add_neg, sub_eq_add_neg, set_to_fun_add hT hf hg.neg, set_to_fun_neg hT g] -lemma set_to_fun_smul [nondiscrete_normed_field 𝕜] +lemma set_to_fun_smul [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] (hT : dominated_fin_meas_additive μ T C) (h_smul : ∀ c : 𝕜, ∀ s x, T s (c • x) = c • T s x) (c : 𝕜) (f : α → E) : set_to_fun μ T hT (c • f) = c • set_to_fun μ T hT f := diff --git a/src/measure_theory/measure/with_density_vector_measure.lean b/src/measure_theory/measure/with_density_vector_measure.lean index 711693e337091..b67c51df4c2b9 100644 --- a/src/measure_theory/measure/with_density_vector_measure.lean +++ b/src/measure_theory/measure/with_density_vector_measure.lean @@ -100,7 +100,7 @@ lemma with_densityᵥ_sub' (hf : integrable f μ) (hg : integrable g μ) : μ.with_densityᵥ (λ x, f x - g x) = μ.with_densityᵥ f - μ.with_densityᵥ g := with_densityᵥ_sub hf hg -@[simp] lemma with_densityᵥ_smul {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +@[simp] lemma with_densityᵥ_smul {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] (f : α → E) (r : 𝕜) : μ.with_densityᵥ (r • f) = r • μ.with_densityᵥ f := begin @@ -115,7 +115,7 @@ begin rwa integrable_smul_iff hr f } } end -lemma with_densityᵥ_smul' {𝕜 : Type*} [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] +lemma with_densityᵥ_smul' {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] (f : α → E) (r : 𝕜) : μ.with_densityᵥ (λ x, r • f x) = r • μ.with_densityᵥ f := with_densityᵥ_smul f r diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index 74c65a8d93439..83f1fca5f7df0 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -784,7 +784,7 @@ end @[simp] lemma norm_p_pow (n : ℤ) : ∥(p^n : ℚ_[p])∥ = p^-n := by rw [norm_zpow, norm_p]; field_simp -instance : nondiscrete_normed_field ℚ_[p] := +instance : nontrivially_normed_field ℚ_[p] := { non_trivial := ⟨p⁻¹, begin rw [norm_inv, norm_p, inv_inv], exact_mod_cast hp.1.one_lt diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index bb3892b157efd..e9f71739575ee 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -83,7 +83,7 @@ variables {R : Type*} {M : Type*} /-- If `M` is a topological module over `R` and `0` is a limit of invertible elements of `R`, then `⊤` is the only submodule of `M` with a nonempty interior. -This is the case, e.g., if `R` is a nondiscrete normed field. -/ +This is the case, e.g., if `R` is a nontrivially normed field. -/ lemma submodule.eq_top_of_nonempty_interior' [ne_bot (𝓝[{x : R | is_unit x}] 0)] (s : submodule R M) (hs : (interior (s:set M)).nonempty) : @@ -104,7 +104,7 @@ end variables (R M) -/-- Let `R` be a topological ring such that zero is not an isolated point (e.g., a nondiscrete +/-- Let `R` be a topological ring such that zero is not an isolated point (e.g., a nontrivially normed field, see `normed_field.punctured_nhds_ne_bot`). Let `M` be a nontrivial module over `R` such that `c • x = 0` implies `c = 0 ∨ x = 0`. Then `M` has no isolated points. We formulate this using `ne_bot (𝓝[≠] x)`. diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 8f03f323b180f..32c6a25a50b19 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -8,7 +8,7 @@ import analysis.locally_convex.balanced_core_hull /-! # Finite dimensional topological vector spaces over complete fields -Let `𝕜` be a nondiscrete and complete normed field, and `E` a topological vector space (TVS) over +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]`). @@ -85,7 +85,7 @@ end field section normed_field -variables {𝕜 : Type u} [hnorm : nondiscrete_normed_field 𝕜] +variables {𝕜 : Type u} [hnorm : nontrivially_normed_field 𝕜] {E : Type v} [add_comm_group E] [module 𝕜 E] [topological_space E] [topological_add_group E] [has_continuous_smul 𝕜 E] {F : Type w} [add_comm_group F] [module 𝕜 F] [topological_space F] @@ -95,8 +95,8 @@ variables {𝕜 : Type u} [hnorm : nondiscrete_normed_field 𝕜] include hnorm -/-- If `𝕜` is a nondiscrete normed field, any T2 topology on `𝕜` which makes it a topological vector - space over itself (with the norm topology) is *equal* to the norm topology. -/ +/-- If `𝕜` is a nontrivially normed field, any T2 topology on `𝕜` which makes it a topological +vector space over itself (with the norm topology) is *equal* to the norm topology. -/ lemma unique_topology_of_t2 {t : topological_space 𝕜} (h₁ : @topological_add_group 𝕜 t _) (h₂ : @has_continuous_smul 𝕜 𝕜 _ hnorm.to_uniform_space.to_topological_space t) @@ -109,7 +109,7 @@ begin refine topological_add_group.ext h₁ infer_instance (le_antisymm _ _), { -- To show `𝓣 ≤ 𝓣₀`, we have to show that closed balls are `𝓣`-neighborhoods of 0. rw metric.nhds_basis_closed_ball.ge_iff, - -- Let `ε > 0`. Since `𝕜` is nondiscrete, we have `0 < ∥ξ₀∥ < ε` for some `ξ₀ : 𝕜`. + -- Let `ε > 0`. Since `𝕜` is nontrivially normed, we have `0 < ∥ξ₀∥ < ε` for some `ξ₀ : 𝕜`. intros ε hε, rcases normed_field.exists_norm_lt 𝕜 hε with ⟨ξ₀, hξ₀, hξ₀ε⟩, -- Since `ξ₀ ≠ 0` and `𝓣` is T2, we know that `{ξ₀}ᶜ` is a `𝓣`-neighborhood of 0. @@ -147,7 +147,7 @@ begin ... = (@nhds 𝕜 t 0) : by rw zero_smul } end -/-- Any linear form on a topological vector space over a nondiscrete normed field is continuous if +/-- Any linear form on a topological vector space over a nontrivially normed field is continuous if its kernel is closed. -/ lemma linear_map.continuous_of_is_closed_ker (l : E →ₗ[𝕜] 𝕜) (hl : is_closed (l.ker : set E)) : continuous l := @@ -191,7 +191,7 @@ begin exact continuous_coinduced_rng } end -/-- Any linear form on a topological vector space over a nondiscrete normed field is continuous if +/-- Any linear form on a topological vector space over a nontrivially normed field is continuous if and only if its kernel is closed. -/ lemma linear_map.continuous_iff_is_closed_ker (l : E →ₗ[𝕜] 𝕜) : continuous l ↔ is_closed (l.ker : set E) := diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 282e1a6ce4625..6f98349849b8b 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -987,7 +987,7 @@ instance [normed_field 𝕜] [normed_space 𝕜 β] : normed_space 𝕜 (α → exact (λ x, trans_rel_right _ (norm_smul _ _) (mul_le_mul_of_nonneg_left (f.norm_coe_le_norm _) (norm_nonneg _))) end⟩ -variables [nondiscrete_normed_field 𝕜] [normed_space 𝕜 β] +variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 β] variables [semi_normed_group γ] [normed_space 𝕜 γ] variables (α) diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 7f0c129b47e09..36df30b9a6952 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -307,7 +307,7 @@ end continuous_map section comp_left variables (X : Type*) {𝕜 β γ : Type*} [topological_space X] [compact_space X] - [nondiscrete_normed_field 𝕜] + [nontrivially_normed_field 𝕜] variables [normed_group β] [normed_space 𝕜 β] [normed_group γ] [normed_space 𝕜 γ] open continuous_map diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index ba277a6b4c64b..45712811344d6 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -402,7 +402,7 @@ end isometric namespace continuous_linear_equiv -variables {𝕜 E F : Type*} [nondiscrete_normed_field 𝕜] +variables {𝕜 E F : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] @[simp] lemma dimH_image (e : E ≃L[𝕜] F) (s : set E) : dimH (e '' s) = dimH s := diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 454a74ea48f20..95b5a4bbe7375 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -480,7 +480,7 @@ section open topological_vector_bundle variables (B) -variables [nondiscrete_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] +variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [normed_group F] [normed_space R F] [topological_space B] [topological_space (total_space E)] [∀ x, topological_space (E x)] @@ -918,7 +918,7 @@ end /-! ### Topological vector prebundle -/ section -variables [nondiscrete_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] +variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [normed_group F] [normed_space R F] [topological_space B] open topological_space diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index ddd1541339077..c9fea76cf57a3 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -70,7 +70,7 @@ by delta_instance bundle.continuous_linear_map end defs -variables {𝕜₁ : Type*} [nondiscrete_normed_field 𝕜₁] {𝕜₂ : Type*} [nondiscrete_normed_field 𝕜₂] +variables {𝕜₁ : Type*} [nontrivially_normed_field 𝕜₁] {𝕜₂ : Type*} [nontrivially_normed_field 𝕜₂] (σ : 𝕜₁ →+* 𝕜₂) variables {B : Type*} [topological_space B] diff --git a/src/topology/vector_bundle/prod.lean b/src/topology/vector_bundle/prod.lean index b0f43c8ed41d7..dc25c0f6daf91 100644 --- a/src/topology/vector_bundle/prod.lean +++ b/src/topology/vector_bundle/prod.lean @@ -56,7 +56,7 @@ lemma prod.inducing_diag : inducing end defs -variables [nondiscrete_normed_field R] [topological_space B] +variables [nontrivially_normed_field R] [topological_space B] variables (F₁ : Type*) [normed_group F₁] [normed_space R F₁] (E₁ : B → Type*) [topological_space (total_space E₁)] diff --git a/src/topology/vector_bundle/pullback.lean b/src/topology/vector_bundle/pullback.lean index 9a2ca6ee1906b..d57ef9165af11 100644 --- a/src/topology/vector_bundle/pullback.lean +++ b/src/topology/vector_bundle/pullback.lean @@ -65,7 +65,7 @@ begin refl end -variables (F) [nondiscrete_normed_field 𝕜] +variables (F) [nontrivially_normed_field 𝕜] [normed_group F] [normed_space 𝕜 F] [topological_space B] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] From a78156aa3deb611e8e90925ff09a6acdfb0dc696 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 23 Jul 2022 19:01:00 +0000 Subject: [PATCH 019/173] fix(*): fix typos I made yesterday (#15627) Fix 4 typos I made yesterday. --- src/data/set/intervals/ord_connected.lean | 2 +- src/data/set/prod.lean | 2 +- src/topology/algebra/module/finite_dimension.lean | 4 ++-- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 3253552bf4ae9..358937c1d2b2c 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -169,7 +169,7 @@ lemma ord_connected.interval_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s [x, y] ⊆ s := hs.out (min_rec' (∈ s) hx hy) (max_rec' (∈ s) hx hy) -lemma ord_interval.interval_oc_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s) : +lemma ord_connected.interval_oc_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 diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index a73543ee67207..aceb213dab742 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -327,7 +327,7 @@ subset_compl_comm.trans $ by simp_rw [← range_diag, range_subset_iff, @[simp] lemma diag_preimage_prod (s t : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ t) = s ∩ t := rfl -lemma diag_preimage_prod_sellf (s : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ s) = s := inter_self s +lemma diag_preimage_prod_self (s : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ s) = s := inter_self s end diagonal diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 32c6a25a50b19..d4d7c22e7667b 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -313,11 +313,11 @@ def to_continuous_linear_map : (E →ₗ[𝕜] F') ≃ₗ[𝕜] E →L[𝕜] F' f.to_continuous_linear_map.det = f.det := rfl -@[simp] lemma ker_to_continuous_linear_map (f : E →ₗ[𝕜] E) : +@[simp] lemma ker_to_continuous_linear_map (f : E →ₗ[𝕜] F') : f.to_continuous_linear_map.ker = f.ker := rfl -@[simp] lemma range_to_continuous_linear_map (f : E →ₗ[𝕜] E) : +@[simp] lemma range_to_continuous_linear_map (f : E →ₗ[𝕜] F') : f.to_continuous_linear_map.range = f.range := rfl From 8e24d4fb879c32f947abe234a8132b48f3d41cea Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Jul 2022 21:26:58 +0000 Subject: [PATCH 020/173] feat(number_theory/legendre_symbol/add_character): add file introducing additive characters (#15499) This adds a file that defines additive characters on commutative rings and proves some relevant properties, e.g., that on finite fields and rings of type `zmod n`, there exists a *primitive* additive character, and that the sum of the character values vanishes when the character is nontrivial. This will be used in a later PR (together with multiplicative characters) to define Gauss sums and prove some resuits on them. There is a [Zulip topic](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Additive.20characters/near/290005886) to discuss this. --- .../legendre_symbol/add_character.lean | 320 ++++++++++++++++++ 1 file changed, 320 insertions(+) create mode 100644 src/number_theory/legendre_symbol/add_character.lean diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean new file mode 100644 index 0000000000000..e94d147983708 --- /dev/null +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -0,0 +1,320 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import tactic.basic +import field_theory.finite.galois_field +import number_theory.cyclotomic.primitive_roots +import field_theory.finite.trace + +/-! +# Additive characters of finite rings and fields + +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'`. + +We use the namespace `add_char`. + +## Main definitions and results + +We define `mul_shift ψ a`, where `ψ : add_char R R'` and `a : R`, to be the +character defined by `x ↦ ψ (a * x)`. An additive character `ψ` is *primitive* +if `mul_shift ψ a` is trivial only when `a = 0`. + +We show that when `ψ` is primitive, then the map `a ↦ mul_shift ψ a` is injective +(`add_char.to_mul_shift_inj_of_is_primitive`) and that `ψ` is primitive when `R` is a field +and `ψ` is nontrivial (`add_char.is_primitive_of_is_nontrivial`). + +We also show that there are primitive additive characters on `R` (with suitable +target `R'`) when `R` is a field or `R = zmod n` (`add_char.primitive_char_finite_field` +and `add_char.primitive_zmod_char`). + +Finally, we show that the sum of all character values is zero when the character +is nontrivial (and the target is a domain); see `add_char.sum_eq_zero_of_is_nontrivial`. + +## Tags + +additive character +-/ + +/-! +### Definitions related to and results on additive characters +-/ + +section add_char_def + +universes u v + +-- The domain of our additive characters +variables (R : Type u) [add_monoid R] +-- The target +variables (R' : Type v) [monoid R'] + +/-- Define `add_char R R'` as `(multiplicative R) →* R'`. +The definition works for an additive monoid `R` and a monoid `R'`, +but we will restrict to the case that both are commutative rings below. +The trivial additive character (sending everything to `1`) is `(1 : add_char R R').` -/ +abbreviation add_char : Type (max u v) := (multiplicative R) →* R' + +end add_char_def + + +section additive + +namespace add_char + +open multiplicative + +universes u v + +-- The domain of our additive characters +variables {R : Type u} [comm_ring R] +-- The target +variables {R' : Type v} [comm_ring R'] + +/-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ +def is_nontrivial (ψ : add_char R R') : Prop := ∃ (a : R), ψ (of_add a) ≠ 1 + +/-- An additive character is nontrivial iff it is not the trivial character. -/ +lemma is_nontrivial_iff_ne_trivial (ψ : add_char R R') : is_nontrivial ψ ↔ ψ ≠ 1 := +begin + refine not_forall.symm.trans (iff.not _), + rw fun_like.ext_iff, + refl, +end + +/-- Define the multiplicative shift of an additive character. +This satisfies `mul_shift ψ a x = ψ (a * x)`. -/ +def mul_shift (ψ : add_char R R') (a : R) : add_char R R' := +ψ.comp (add_monoid_hom.mul_left a).to_multiplicative + +@[simp] lemma mul_shift_apply {ψ : add_char R R'} {a : R} {x : multiplicative R} : + mul_shift ψ a x = ψ (of_add (a * to_add x)) := rfl + +/-- If `n` is a natural number, then `mul_shift ψ n x = (ψ x) ^ n`. -/ +lemma mul_shift_spec' (ψ : add_char R R') (n : ℕ) (x : multiplicative R) : + mul_shift ψ n x = (ψ x) ^ n := +by rw [mul_shift_apply, ← nsmul_eq_mul, of_add_nsmul, map_pow, of_add_to_add] + +/-- The product of `mul_shift ψ a` and `mul_shift ψ b` is `mul_shift ψ (a + b)`. -/ +lemma mul_shift_mul (ψ : add_char R R') (a b : R) : + mul_shift ψ a * mul_shift ψ b = mul_shift ψ (a + b) := +begin + ext, + simp only [monoid_hom.mul_apply, mul_shift_apply, add_mul, of_add_add, ψ.map_mul], +end + +/-- `mul_shift ψ 0` is the trivial character. -/ +@[simp] +lemma mul_shift_zero (ψ : add_char R R') : mul_shift ψ 0 = 1 := +begin + ext, + simp only [mul_shift_apply, zero_mul, of_add_zero, map_one, monoid_hom.one_apply], +end + +/-- An additive character is *primitive* iff all its multiplicative shifts by nonzero +elements are nontrivial. -/ +def is_primitive (ψ : add_char R R') : Prop := +∀ (a : R), a ≠ 0 → is_nontrivial (mul_shift ψ a) + +/-- The map associating to `a : R` the multiplicative shift of `ψ` by `a` +is injective when `ψ` is primitive. -/ +lemma to_mul_shift_inj_of_is_primitive {ψ : add_char R R'} (hψ : is_primitive ψ) : + function.injective ψ.mul_shift := +begin + intros a b h, + apply_fun (λ x, x * mul_shift ψ (-b)) at h, + simp only [mul_shift_mul, mul_shift_zero, add_right_neg] at h, + have h₂ := hψ (a + (- b)), + rw [h, is_nontrivial_iff_ne_trivial, ← sub_eq_add_neg, sub_ne_zero] at h₂, + exact not_not.mp (λ h, h₂ h rfl), +end + +-- `add_comm_group.equiv_direct_sum_zmod_of_fintype` +-- gives the structure theorem for finite abelian groups. +-- This could be used to show that the map above is a bijection. +-- We leave this for a later occasion. + +/-- When `R` is a field `F`, then a nontrivial additive character is primitive -/ +lemma is_nontrivial.is_primitive {F : Type u} [field F] + (ψ : add_char F R') (hψ : is_nontrivial ψ) : + is_primitive ψ := +begin + intros a ha, + cases hψ with x h, + use (a⁻¹ * x), + rwa [mul_shift_apply, to_add_of_add, mul_inv_cancel_left₀ ha], +end + +/-- Structure 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. -/ +@[nolint has_inhabited_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 : pnat) +(char : add_char R (cyclotomic_field n R')) +(prim : is_primitive char) + +/-! +### Additive characters on `zmod n` +-/ + +variables {C : Type v} [comm_ring C] + +/-- We can define an additive character on `zmod n` when we have an `n`th root of unity `ζ : C`. -/ +def zmod_char (n : ℕ) [fact (0 < n)] {ζ : C} (hζ : ζ ^ n = 1) : add_char (zmod n) C := +{ to_fun := λ (a : multiplicative (zmod n)), ζ ^ a.to_add.val, + map_one' := by simp only [to_add_one, zmod.val_zero, pow_zero], + map_mul' := λ x y, by rw [to_add_mul, ← pow_add, zmod.val_add (to_add x) (to_add y), + ← pow_eq_pow_mod _ hζ] } + +/-- An additive character on `zmod n` is nontrivial iff it takes a value `≠ 1` on `1`. -/ +lemma zmod_char_is_nontrivial_iff (n : ℕ) [fact (0 < n)] (ψ : add_char (zmod n) C) : + is_nontrivial ψ ↔ ψ (of_add 1) ≠ 1 := +begin + refine ⟨_, λ h, ⟨1, h⟩⟩, + contrapose!, + rintros h₁ ⟨a, ha⟩, + have ha₁ : a = a.val • 1, + { rw [nsmul_eq_mul, mul_one], exact (zmod.nat_cast_zmod_val a).symm }, + have ha₂ : of_add a = (of_add 1) ^ a.val := by rw [← of_add_nsmul _ a.val, ← ha₁], + rw [ha₂, map_pow, h₁, one_pow] at ha, + exact ha rfl, +end + +/-- A primitive additive character on `zmod n` takes the value `1` only at `0`. -/ +lemma is_primitive.zmod_char_eq_one_iff (n : ℕ) [fact (0 < n)] + {ψ : add_char (zmod n) C} (hψ : is_primitive ψ) (a : zmod n) : + ψ (of_add a) = 1 ↔ a = 0 := +begin + refine ⟨λ h, not_imp_comm.mp (hψ a) _, λ ha, (by rw [ha, of_add_zero, map_one])⟩, + rw [zmod_char_is_nontrivial_iff n (mul_shift ψ a)], + simpa only [mul_shift, monoid_hom.coe_comp, function.comp_app, add_monoid_hom.coe_mul_left, + add_monoid_hom.to_multiplicative_apply_apply, to_add_of_add, mul_one, not_not], +end + +/-- The converse: if the additive character takes the value `1` only at `0`, +then it is primitive. -/ +lemma zmod_char_primitive_of_eq_one_only_at_zero (n : ℕ) + (ψ : add_char (zmod n) C) (hψ : ∀ a, ψ (of_add a) = 1 → a = 0) : + is_primitive ψ := +begin + refine λ a ha, (is_nontrivial_iff_ne_trivial _).mpr (λ hf, _), + have h : mul_shift ψ a (of_add 1) = (1 : add_char (zmod n) C) (of_add (1 : zmod n)) := + congr_fun (congr_arg coe_fn hf) 1, + simp only [mul_shift, monoid_hom.coe_comp, function.comp_app, + add_monoid_hom.to_multiplicative_apply_apply, to_add_of_add, + add_monoid_hom.coe_mul_left, mul_one, monoid_hom.one_apply] at h, + exact ha (hψ a h), +end + +/-- The additive character on `zmod n` associated to a primitive `n`th root of unity +is primitive -/ +lemma zmod_char_primitive_of_primitive_root (n : ℕ) [hn : fact (0 < n)] + {ζ : C} (h : is_primitive_root ζ n) : + is_primitive (zmod_char n ((is_primitive_root.iff_def ζ n).mp h).left) := +begin + apply zmod_char_primitive_of_eq_one_only_at_zero, + intros a ha, + simp only [zmod_char, monoid_hom.coe_mk, to_add_of_add, ← pow_zero ζ] at ha, + exact (zmod.val_eq_zero a).mp (is_primitive_root.pow_inj h (zmod.val_lt a) hn.1 ha), +end + +/-- There is a primitive additive character on `zmod n` if the characteristic of the target +does not divide `n` -/ +noncomputable +def primitive_zmod_char (n : ℕ) [hn : fact (0 < n)] (F' : Type v) [field F'] (h : (n : F') ≠ 0) : + primitive_add_char (zmod n) F' := +begin + let nn := n.to_pnat hn.1, + haveI : ne_zero ((nn : ℕ) : F') := ⟨h⟩, + haveI : ne_zero ((nn : ℕ) : cyclotomic_field nn F') := ne_zero.of_no_zero_smul_divisors F' _ n, + exact +{ n := nn, + char := zmod_char n (is_cyclotomic_extension.zeta_pow nn F' _), + prim := zmod_char_primitive_of_primitive_root n (is_cyclotomic_extension.zeta_spec nn F' _) } +end + +/-! +### Existence of a primitive additive character on a finite field +-/ + +/-- There is a primitive additive character on the finite field `F` if the characteristic +of the target is different from that of `F`. +We obtain it as the composition of the trace from `F` to `zmod p` with a primitive +additive character on `zmod p`, where `p` is the characteristic of `F`. -/ +noncomputable +def primitive_char_finite_field (F F': Type) [field F] [fintype F] [field F'] + (h : ring_char F' ≠ ring_char F) : + primitive_add_char F F' := +begin + let p := ring_char F, + haveI hp : fact p.prime := ⟨char_p.char_is_prime F _⟩, + haveI hp₁ : fact (0 < p) := ⟨hp.1.pos⟩, + have hp₂ : ¬ ring_char F' ∣ p := + begin + cases char_p.char_is_prime_or_zero F' (ring_char F') with hq hq, + { exact mt (nat.prime.dvd_iff_eq hp.1 (nat.prime.ne_one hq)).mp h.symm, }, + { rw [hq], + exact λ hf, nat.prime.ne_zero hp.1 (zero_dvd_iff.mp hf), }, + end, + let ψ := primitive_zmod_char p F' (ne_zero_iff.mp (ne_zero.of_not_dvd F' hp₂)), + let ψ' := ψ.char.comp (add_monoid_hom.to_multiplicative + (algebra.trace (zmod p) F).to_add_monoid_hom), + have hψ' : is_nontrivial ψ' := + begin + obtain ⟨a, ha⟩ := finite_field.trace_to_zmod_nondegenerate F one_ne_zero, + rw one_mul at ha, + exact ⟨a, λ hf, ha $ (ψ.prim.zmod_char_eq_one_iff p $ algebra.trace (zmod p) F a).mp hf⟩, + end, + exact +{ n := ψ.n, + char := ψ', + prim := hψ'.is_primitive ψ' }, +end + +/-! +### The sum of all character values +-/ + +open_locale big_operators + +variables [fintype R] + +/-- The sum over the values of a nontrivial additive character vanishes if the target ring +is a domain. -/ +lemma sum_eq_zero_of_is_nontrivial' [is_domain R'] {ψ : add_char R R'} (hψ : is_nontrivial ψ) : + ∑ a, ψ a = 0 := +begin + rcases hψ with ⟨b, hb⟩, + have h₁ : ∑ (a : R), ψ (of_add (b + a)) = ∑ (a : R), ψ a := + fintype.sum_bijective _ (add_group.add_left_bijective b) _ _ (λ x, rfl), + simp_rw [of_add_add, ψ.map_mul] at h₁, + have h₂ : ∑ (a : R), ψ (of_add a) = finset.univ.sum ⇑ψ := rfl, + rw [← finset.mul_sum, h₂] at h₁, + exact eq_zero_of_mul_eq_self_left hb h₁, +end + +lemma sum_eq_zero_of_is_nontrivial [is_domain R'] {ψ : add_char R R'} (hψ : is_nontrivial ψ) : + ∑ a, ψ (of_add a) = 0 := +sum_eq_zero_of_is_nontrivial' hψ + +/-- The sum over the values of the trivial additive character is the cardinality of the source. -/ +lemma sum_eq_card_of_is_trivial {ψ : add_char R R'} (hψ : ¬ is_nontrivial ψ) : + ∑ a, ψ (of_add a) = fintype.card R := +begin + simp only [is_nontrivial] at hψ, + push_neg at hψ, + simp only [hψ, finset.sum_const, nat.smul_one_eq_coe], + refl, +end + +lemma sum_eq_card_of_is_trivial' {ψ : add_char R R'} (hψ : ¬ is_nontrivial ψ) : + ∑ a, ψ a = fintype.card R := +sum_eq_card_of_is_trivial hψ + +end add_char + +end additive From 07ed95f4e2754e293250a54f93aabc9149173822 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sat, 23 Jul 2022 22:32:58 +0000 Subject: [PATCH 021/173] feat(analysis/inner_product_space/positive): definition and basic facts about positive operators (#15470) --- src/analysis/inner_product_space/adjoint.lean | 72 ++++++++++ .../inner_product_space/positive.lean | 126 ++++++++++++++++++ 2 files changed, 198 insertions(+) create mode 100644 src/analysis/inner_product_space/positive.lean diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index a080577e7ce11..d9f6ea3d84a1a 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -77,6 +77,26 @@ by rw [hT x y, inner_conj_sym] ⟪T x, y⟫ = ⟪x, T y⟫ := hT x y +lemma is_self_adjoint_zero : is_self_adjoint (0 : E →ₗ[𝕜] E) := +λ x y, (inner_zero_right : ⟪x, 0⟫ = 0).symm ▸ (inner_zero_left : ⟪0, y⟫ = 0) + +lemma is_self_adjoint_id : is_self_adjoint (linear_map.id : E →ₗ[𝕜] E) := +λ x y, rfl + +lemma is_self_adjoint.add {T S : E →ₗ[𝕜] E} (hT : is_self_adjoint T) + (hS : is_self_adjoint S) : is_self_adjoint (T + S) := +begin + intros x y, + rw [linear_map.add_apply, inner_add_left, hT x y, hS x y, ← inner_add_right], + refl +end + +/-- The orthogonal projection is self-adjoint. -/ +lemma orthogonal_projection_is_self_adjoint [complete_space E] (U : submodule 𝕜 E) + [complete_space U] : + is_self_adjoint (U.subtypeL ∘L orthogonal_projection U : E →ₗ[𝕜] E):= +inner_orthogonal_projection_left_eq_right U + /-- The **Hellinger--Toeplitz theorem**: if a symmetric operator is defined everywhere, then it is automatically continuous. -/ lemma is_self_adjoint.continuous [complete_space E] {T : E →ₗ[𝕜] E} (hT : is_self_adjoint T) : @@ -254,6 +274,58 @@ begin exact ext_inner_right 𝕜 (λ y, by simp only [adjoint_inner_left, h x y]) end +@[simp] lemma is_self_adjoint_iff_adjoint_eq (A : E →L[𝕜] E) : + is_self_adjoint (A : E →ₗ[𝕜] E) ↔ A† = A := +by simp_rw [is_self_adjoint, coe_coe, ← eq_adjoint_iff, eq_comm] + +lemma _root_.inner_product_space.is_self_adjoint.adjoint_eq {A : E →L[𝕜] E} + (hA : is_self_adjoint (A : E →ₗ[𝕜] E)) : A† = A := +by rwa is_self_adjoint_iff_adjoint_eq at hA + +lemma _root_.inner_product_space.is_self_adjoint.conj_adjoint {T : E →L[𝕜] E} + (hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (S : E →L[𝕜] F) : + is_self_adjoint (S ∘L T ∘L S† : F →ₗ[𝕜] F) := +begin + intros x y, + rw [coe_coe, comp_apply, comp_apply, ← adjoint_inner_right, ← coe_coe, hT, coe_coe, + adjoint_inner_left], + refl +end + +lemma _root_.inner_product_space.is_self_adjoint.adjoint_conj {T : E →L[𝕜] E} + (hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (S : F →L[𝕜] E) : + is_self_adjoint (S† ∘L T ∘L S : F →ₗ[𝕜] F) := +begin + convert hT.conj_adjoint (S†), + rw adjoint_adjoint +end + +lemma _root_.inner_product_space.is_self_adjoint.conj_orthogonal_projection {T : E →L[𝕜] E} + (hT : is_self_adjoint (T : E →ₗ[𝕜] E)) (U : submodule 𝕜 E) [complete_space U] : + is_self_adjoint (U.subtypeL ∘L orthogonal_projection U ∘L T ∘L U.subtypeL ∘L + orthogonal_projection U : E →ₗ[𝕜] E) := +begin + have := hT.conj_adjoint (U.subtypeL ∘L orthogonal_projection U), + rwa (orthogonal_projection_is_self_adjoint U).adjoint_eq at this +end + +lemma _root_.submodule.adjoint_subtypeL (U : submodule 𝕜 E) + [complete_space U] : + (U.subtypeL)† = orthogonal_projection U := +begin + symmetry, + rw eq_adjoint_iff, + intros x u, + rw [U.coe_inner, inner_orthogonal_projection_left_eq_right, + orthogonal_projection_mem_subspace_eq_self], + refl +end + +lemma _root_.submodule.adjoint_orthogonal_projection (U : submodule 𝕜 E) + [complete_space U] : + (orthogonal_projection U : E →L[𝕜] U)† = U.subtypeL := +by rw [← U.adjoint_subtypeL, adjoint_adjoint] + /-- `E →L[𝕜] E` is a star algebra with the adjoint as the star operation. -/ instance : has_star (E →L[𝕜] E) := ⟨adjoint⟩ instance : has_involutive_star (E →L[𝕜] E) := ⟨adjoint_adjoint⟩ diff --git a/src/analysis/inner_product_space/positive.lean b/src/analysis/inner_product_space/positive.lean new file mode 100644 index 0000000000000..8953f3da0eff7 --- /dev/null +++ b/src/analysis/inner_product_space/positive.lean @@ -0,0 +1,126 @@ +/- +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.inner_product_space.l2_space +import analysis.inner_product_space.adjoint + +/-! +# Positive operators + +In this file we define positive operators in a Hilbert space. We follow Bourbaki's choice +of requiring self adjointness in the definition. + +## Main definitions + +* `is_positive` : a continuous linear map is positive if it is self adjoint and + `∀ x, 0 ≤ re ⟪T x, x⟫` + +## Main statements + +* `continuous_linear_map.is_positive.conj_adjoint` : if `T : E →L[𝕜] E` is positive, + then for any `S : E →L[𝕜] F`, `S ∘L T ∘L S†` is also positive. +* `continuous_linear_map.is_positive_iff_complex` : in a ***complex*** hilbert space, + checking that `⟪T x, x⟫` is a nonnegative real number for all `x` suffices to prove that + `T` is positive + +## References + +* [Bourbaki, *Topological Vector Spaces*][bourbaki1987] + +## Tags + +Positive operator +-/ + +open inner_product_space is_R_or_C continuous_linear_map +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] +local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y + +/-- A continuous linear endomorphism `T` of a Hilbert space is **positive** if it is self adjoint + and `∀ x, 0 ≤ re ⟪T x, x⟫`. -/ +def is_positive (T : E →L[𝕜] E) : Prop := + is_self_adjoint (T : E →ₗ[𝕜] E) ∧ ∀ x, 0 ≤ T.re_apply_inner_self x + +lemma is_positive.is_self_adjoint {T : E →L[𝕜] E} (hT : is_positive T) : + is_self_adjoint (T : E →ₗ[𝕜] E) := +hT.1 + +lemma is_positive.inner_nonneg_left {T : E →L[𝕜] E} (hT : is_positive T) (x : E) : + 0 ≤ re ⟪T x, x⟫ := +hT.2 x + +lemma is_positive.inner_nonneg_right {T : E →L[𝕜] E} (hT : is_positive T) (x : E) : + 0 ≤ re ⟪x, T x⟫ := +by rw inner_re_symm; exact hT.inner_nonneg_left x + +lemma is_positive_zero : is_positive (0 : E →L[𝕜] E) := +begin + refine ⟨is_self_adjoint_zero, λ x, _⟩, + change 0 ≤ re ⟪_, _⟫, + rw [zero_apply, inner_zero_left, zero_hom_class.map_zero] +end + +lemma is_positive_id : is_positive (1 : E →L[𝕜] E) := +⟨λ x y, rfl, λ x, inner_self_nonneg⟩ + +lemma is_positive.add {T S : E →L[𝕜] E} (hT : T.is_positive) + (hS : S.is_positive) : (T + S).is_positive := +begin + refine ⟨hT.is_self_adjoint.add hS.is_self_adjoint, λ x, _⟩, + rw [re_apply_inner_self, add_apply, inner_add_left, map_add], + exact add_nonneg (hT.inner_nonneg_left x) (hS.inner_nonneg_left x) +end + +lemma is_positive.conj_adjoint [complete_space E] [complete_space F] {T : E →L[𝕜] E} + (hT : T.is_positive) (S : E →L[𝕜] F) : (S ∘L T ∘L S†).is_positive := +begin + refine ⟨hT.is_self_adjoint.conj_adjoint S, λ x, _⟩, + rw [re_apply_inner_self, comp_apply, ← adjoint_inner_right], + exact hT.inner_nonneg_left _ +end + +lemma is_positive.adjoint_conj [complete_space E] [complete_space F] {T : E →L[𝕜] E} + (hT : T.is_positive) (S : F →L[𝕜] E) : (S† ∘L T ∘L S).is_positive := +begin + convert hT.conj_adjoint (S†), + rw adjoint_adjoint +end + +lemma is_positive.conj_orthogonal_projection [complete_space E] (U : submodule 𝕜 E) {T : E →L[𝕜] E} + (hT : T.is_positive) [complete_space U] : + (U.subtypeL ∘L orthogonal_projection U ∘L T ∘L U.subtypeL ∘L + orthogonal_projection U).is_positive := +begin + have := hT.conj_adjoint (U.subtypeL ∘L orthogonal_projection U), + rwa (orthogonal_projection_is_self_adjoint U).adjoint_eq at this +end + +lemma is_positive.orthogonal_projection_comp [complete_space E] {T : E →L[𝕜] E} + (hT : T.is_positive) (U : submodule 𝕜 E) [complete_space U] : + (orthogonal_projection U ∘L T ∘L U.subtypeL).is_positive := +begin + have := hT.conj_adjoint (orthogonal_projection U : E →L[𝕜] U), + rwa [U.adjoint_orthogonal_projection] at this, +end + +section complex + +variables {E' : Type*} [inner_product_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⟫_ℂ := +begin + simp_rw [is_positive, forall_and_distrib, is_self_adjoint_iff_inner_map_self_real, + eq_conj_iff_re], + refl +end + +end complex + +end continuous_linear_map From de3de1fd7f41836e5662c043b8e13fe82387f0dd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 24 Jul 2022 02:48:42 +0000 Subject: [PATCH 022/173] feat(analysis/normed_space/linear_isometry): add defs and lemmas (#15650) * add `linear_isometry_equiv.of_top` and `linear_isometry.equiv_range`; * add lemmas about (pre)images of balls and spheres. --- .../normed_space/linear_isometry.lean | 54 +++++++++++++++++++ 1 file changed, 54 insertions(+) diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 87030795cef45..fdac94b0e337d 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -145,6 +145,18 @@ instance : continuous_semilinear_map_class (E →ₛₗᵢ[σ₁₂] E₂) σ₁ map_continuous := λ f, f.continuous, ..linear_isometry.add_monoid_hom_class } +@[simp] lemma preimage_ball (x : E) (r : ℝ) : + f ⁻¹' (metric.ball (f x) r) = metric.ball x r := +f.isometry.preimage_ball x r + +@[simp] lemma preimage_sphere (x : E) (r : ℝ) : + f ⁻¹' (metric.sphere (f x) r) = metric.sphere x r := +f.isometry.preimage_sphere x r + +@[simp] lemma preimage_closed_ball (x : E) (r : ℝ) : + f ⁻¹' (metric.closed_ball (f x) r) = metric.closed_ball x r := +f.isometry.preimage_closed_ball x r + lemma ediam_image (s : set E) : emetric.diam (f '' s) = emetric.diam s := f.isometry.ediam_image s @@ -539,12 +551,39 @@ protected lemma lipschitz : lipschitz_with 1 e := e.isometry.lipschitz protected lemma antilipschitz : antilipschitz_with 1 e := e.isometry.antilipschitz +lemma image_eq_preimage (s : set E) : e '' s = e.symm ⁻¹' s := +e.to_linear_equiv.image_eq_preimage s + @[simp] lemma ediam_image (s : set E) : emetric.diam (e '' s) = emetric.diam s := e.isometry.ediam_image s @[simp] lemma diam_image (s : set E) : metric.diam (e '' s) = metric.diam s := 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 + +@[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 + +@[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 + +@[simp] lemma image_ball (x : E) (r : ℝ) : + e '' (metric.ball x r) = metric.ball (e x) r := +e.to_isometric.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 + +@[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 + variables {α : Type*} [topological_space α] @[simp] lemma comp_continuous_on_iff {f : α → E} {s : set α} : @@ -607,6 +646,13 @@ rfl ((prod_assoc R E E₂ E₃).symm : E × E₂ × E₃ → (E × E₂) × E₃) = (equiv.prod_assoc E E₂ E₃).symm := rfl +/-- If `p` is a submodule that is equal to `⊤`, then `linear_isometry_equiv.of_top p hp` is the +"identity" equivalence between `p` and `E`. -/ +@[simps to_linear_equiv apply symm_apply_coe] +def of_top {R : Type*} [ring R] [module R E] (p : submodule R E) (hp : p = ⊤) : + p ≃ₗᵢ[R] E := +{ to_linear_equiv := linear_equiv.of_top p hp, .. p.subtypeₗᵢ } + variables {R E E₂ E₃} {R' : Type*} [ring R'] [module R' E] (p q : submodule R' E) /-- `linear_equiv.of_eq` as a `linear_isometry_equiv`. -/ @@ -636,3 +682,11 @@ lemma basis.ext_linear_isometry_equiv {ι : Type*} (b : basis ι R E) {f₁ f₂ linear_isometry_equiv.to_linear_equiv_injective $ b.ext' h omit σ₂₁ + +/-- Reinterpret a `linear_isometry` as a `linear_isometry_equiv` to the range. -/ +@[simps to_linear_equiv apply_coe] +noncomputable def linear_isometry.equiv_range {R S : Type*} [semiring R] [ring S] [module S E] + [module R F] {σ₁₂ : R →+* S} {σ₂₁ : S →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] + [ring_hom_inv_pair σ₂₁ σ₁₂] (f : F →ₛₗᵢ[σ₁₂] E) : + F ≃ₛₗᵢ[σ₁₂] f.to_linear_map.range := +{ to_linear_equiv := linear_equiv.of_injective f.to_linear_map f.injective, .. f } From 8e4d5ef4d5ac10570edc992bd812cc6a0bf753ec Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sun, 24 Jul 2022 04:00:22 +0000 Subject: [PATCH 023/173] chore(scripts): update nolints.txt (#15656) 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 ed8775d4740fa..c7a1ccb0979f5 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -547,9 +547,6 @@ apply_nolint wseq.mem doc_blame apply_nolint wseq.tail.aux doc_blame apply_nolint wseq.think_congr unused_arguments --- data/set/countable.lean -apply_nolint set.countable_pi fintype_finite - -- data/set/finite.lean apply_nolint infi_supr_of_antitone fintype_finite apply_nolint infi_supr_of_monotone fintype_finite From a17f46f6248e654f77bcfb16fa26ba57ea61ad3e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 24 Jul 2022 14:07:55 +0000 Subject: [PATCH 024/173] feat(data/list/of_fn): lemmas to turn quantifiers over lists to quantifiers over tuples (#15433) In order to prove a property of the recursor, this adds some helper lemmas to `function.left_inverse` --- src/data/fin/tuple/basic.lean | 11 +++++++++++ src/data/list/of_fn.lean | 33 ++++++++++++++++++++++++++++++++- src/logic/function/basic.lean | 15 +++++++++++++++ 3 files changed, 58 insertions(+), 1 deletion(-) diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index 4c7776207ab0d..d9b7b9aaf2eaf 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -690,4 +690,15 @@ mem_find_iff.2 ⟨hi, λ j hj, le_of_eq $ h i j hi hj⟩ end find +/-- To show two sigma pairs of tuples agree, it to show the second elements are related via +`fin.cast`. -/ +lemma sigma_eq_of_eq_comp_cast {α : Type*} : + ∀ {a b : Σ ii, fin ii → α} (h : a.fst = b.fst), a.snd = b.snd ∘ fin.cast h → a = b +| ⟨ai, a⟩ ⟨bi, b⟩ hi h := +begin + dsimp only at hi, + subst hi, + simpa using h, +end + end fin diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 5197f8bb9c1b6..313b8c1d98cf3 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import data.fin.basic +import data.fin.tuple.basic import data.list.basic import data.list.join @@ -20,6 +20,8 @@ The main statements pertain to lists generated using `of_fn` - `list.length_of_fn`, which tells us the length of such a list - `list.nth_of_fn`, which tells us the nth element of such a list - `list.array_eq_of_fn`, which interprets the list form of an array as such a list. +- `list.equiv_sigma_tuple`, which is an `equiv` between lists and the functions that generate them + via `list.of_fn`. -/ universes u @@ -157,4 +159,33 @@ by simp only [mem_of_fn, set.forall_range_iff] of_fn (λ i : fin n, c) = repeat c n := nat.rec_on n (by simp) $ λ n ihn, by simp [ihn] +/-- Lists are equivalent to the sigma type of tuples of a given length. -/ +@[simps] +def equiv_sigma_tuple : list α ≃ Σ n, fin n → α := +{ to_fun := λ l, ⟨l.length, λ i, l.nth_le ↑i i.2⟩, + inv_fun := λ f, list.of_fn f.2, + left_inv := list.of_fn_nth_le, + right_inv := λ ⟨n, f⟩, fin.sigma_eq_of_eq_comp_cast (length_of_fn _) $ funext $ λ i, + nth_le_of_fn' f i.prop } + +/-- A recursor for lists that expands a list into a function mapping to its elements. + +This can be used with `induction l using list.of_fn_rec`. -/ +@[elab_as_eliminator] +def of_fn_rec {C : list α → Sort*} (h : Π n (f : fin n → α), C (list.of_fn f)) (l : list α) : C l := +cast (congr_arg _ l.of_fn_nth_le) $ h l.length (λ i, l.nth_le ↑i i.2) + +@[simp] +lemma of_fn_rec_of_fn {C : list α → Sort*} (h : Π n (f : fin n → α), C (list.of_fn f)) + {n : ℕ} (f : fin n → α) : @of_fn_rec _ C h (list.of_fn f) = h _ f := +equiv_sigma_tuple.right_inverse_symm.cast_eq (λ s, h s.1 s.2) ⟨n, f⟩ + +lemma exists_iff_exists_tuple {P : list α → Prop} : + (∃ l : list α, P l) ↔ ∃ n (f : fin n → α), P (list.of_fn f) := +equiv_sigma_tuple.symm.surjective.exists.trans sigma.exists + +lemma forall_iff_forall_tuple {P : list α → Prop} : + (∀ l : list α, P l) ↔ ∀ n (f : fin n → α), P (list.of_fn f) := +equiv_sigma_tuple.symm.surjective.forall.trans sigma.forall + end list diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index a9ca7d4410e2c..7af69f462df50 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -774,6 +774,21 @@ lemma eq_rec_inj {α : Sort*} {a a' : α} (h : a = a') {C : α → Type*} (x y : lemma cast_inj {α β : Type*} (h : α = β) {x y : α} : cast h x = cast h y ↔ x = y := (cast_bijective h).injective.eq_iff +lemma function.left_inverse.eq_rec_eq {α β : Sort*} {γ : β → Sort v} {f : α → β} {g : β → α} + (h : function.left_inverse g f) (C : Π a : α, γ (f a)) (a : α) : + (congr_arg f (h a)).rec (C (g (f a))) = C a := +eq_of_heq $ (eq_rec_heq _ _).trans $ by rw h + +lemma function.left_inverse.eq_rec_on_eq {α β : Sort*} {γ : β → Sort v} {f : α → β} {g : β → α} + (h : function.left_inverse g f) (C : Π a : α, γ (f a)) (a : α) : + (congr_arg f (h a)).rec_on (C (g (f a))) = C a := +h.eq_rec_eq _ _ + +lemma function.left_inverse.cast_eq {α β : Sort*} {γ : β → Sort v} {f : α → β} {g : β → α} + (h : function.left_inverse g f) (C : Π a : α, γ (f a)) (a : α) : + cast (congr_arg (λ a, γ (f a)) (h a)) (C (g (f a))) = C a := +eq_of_heq $ (eq_rec_heq _ _).trans $ by rw h + /-- A set of functions "separates points" if for each pair of distinct points there is a function taking different values on them. -/ def set.separates_points {α β : Type*} (A : set (α → β)) : Prop := From a5272f940ff867c9d5c8f1cfada54725a0e11653 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 24 Jul 2022 16:34:38 +0000 Subject: [PATCH 025/173] =?UTF-8?q?feat(set=5Ftheory/zfc/basic):=20`?= =?UTF-8?q?=E2=88=88`=20is=20well-founded=20on=20classes=20(#15544)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/set_theory/zfc/basic.lean | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 6e4cf52262186..779c333d8bdec 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -627,6 +627,7 @@ iff.trans mem_sUnion @@mem_sep (λ z : Set.{u}, z ∉ y) /-- 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 @@ -785,6 +786,26 @@ instance : has_mem Class Class := ⟨Class.mem⟩ theorem mem_univ {A : Class.{u}} : A ∈ univ.{u} ↔ ∃ x : Set.{u}, ↑x = A := exists_congr $ λ x, and_true _ +theorem mem_wf : @well_founded Class.{u} (∈) := +⟨begin + have H : ∀ x : Set.{u}, @acc Class.{u} (∈) ↑x, + { refine λ a, Set.induction_on a (λ x IH, ⟨x, _⟩), + rintros A ⟨z, rfl, hz⟩, + exact IH z hz }, + { refine λ A, ⟨A, _⟩, + rintros B ⟨x, rfl, hx⟩, + exact H x } +end⟩ + +instance : has_well_founded Class := ⟨_, mem_wf⟩ +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. -/ +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} From abd55192035cc5f0e24747d3a664b04fc35deb86 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 24 Jul 2022 16:34:39 +0000 Subject: [PATCH 026/173] chore(*): Rename `normed_group` to `normed_add_comm_group` (#15619) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `normed_group` → `normed_add_comm_group` * `semi_normed_group` → `seminormed_add_comm_group`. Elision of the underscore corresponds to `seminorm` (and to counterbalance the name being sensibly longer). * `normed_group_hom` → `normed_add_group_hom` --- counterexamples/phillips.lean | 18 +- src/algebra/module/pi.lean | 3 +- src/analysis/ODE/gronwall.lean | 4 +- src/analysis/ODE/picard_lindelof.lean | 4 +- src/analysis/analytic/basic.lean | 6 +- src/analysis/analytic/composition.lean | 8 +- src/analysis/analytic/inverse.lean | 4 +- src/analysis/analytic/linear.lean | 6 +- src/analysis/analytic/radius_liminf.lean | 4 +- .../asymptotics/asymptotic_equivalent.lean | 16 +- src/analysis/asymptotics/asymptotics.lean | 22 +- .../asymptotics/specific_asymptotics.lean | 6 +- src/analysis/asymptotics/theta.lean | 6 +- src/analysis/box_integral/basic.lean | 4 +- .../box_integral/divergence_theorem.lean | 2 +- src/analysis/box_integral/integrability.lean | 2 +- .../box_integral/partition/additive.lean | 2 +- .../box_integral/partition/measure.lean | 4 +- src/analysis/calculus/affine_map.lean | 4 +- .../calculus/conformal/normed_space.lean | 4 +- src/analysis/calculus/cont_diff.lean | 55 +-- src/analysis/calculus/deriv.lean | 21 +- src/analysis/calculus/diff_on_int_cont.lean | 6 +- src/analysis/calculus/dslope.lean | 4 +- src/analysis/calculus/extend_deriv.lean | 4 +- src/analysis/calculus/fderiv.lean | 28 +- src/analysis/calculus/fderiv_analytic.lean | 4 +- src/analysis/calculus/fderiv_measurable.lean | 8 +- src/analysis/calculus/fderiv_symmetric.lean | 4 +- .../calculus/formal_multilinear_series.lean | 6 +- src/analysis/calculus/implicit.lean | 20 +- src/analysis/calculus/inverse.lean | 16 +- src/analysis/calculus/iterated_deriv.lean | 4 +- .../calculus/lagrange_multipliers.lean | 4 +- src/analysis/calculus/local_extr.lean | 2 +- src/analysis/calculus/mean_value.lean | 15 +- .../calculus/parametric_integral.lean | 4 +- .../parametric_interval_integral.lean | 4 +- src/analysis/calculus/specific_functions.lean | 9 +- src/analysis/calculus/tangent_cone.lean | 16 +- src/analysis/complex/abs_max.lean | 4 +- src/analysis/complex/basic.lean | 14 +- src/analysis/complex/cauchy_integral.lean | 2 +- src/analysis/complex/conformal.lean | 2 +- src/analysis/complex/liouville.lean | 4 +- src/analysis/complex/phragmen_lindelof.lean | 2 +- src/analysis/complex/real_deriv.lean | 4 +- .../complex/removable_singularity.lean | 2 +- src/analysis/complex/schwarz.lean | 3 +- src/analysis/convex/exposed.lean | 2 +- src/analysis/convex/gauge.lean | 2 +- src/analysis/convex/integral.lean | 4 +- src/analysis/convex/krein_milman.lean | 2 +- src/analysis/convex/measure.lean | 4 +- src/analysis/convex/strict_convex_space.lean | 6 +- src/analysis/convex/topology.lean | 2 +- src/analysis/convex/uniform.lean | 10 +- src/analysis/convolution.lean | 18 +- src/analysis/inner_product_space/basic.lean | 16 +- .../inner_product_space/calculus.lean | 4 +- .../inner_product_space/euclidean_dist.lean | 4 +- .../locally_convex/balanced_core_hull.lean | 3 +- src/analysis/locally_convex/bounded.lean | 2 +- .../locally_convex/with_seminorms.lean | 8 +- src/analysis/matrix.lean | 80 ++-- src/analysis/normed/field/basic.lean | 67 ++-- .../normed/group/SemiNormedGroup.lean | 41 +- .../group/SemiNormedGroup/completion.lean | 16 +- .../normed/group/SemiNormedGroup/kernels.lean | 52 +-- src/analysis/normed/group/add_torsor.lean | 19 +- src/analysis/normed/group/ball_sphere.lean | 2 +- src/analysis/normed/group/basic.lean | 215 ++++++----- src/analysis/normed/group/completion.lean | 4 +- src/analysis/normed/group/hom.lean | 352 +++++++++--------- src/analysis/normed/group/hom_completion.lean | 136 +++---- src/analysis/normed/group/infinite_sum.lean | 2 +- src/analysis/normed/group/pointwise.lean | 6 +- src/analysis/normed/group/quotient.lean | 88 ++--- src/analysis/normed_space/M_structure.lean | 2 +- src/analysis/normed_space/add_torsor.lean | 6 +- .../normed_space/add_torsor_bases.lean | 9 +- .../normed_space/affine_isometry.lean | 8 +- src/analysis/normed_space/ball_action.lean | 2 +- src/analysis/normed_space/banach.lean | 4 +- .../normed_space/banach_steinhaus.lean | 2 +- src/analysis/normed_space/basic.lean | 54 +-- .../normed_space/bounded_linear_maps.lean | 24 +- src/analysis/normed_space/complemented.lean | 5 +- src/analysis/normed_space/completion.lean | 6 +- .../normed_space/conformal_linear_map.lean | 6 +- .../normed_space/continuous_affine_map.lean | 10 +- src/analysis/normed_space/dual.lean | 20 +- src/analysis/normed_space/enorm.lean | 2 +- src/analysis/normed_space/extend.lean | 2 +- src/analysis/normed_space/extr.lean | 2 +- .../normed_space/finite_dimension.lean | 46 +-- .../normed_space/hahn_banach/extension.lean | 6 +- .../normed_space/hahn_banach/separation.lean | 4 +- .../normed_space/indicator_function.lean | 2 +- src/analysis/normed_space/is_R_or_C.lean | 4 +- .../normed_space/lattice_ordered_group.lean | 4 +- .../normed_space/linear_isometry.lean | 27 +- src/analysis/normed_space/lp_space.lean | 19 +- .../normed_space/matrix_exponential.lean | 9 +- src/analysis/normed_space/mazur_ulam.lean | 6 +- src/analysis/normed_space/multilinear.lean | 33 +- src/analysis/normed_space/operator_norm.lean | 45 +-- src/analysis/normed_space/ordered.lean | 6 +- src/analysis/normed_space/pi_Lp.lean | 40 +- src/analysis/normed_space/pointwise.lean | 12 +- src/analysis/normed_space/ray.lean | 4 +- src/analysis/normed_space/riesz_lemma.lean | 4 +- src/analysis/normed_space/star/basic.lean | 8 +- src/analysis/normed_space/star/matrix.lean | 2 +- src/analysis/normed_space/weak_dual.lean | 2 +- src/analysis/seminorm.lean | 2 +- src/analysis/special_functions/arsinh.lean | 2 +- .../special_functions/complex/log_deriv.lean | 2 +- src/analysis/special_functions/exp_deriv.lean | 4 +- src/analysis/special_functions/log/deriv.lean | 4 +- .../special_functions/non_integrable.lean | 4 +- .../special_functions/polar_coord.lean | 2 +- src/analysis/special_functions/pow_deriv.lean | 4 +- src/analysis/special_functions/sqrt.lean | 2 +- .../trigonometric/arctan_deriv.lean | 4 +- .../trigonometric/deriv.lean | 4 +- src/analysis/specific_limits/normed.lean | 27 +- src/analysis/von_neumann_algebra/basic.lean | 2 +- src/combinatorics/additive/salem_spencer.lean | 4 +- src/data/complex/is_R_or_C.lean | 2 +- .../algebra/left_invariant_derivation.lean | 2 +- src/geometry/manifold/algebra/lie_group.lean | 18 +- src/geometry/manifold/algebra/monoid.lean | 20 +- .../manifold/algebra/smooth_functions.lean | 22 +- src/geometry/manifold/algebra/structures.lean | 4 +- src/geometry/manifold/bump_function.lean | 4 +- src/geometry/manifold/conformal_groupoid.lean | 2 +- src/geometry/manifold/cont_mdiff.lean | 14 +- src/geometry/manifold/cont_mdiff_map.lean | 6 +- src/geometry/manifold/derivation_bundle.lean | 6 +- src/geometry/manifold/diffeomorph.lean | 6 +- src/geometry/manifold/instances/sphere.lean | 2 +- src/geometry/manifold/metrizable.lean | 2 +- src/geometry/manifold/mfderiv.lean | 32 +- src/geometry/manifold/partition_of_unity.lean | 6 +- .../smooth_manifold_with_corners.lean | 70 ++-- src/geometry/manifold/tangent_bundle.lean | 14 +- src/geometry/manifold/whitney_embedding.lean | 2 +- src/information_theory/hamming.lean | 9 +- .../constructions/borel_space.lean | 24 +- src/measure_theory/constructions/prod.lean | 4 +- .../covering/besicovitch_vector_space.lean | 4 +- .../function/ae_eq_of_integral.lean | 8 +- .../function/conditional_expectation.lean | 10 +- .../function/continuous_map_dense.lean | 4 +- .../function/convergence_in_measure.lean | 10 +- src/measure_theory/function/jacobian.lean | 4 +- src/measure_theory/function/l1_space.lean | 27 +- src/measure_theory/function/l2_space.lean | 4 +- .../function/locally_integrable.lean | 2 +- src/measure_theory/function/lp_order.lean | 2 +- src/measure_theory/function/lp_space.lean | 30 +- .../function/simple_func_dense_lp.lean | 20 +- .../function/strongly_measurable.lean | 31 +- .../function/strongly_measurable_lp.lean | 2 +- .../function/uniform_integrable.lean | 2 +- .../group/fundamental_domain.lean | 2 +- src/measure_theory/group/integration.lean | 2 +- src/measure_theory/group/measure.lean | 5 +- src/measure_theory/integral/average.lean | 4 +- src/measure_theory/integral/bochner.lean | 29 +- .../integral/circle_integral.lean | 2 +- .../integral/circle_integral_transform.lean | 2 +- .../integral/divergence_theorem.lean | 4 +- .../integral/integrable_on.lean | 14 +- .../integral/integral_eq_improper.lean | 8 +- .../integral/interval_average.lean | 2 +- .../integral/interval_integral.lean | 22 +- src/measure_theory/integral/lebesgue.lean | 4 +- src/measure_theory/integral/periodic.lean | 2 +- src/measure_theory/integral/set_integral.lean | 18 +- src/measure_theory/integral/set_to_l1.lean | 20 +- .../integral/torus_integral.lean | 2 +- src/measure_theory/measure/haar_lebesgue.lean | 30 +- .../measure/with_density_vector_measure.lean | 2 +- src/number_theory/modular.lean | 2 +- src/probability/conditional_expectation.lean | 2 +- src/probability/ident_distrib.lean | 17 +- src/probability/martingale.lean | 2 +- src/probability/stopping.lean | 6 +- src/topology/continuous_function/bounded.lean | 94 ++--- src/topology/continuous_function/compact.lean | 25 +- .../continuous_function/zero_at_infty.lean | 10 +- src/topology/metric_space/algebra.lean | 4 +- .../metric_space/hausdorff_dimension.lean | 9 +- src/topology/vector_bundle/basic.lean | 4 +- src/topology/vector_bundle/hom.lean | 4 +- src/topology/vector_bundle/prod.lean | 4 +- src/topology/vector_bundle/pullback.lean | 2 +- test/calc.lean | 2 +- test/measurability.lean | 2 +- 201 files changed, 1469 insertions(+), 1406 deletions(-) diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index cc507660d0c2b..63dd287b5b9ae 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -379,18 +379,18 @@ by applying the functional to the indicator functions. -/ def _root_.continuous_linear_map.to_bounded_additive_measure [topological_space α] [discrete_topology α] (f : (α →ᵇ ℝ) →L[ℝ] ℝ) : bounded_additive_measure α := -{ to_fun := λ s, f (of_normed_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)), +{ to_fun := λ s, f (of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)), additive' := λ s t hst, begin - have : of_normed_group_discrete (indicator (s ∪ t) 1) 1 (norm_indicator_le_one (s ∪ t)) - = of_normed_group_discrete (indicator s 1) 1 (norm_indicator_le_one s) - + of_normed_group_discrete (indicator t 1) 1 (norm_indicator_le_one t), + have : of_normed_add_comm_group_discrete (indicator (s ∪ t) 1) 1 (norm_indicator_le_one _) + = of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s) + + of_normed_add_comm_group_discrete (indicator t 1) 1 (norm_indicator_le_one t), by { ext x, simp [indicator_union_of_disjoint hst], }, rw [this, f.map_add], end, exists_bound := ⟨∥f∥, λ s, begin - have I : ∥of_normed_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)∥ ≤ 1, - by apply norm_of_normed_group_le _ zero_le_one, + have I : ∥of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)∥ ≤ 1, + by apply norm_of_normed_add_comm_group_le _ zero_le_one, apply le_trans (f.le_op_norm _), simpa using mul_le_mul_of_nonneg_left I (norm_nonneg f), end⟩ } @@ -410,7 +410,7 @@ lemma to_functions_to_measure [measurable_space α] (μ : measure α) [is_finite μ.extension_to_bounded_functions.to_bounded_additive_measure s = (μ s).to_real := begin change μ.extension_to_bounded_functions - (of_normed_group_discrete (indicator s (λ x, 1)) 1 (norm_indicator_le_one s)) = (μ s).to_real, + (of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)) = (μ s).to_real, rw extension_to_bounded_functions_apply, { change ∫ x, s.indicator (λ y, (1 : ℝ)) x ∂μ = _, simp [integral_indicator hs] }, @@ -500,7 +500,7 @@ which is large (it has countable complement), as in the Sierpinski pathological taking values in `{0, 1}`), indexed by a real parameter `x`, corresponding to the characteristic functions of the different fibers of the Sierpinski pathological family -/ def f (Hcont : #ℝ = aleph 1) (x : ℝ) : (discrete_copy ℝ →ᵇ ℝ) := -of_normed_group_discrete (indicator (spf Hcont x) 1) 1 (norm_indicator_le_one _) +of_normed_add_comm_group_discrete (indicator (spf Hcont x) 1) 1 (norm_indicator_le_one _) lemma apply_f_eq_continuous_part (Hcont : #ℝ = aleph 1) (φ : (discrete_copy ℝ →ᵇ ℝ) →L[ℝ] ℝ) (x : ℝ) @@ -579,7 +579,7 @@ end /-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` is uniformly bounded by `1` in norm. -/ lemma norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : ∥f Hcont x∥ ≤ 1 := -norm_of_normed_group_le _ zero_le_one _ +norm_of_normed_add_comm_group_le _ zero_le_one _ /-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` has no Pettis integral. -/ theorem no_pettis_integral (Hcont : #ℝ = aleph 1) : diff --git a/src/algebra/module/pi.lean b/src/algebra/module/pi.lean index 335b179e5ab26..40d470756995d 100644 --- a/src/algebra/module/pi.lean +++ b/src/algebra/module/pi.lean @@ -63,7 +63,8 @@ instance module (α) {r : semiring α} {m : ∀ i, add_comm_monoid $ f i} /- Extra instance to short-circuit type class resolution. For unknown reasons, this is necessary for certain inference problems. E.g., for this to succeed: ```lean -example (β X : Type*) [normed_group β] [normed_space ℝ β] : module ℝ (X → β) := by apply_instance +example (β X : Type*) [normed_add_comm_group β] [normed_space ℝ β] : module ℝ (X → β) := +infer_instance ``` See: https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Typeclass.20resolution.20under.20binders/near/281296989 -/ diff --git a/src/analysis/ODE/gronwall.lean b/src/analysis/ODE/gronwall.lean index 06a3849e4cc4b..a68f9966e0bf3 100644 --- a/src/analysis/ODE/gronwall.lean +++ b/src/analysis/ODE/gronwall.lean @@ -27,8 +27,8 @@ Sec. 4.5][HubbardWest-ode], where `norm_le_gronwall_bound_of_norm_deriv_right_le of `K x` and `f x`. -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] - {F : Type*} [normed_group F] [normed_space ℝ F] +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 diff --git a/src/analysis/ODE/picard_lindelof.lean b/src/analysis/ODE/picard_lindelof.lean index 5b1980db1c81a..70d4c047304ff 100644 --- a/src/analysis/ODE/picard_lindelof.lean +++ b/src/analysis/ODE/picard_lindelof.lean @@ -33,12 +33,12 @@ open_locale filter topological_space nnreal ennreal nat interval noncomputable theory -variables {E : Type*} [normed_group E] [normed_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] /-- This structure holds arguments of the Picard-Lipschitz (Cauchy-Lipschitz) theorem. Unless you want to use one of the auxiliary lemmas, use `exists_forall_deriv_within_Icc_eq_of_lipschitz_of_continuous` instead of using this structure. -/ -structure picard_lindelof (E : Type*) [normed_group E] [normed_space ℝ E] := +structure picard_lindelof (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] := (to_fun : ℝ → E → E) (t_min t_max : ℝ) (t₀ : Icc t_min t_max) diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index 6b18d9f34ad18..dded622c580d3 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -101,9 +101,9 @@ end formal_multilinear_series variables [nontrivially_normed_field 𝕜] -[normed_group E] [normed_space 𝕜 E] -[normed_group F] [normed_space 𝕜 F] -[normed_group G] [normed_space 𝕜 G] +[normed_add_comm_group E] [normed_space 𝕜 E] +[normed_add_comm_group F] [normed_space 𝕜 F] +[normed_add_comm_group G] [normed_space 𝕜 G] namespace formal_multilinear_series diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 6a38ef670628b..2c07461779db9 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -289,10 +289,10 @@ end formal_multilinear_series end topological variables [nontrivially_normed_field 𝕜] - [normed_group E] [normed_space 𝕜 E] - [normed_group F] [normed_space 𝕜 F] - [normed_group G] [normed_space 𝕜 G] - [normed_group H] [normed_space 𝕜 H] + [normed_add_comm_group E] [normed_space 𝕜 E] + [normed_add_comm_group F] [normed_space 𝕜 F] + [normed_add_comm_group G] [normed_space 𝕜 G] + [normed_add_comm_group H] [normed_space 𝕜 H] namespace formal_multilinear_series diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index ee7453d36ddbb..b621409195275 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -32,8 +32,8 @@ open finset filter namespace formal_multilinear_series variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] /-! ### The left inverse of a formal multilinear series -/ diff --git a/src/analysis/analytic/linear.lean b/src/analysis/analytic/linear.lean index 009ea64a77746..7d88f847c567f 100644 --- a/src/analysis/analytic/linear.lean +++ b/src/analysis/analytic/linear.lean @@ -13,9 +13,9 @@ the formal power series `f x = f a + f (x - a)`. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] -{G : Type*} [normed_group G] [normed_space 𝕜 G] +{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] open_locale topological_space classical big_operators nnreal ennreal open set filter asymptotics diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index 0220c26796720..429adbf102d1a 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -14,8 +14,8 @@ $\liminf_{n\to\infty} \frac{1}{\sqrt[n]{∥p n∥}}$. This lemma can't go to `ba would create a circular dependency once we redefine `exp` using `formal_multilinear_series`. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{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 filter asymptotics diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index e25e2427823f3..b7d397a74d148 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -13,8 +13,8 @@ In this file, we define the relation `is_equivalent l u v`, which means that `u- `v` along the filter `l`. Unlike `is_[oO]` relations, this one requires `u` and `v` to have the same codomain `β`. While the -definition only requires `β` to be a `normed_group`, most interesting properties require it to be a -`normed_field`. +definition only requires `β` to be a `normed_add_comm_group`, most interesting properties require it +to be a `normed_field`. ## Notations @@ -23,7 +23,7 @@ We introduce the notation `u ~[l] v := is_equivalent l u v`, which you can use b ## Main results -If `β` is a `normed_group` : +If `β` is a `normed_add_comm_group` : - `_ ~[l] _` is an equivalence relation - Equivalent statements for `u ~[l] const _ c` : @@ -59,9 +59,9 @@ namespace asymptotics open filter function open_locale topological_space -section normed_group +section normed_add_comm_group -variables {α β : Type*} [normed_group β] +variables {α β : Type*} [normed_add_comm_group β] /-- Two functions `u` and `v` are said to be asymptotically equivalent along a filter `l` when `u x - v x = o(v x)` as x converges along `l`. -/ @@ -162,7 +162,7 @@ begin simp, end -end normed_group +end normed_add_comm_group open_locale asymptotics @@ -217,7 +217,7 @@ end normed_field section smul -lemma is_equivalent.smul {α E 𝕜 : Type*} [normed_field 𝕜] [normed_group E] +lemma is_equivalent.smul {α E 𝕜 : Type*} [normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] {a b : α → 𝕜} {u v : α → E} {l : filter α} (hab : a ~[l] b) (huv : u ~[l] v) : (λ x, a x • u x) ~[l] (λ x, b x • v x) := begin @@ -310,7 +310,7 @@ end asymptotics open filter asymptotics open_locale asymptotics -variables {α β : Type*} [normed_group β] +variables {α β : Type*} [normed_add_comm_group β] lemma filter.eventually_eq.is_equivalent {u v : α → β} {l : filter α} (h : u =ᶠ[l] v) : u ~[l] v := is_equivalent.congr_right (is_o_refl_left _ _) h diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index 4d395375e0d04..a452a1bc2d16a 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -54,9 +54,9 @@ variables {α : Type*} {β : Type*} {E : Type*} {F : Type*} {G : Type*} {R : Type*} {R' : Type*} {𝕜 : Type*} {𝕜' : Type*} variables [has_norm E] [has_norm F] [has_norm G] -variables [semi_normed_group E'] [semi_normed_group F'] [semi_normed_group G'] -variables [normed_group E''] [normed_group F''] [normed_group G''] -variables [semi_normed_ring R] [semi_normed_ring R'] +variables [seminormed_add_comm_group E'] [seminormed_add_comm_group F'] + [seminormed_add_comm_group G'] [normed_add_comm_group E''] [normed_add_comm_group F''] + [normed_add_comm_group G''] [semi_normed_ring R] [semi_normed_ring R'] variables [normed_field 𝕜] [normed_field 𝕜'] variables {c c' c₁ c₂ : ℝ} {f : α → E} {g : α → F} {k : α → G} variables {f' : α → E'} {g' : α → F'} {k' : α → G'} @@ -1457,7 +1457,7 @@ begin split, { exact λ h, ⟨λ x, u x / v x, h.tendsto_div_nhds_zero, h.eventually_mul_div_cancel.symm⟩ }, { unfold is_o, rintros ⟨φ, hφ, huvφ⟩ c hpos, - rw normed_group.tendsto_nhds_zero at hφ, + rw normed_add_comm_group.tendsto_nhds_zero at hφ, exact is_O_with_of_eq_mul _ ((hφ c hpos).mono $ λ x, le_of_lt) huvφ } end @@ -1496,8 +1496,8 @@ theorem is_O_of_div_tendsto_nhds {α : Type*} {l : filter α} f =O[l] g := (is_O_iff_div_is_bounded_under hgf).2 $ H.norm.is_bounded_under_le -lemma is_o.tendsto_zero_of_tendsto {α E 𝕜 : Type*} [normed_group E] [normed_field 𝕜] {u : α → E} - {v : α → 𝕜} {l : filter α} {y : 𝕜} (huv : u =o[l] v) (hv : tendsto v l (𝓝 y)) : +lemma is_o.tendsto_zero_of_tendsto {α E 𝕜 : Type*} [normed_add_comm_group E] [normed_field 𝕜] + {u : α → E} {v : α → 𝕜} {l : filter α} {y : 𝕜} (huv : u =o[l] v) (hv : tendsto v l (𝓝 y)) : tendsto u l (𝓝 0) := begin suffices h : u =o[l] (λ x, (1 : 𝕜)), @@ -1582,13 +1582,13 @@ theorem is_O_one_nat_at_top_iff {f : ℕ → E''} : iff.trans (is_O_nat_at_top_iff (λ n h, (one_ne_zero h).elim)) $ by simp only [norm_one, mul_one] -theorem is_O_with_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_group (E' i)] +theorem is_O_with_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] {f : α → Π i, E' i} {C : ℝ} (hC : 0 ≤ C) : is_O_with C l f g' ↔ ∀ i, is_O_with C l (λ x, f x i) g' := have ∀ x, 0 ≤ C * ∥g' x∥, from λ x, mul_nonneg hC (norm_nonneg _), by simp only [is_O_with_iff, pi_norm_le_iff (this _), eventually_all] -@[simp] theorem is_O_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_group (E' i)] +@[simp] theorem is_O_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] {f : α → Π i, E' i} : f =O[l] g' ↔ ∀ i, (λ x, f x i) =O[l] g' := begin @@ -1596,7 +1596,7 @@ begin exact eventually_congr (eventually_at_top.2 ⟨0, λ c, is_O_with_pi⟩) end -@[simp] theorem is_o_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_group (E' i)] +@[simp] theorem is_o_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] {f : α → Π i, E' i} : f =o[l] g' ↔ ∀ i, (λ x, f x i) =o[l] g' := begin @@ -1608,12 +1608,12 @@ end asymptotics open asymptotics -lemma summable_of_is_O {ι E} [normed_group E] [complete_space E] {f : ι → E} {g : ι → ℝ} +lemma summable_of_is_O {ι E} [normed_add_comm_group E] [complete_space E] {f : ι → E} {g : ι → ℝ} (hg : summable g) (h : f =O[cofinite] g) : summable f := let ⟨C, hC⟩ := h.is_O_with in summable_of_norm_bounded_eventually (λ x, C * ∥g x∥) (hg.abs.mul_left _) hC.bound -lemma summable_of_is_O_nat {E} [normed_group E] [complete_space E] {f : ℕ → E} {g : ℕ → ℝ} +lemma summable_of_is_O_nat {E} [normed_add_comm_group E] [complete_space E] {f : ℕ → E} {g : ℕ → ℝ} (hg : summable g) (h : f =O[at_top] g) : summable f := summable_of_is_O hg $ nat.cofinite_eq_at_top.symm ▸ h diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index 7c2c6fa3ea144..5f4d69cbadd0b 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -104,7 +104,7 @@ section real open_locale big_operators open finset -lemma asymptotics.is_o.sum_range {α : Type*} [normed_group α] +lemma asymptotics.is_o.sum_range {α : Type*} [normed_add_comm_group α] {f : ℕ → α} {g : ℕ → ℝ} (h : f =o[at_top] g) (hg : 0 ≤ g) (h'g : tendsto (λ n, ∑ i in range n, g i) at_top at_top) : (λ n, ∑ i in range n, f i) =o[at_top] (λ n, ∑ i in range n, g i) := @@ -143,7 +143,7 @@ begin ... = ε * ∥(∑ i in range n, g i)∥ : by { simp [B], ring } end -lemma asymptotics.is_o_sum_range_of_tendsto_zero {α : Type*} [normed_group α] +lemma asymptotics.is_o_sum_range_of_tendsto_zero {α : Type*} [normed_add_comm_group α] {f : ℕ → α} (h : tendsto f at_top (𝓝 0)) : (λ n, ∑ i in range n, f i) =o[at_top] (λ n, (n : ℝ)) := begin @@ -153,7 +153,7 @@ begin end /-- The Cesaro average of a converging sequence converges to the same limit. -/ -lemma filter.tendsto.cesaro_smul {E : Type*} [normed_group E] [normed_space ℝ E] +lemma filter.tendsto.cesaro_smul {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {u : ℕ → E} {l : E} (h : tendsto u at_top (𝓝 l)) : tendsto (λ (n : ℕ), (n ⁻¹ : ℝ) • (∑ i in range n, u i)) at_top (𝓝 l) := begin diff --git a/src/analysis/asymptotics/theta.lean b/src/analysis/asymptotics/theta.lean index f5cca8f17b8c1..43117658076ce 100644 --- a/src/analysis/asymptotics/theta.lean +++ b/src/analysis/asymptotics/theta.lean @@ -23,9 +23,9 @@ variables {α : Type*} {β : Type*} {E : Type*} {F : Type*} {G : Type*} {R : Type*} {R' : Type*} {𝕜 : Type*} {𝕜' : Type*} variables [has_norm E] [has_norm F] [has_norm G] -variables [semi_normed_group E'] [semi_normed_group F'] [semi_normed_group G'] -variables [normed_group E''] [normed_group F''] [normed_group G''] -variables [semi_normed_ring R] [semi_normed_ring R'] +variables [seminormed_add_comm_group E'] [seminormed_add_comm_group F'] + [seminormed_add_comm_group G'] [normed_add_comm_group E''] [normed_add_comm_group F''] + [normed_add_comm_group G''] [semi_normed_ring R] [semi_normed_ring R'] variables [normed_field 𝕜] [normed_field 𝕜'] variables {c c' c₁ c₂ : ℝ} {f : α → E} {g : α → F} {k : α → G} variables {f' : α → E'} {g' : α → F'} {k' : α → G'} diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index 0449ea5728472..2245ab9942ac9 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -57,8 +57,8 @@ namespace box_integral universes u v w -variables {ι : Type u} {E : Type v} {F : Type w} [normed_group E] [normed_space ℝ E] - [normed_group F] [normed_space ℝ F] {I J : box ι} {π : tagged_prepartition I} +variables {ι : Type u} {E : Type v} {F : Type w} [normed_add_comm_group E] [normed_space ℝ E] + [normed_add_comm_group F] [normed_space ℝ F] {I J : box ι} {π : tagged_prepartition I} open tagged_prepartition diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index a5c0acde4b54e..3b1d6577ad758 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -43,7 +43,7 @@ open continuous_linear_map (lsmul) filter set finset metric noncomputable theory universes u -variables {E : Type u} [normed_group E] [normed_space ℝ E] {n : ℕ} +variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {n : ℕ} namespace box_integral diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index df7802c704605..4573252fca6dd 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -22,7 +22,7 @@ open_locale classical nnreal ennreal topological_space big_operators universes u v -variables {ι : Type u} {E : Type v} [fintype ι] [normed_group E] [normed_space ℝ E] +variables {ι : Type u} {E : Type v} [fintype ι] [normed_add_comm_group E] [normed_space ℝ E] open measure_theory metric set finset filter box_integral diff --git a/src/analysis/box_integral/partition/additive.lean b/src/analysis/box_integral/partition/additive.lean index b08c6d734eecb..e16b9a790d099 100644 --- a/src/analysis/box_integral/partition/additive.lean +++ b/src/analysis/box_integral/partition/additive.lean @@ -153,7 +153,7 @@ end section to_smul -variables {E : Type*} [normed_group E] [normed_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] /-- If `f` is a box-additive map, then so is the map sending `I` to the scalar multiplication by `f I` as a continuous linear map from `E` to itself. -/ diff --git a/src/analysis/box_integral/partition/measure.lean b/src/analysis/box_integral/partition/measure.lean index 59650c2282b00..886f829a265d6 100644 --- a/src/analysis/box_integral/partition/measure.lean +++ b/src/analysis/box_integral/partition/measure.lean @@ -115,11 +115,11 @@ namespace box_additive_map /-- Box-additive map sending each box `I` to the continuous linear endomorphism `x ↦ (volume I).to_real • x`. -/ -protected def volume {E : Type*} [normed_group E] [normed_space ℝ E] : +protected def volume {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] : ι →ᵇᵃ (E →L[ℝ] E) := (volume : measure (ι → ℝ)).to_box_additive.to_smul -lemma volume_apply {E : Type*} [normed_group E] [normed_space ℝ E] (I : box ι) (x : E) : +lemma volume_apply {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] (I : box ι) (x : E) : box_additive_map.volume I x = (∏ j, (I.upper j - I.lower j)) • x := congr_arg2 (•) I.volume_apply rfl diff --git a/src/analysis/calculus/affine_map.lean b/src/analysis/calculus/affine_map.lean index 331bccf98693f..10a08f23164a3 100644 --- a/src/analysis/calculus/affine_map.lean +++ b/src/analysis/calculus/affine_map.lean @@ -20,8 +20,8 @@ This file contains results about smoothness of affine maps. namespace continuous_affine_map variables {𝕜 V W : Type*} [nontrivially_normed_field 𝕜] -variables [normed_group V] [normed_space 𝕜 V] -variables [normed_group W] [normed_space 𝕜 W] +variables [normed_add_comm_group V] [normed_space 𝕜 V] +variables [normed_add_comm_group W] [normed_space 𝕜 W] /-- A continuous affine map between normed vector spaces is smooth. -/ lemma cont_diff {n : with_top ℕ} (f : V →A[𝕜] W) : diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index ed631ac4bab84..5fdcd41193683 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -42,8 +42,8 @@ Maps such as the complex conjugate are considered to be conformal. noncomputable theory -variables {X Y Z : Type*} [normed_group X] [normed_group Y] [normed_group Z] - [normed_space ℝ X] [normed_space ℝ Y] [normed_space ℝ Z] +variables {X Y Z : Type*} [normed_add_comm_group X] [normed_add_comm_group Y] + [normed_add_comm_group Z] [normed_space ℝ X] [normed_space ℝ Y] [normed_space ℝ Z] section loc_conformality diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index c98832dce5957..aa6a6ca3041d0 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -165,16 +165,16 @@ local notation `∞` := (⊤ : with_top ℕ) universes u v w local attribute [instance, priority 1001] -normed_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid +normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid open set fin filter open_locale topological_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] -{G : Type*} [normed_group G] [normed_space 𝕜 G] -{X : Type*} [normed_group X] [normed_space 𝕜 X] +{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 : E} {c : F} {b : E × F → G} {m n : with_top ℕ} @@ -1936,9 +1936,9 @@ which we have already proved previously. spaces live in the same universe. Use instead `cont_diff_on.comp` which removes the universe assumption (but is deduced from this one). -/ private lemma cont_diff_on.comp_same_univ - {Eu : Type u} [normed_group Eu] [normed_space 𝕜 Eu] - {Fu : Type u} [normed_group Fu] [normed_space 𝕜 Fu] - {Gu : Type u} [normed_group Gu] [normed_space 𝕜 Gu] + {Eu : Type u} [normed_add_comm_group Eu] [normed_space 𝕜 Eu] + {Fu : Type u} [normed_add_comm_group Fu] [normed_space 𝕜 Fu] + {Gu : Type u} [normed_add_comm_group Gu] [normed_space 𝕜 Gu] {s : set Eu} {t : set Fu} {g : Fu → Gu} {f : Eu → Fu} (hg : cont_diff_on 𝕜 n g t) (hf : cont_diff_on 𝕜 n f s) (st : s ⊆ f ⁻¹' t) : cont_diff_on 𝕜 n (g ∘ f) s := @@ -1994,13 +1994,13 @@ begin 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_group Eu := by apply_instance, + 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_group Fu := by apply_instance, + 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_group Gu := by apply_instance, + letI : normed_add_comm_group Gu := by apply_instance, letI : normed_space 𝕜 Gu := by apply_instance, -- declare the isomorphisms let isoE : Eu ≃L[𝕜] E := continuous_multilinear_curry_fin0 𝕜 (E × F × G) E, @@ -2204,8 +2204,9 @@ cont_diff_snd.cont_diff_within_at section n_ary variables {E₁ E₂ E₃ E₄ : Type*} -variables [normed_group E₁] [normed_group E₂] [normed_group E₃] [normed_group E₄] -variables [normed_space 𝕜 E₁] [normed_space 𝕜 E₂] [normed_space 𝕜 E₃] [normed_space 𝕜 E₄] +variables [normed_add_comm_group E₁] [normed_add_comm_group E₂] [normed_add_comm_group E₃] + [normed_add_comm_group E₄] [normed_space 𝕜 E₁] [normed_space 𝕜 E₂] [normed_space 𝕜 E₃] + [normed_space 𝕜 E₄] lemma cont_diff.comp₂ {g : E₁ × E₂ → G} {f₁ : F → E₁} {f₂ : F → E₂} (hg : cont_diff 𝕜 n g) (hf₁ : cont_diff 𝕜 n f₁) (hf₂ : cont_diff 𝕜 n f₂) : @@ -2287,8 +2288,8 @@ end section pi -variables {ι ι' : Type*} [fintype ι] [fintype ι'] {F' : ι → Type*} [Π i, normed_group (F' i)] - [Π i, normed_space 𝕜 (F' i)] {φ : Π i, E → F' i} +variables {ι ι' : Type*} [fintype ι] [fintype ι'] {F' : ι → Type*} + [Π i, normed_add_comm_group (F' i)] [Π i, normed_space 𝕜 (F' i)] {φ : Π i, E → F' i} {p' : Π i, E → formal_multilinear_series 𝕜 E (F' i)} {Φ : E → Π i, F' i} {P' : E → formal_multilinear_series 𝕜 E (Π i, F' i)} @@ -2606,8 +2607,8 @@ lemma cont_diff_on.smul {s : set E} {f : E → 𝕜} {g : E → F} /-! ### Cartesian product of two functions -/ section prod_map -variables {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] -variables {F' : Type*} [normed_group F'] [normed_space 𝕜 F'] +variables {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] +variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] /-- The product map of two `C^n` functions within a set at a point is `C^n` within the product set at the product point. -/ @@ -2625,8 +2626,8 @@ lemma cont_diff_within_at.prod_map cont_diff_within_at.prod_map' hf hg /-- The product map of two `C^n` functions on a set is `C^n` on the product set. -/ -lemma cont_diff_on.prod_map {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] - {F' : Type*} [normed_group F'] [normed_space 𝕜 F'] +lemma cont_diff_on.prod_map {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] + {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] {s : set E} {t : set E'} {f : E → F} {g : E' → F'} (hf : cont_diff_on 𝕜 n f s) (hg : cont_diff_on 𝕜 n g t) : cont_diff_on 𝕜 n (prod.map f g) (s ×ˢ t) := @@ -2965,8 +2966,8 @@ section real variables {𝕂 : Type*} [is_R_or_C 𝕂] -{E' : Type*} [normed_group E'] [normed_space 𝕂 E'] -{F' : Type*} [normed_group F'] [normed_space 𝕂 F'] +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕂 E'] +{F' : Type*} [normed_add_comm_group F'] [normed_space 𝕂 F'] /-- If a function has a Taylor series at order at least 1, then at points in the interior of the domain of definition, the term of order 1 of this series is a strict derivative of `f`. -/ @@ -3027,8 +3028,8 @@ hf.cont_diff_at.has_strict_deriv_at hn /-- If `f` has a formal Taylor series `p` up to order `1` on `{x} ∪ s`, where `s` is a convex set, and `∥p x 1∥₊ < K`, then `f` is `K`-Lipschitz in a neighborhood of `x` within `s`. -/ lemma has_ftaylor_series_up_to_on.exists_lipschitz_on_with_of_nnnorm_lt {E F : Type*} - [normed_group E] [normed_space ℝ E] [normed_group F] [normed_space ℝ F] {f : E → F} - {p : E → formal_multilinear_series ℝ E F} {s : set E} {x : E} + [normed_add_comm_group E] [normed_space ℝ E] [normed_add_comm_group F] [normed_space ℝ F] + {f : E → F} {p : E → formal_multilinear_series ℝ E F} {s : set E} {x : E} (hf : has_ftaylor_series_up_to_on 1 f p (insert x s)) (hs : convex ℝ s) (K : ℝ≥0) (hK : ∥p x 1∥₊ < K) : ∃ t ∈ 𝓝[s] x, lipschitz_on_with K f t := @@ -3046,8 +3047,8 @@ end /-- If `f` has a formal Taylor series `p` up to order `1` on `{x} ∪ s`, where `s` is a convex set, then `f` is Lipschitz in a neighborhood of `x` within `s`. -/ -lemma has_ftaylor_series_up_to_on.exists_lipschitz_on_with {E F : Type*} - [normed_group E] [normed_space ℝ E] [normed_group F] [normed_space ℝ F] {f : E → F} +lemma has_ftaylor_series_up_to_on.exists_lipschitz_on_with {E F : Type*} [normed_add_comm_group E] + [normed_space ℝ E] [normed_add_comm_group F] [normed_space ℝ F] {f : E → F} {p : E → formal_multilinear_series ℝ E F} {s : set E} {x : E} (hf : has_ftaylor_series_up_to_on 1 f p (insert x s)) (hs : convex ℝ s) : ∃ K (t ∈ 𝓝[s] x), lipschitz_on_with K f t := @@ -3055,8 +3056,8 @@ lemma has_ftaylor_series_up_to_on.exists_lipschitz_on_with {E F : Type*} /-- If `f` is `C^1` within a conves set `s` at `x`, then it is Lipschitz on a neighborhood of `x` within `s`. -/ -lemma cont_diff_within_at.exists_lipschitz_on_with {E F : Type*} [normed_group E] - [normed_space ℝ E] [normed_group F] [normed_space ℝ F] {f : E → F} {s : set E} +lemma cont_diff_within_at.exists_lipschitz_on_with {E F : Type*} [normed_add_comm_group E] + [normed_space ℝ E] [normed_add_comm_group F] [normed_space ℝ F] {f : E → F} {s : set E} {x : E} (hf : cont_diff_within_at ℝ 1 f s x) (hs : convex ℝ s) : ∃ (K : ℝ≥0) (t ∈ 𝓝[s] x), lipschitz_on_with K f t := begin diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 81e08abcf102b..15b565b864e92 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -91,8 +91,8 @@ open continuous_linear_map (smul_right smul_right_one_eq_iff) variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] section -variables {F : Type v} [normed_group F] [normed_space 𝕜 F] -variables {E : Type w} [normed_group E] [normed_space 𝕜 E] +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`. @@ -461,7 +461,8 @@ lemma differentiable_within_at_Ioi_iff_Ici [partial_order 𝕜] : ⟨λ 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_group E] [normed_space ℝ E] (f : ℝ → E) (x : ℝ) : +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, @@ -755,7 +756,7 @@ section pi /-! ### Derivatives of functions `f : 𝕜 → Π i, E i` -/ -variables {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_group (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 : @@ -1075,7 +1076,7 @@ end continuous section cartesian_product /-! ### Derivative of the cartesian product of two functions -/ -variables {G : Type w} [normed_group G] [normed_space 𝕜 G] +variables {G : Type w} [normed_add_comm_group G] [normed_space 𝕜 G] variables {f₂ : 𝕜 → G} {f₂' : G} lemma has_deriv_at_filter.prod @@ -1670,8 +1671,8 @@ section clm_comp_apply open continuous_linear_map -variables {G : Type*} [normed_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} +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) : @@ -1965,7 +1966,7 @@ end pow section zpow /-! ### Derivative of `x ↦ x^m` for `m : ℤ` -/ -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] {x : 𝕜} {s : set 𝕜} {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 := @@ -2089,7 +2090,7 @@ end zpow section support open function -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] {f : 𝕜 → F} +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {f : 𝕜 → F} lemma support_deriv_subset : support (deriv f) ⊆ tsupport f := begin @@ -2131,7 +2132,7 @@ section real_space open metric -variables {E : Type u} [normed_group E] [normed_space ℝ E] {f : ℝ → E} {f' : E} {s : set ℝ} +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 diff --git a/src/analysis/calculus/diff_on_int_cont.lean b/src/analysis/calculus/diff_on_int_cont.lean index 465b80892621c..bcb209415ff0c 100644 --- a/src/analysis/calculus/diff_on_int_cont.lean +++ b/src/analysis/calculus/diff_on_int_cont.lean @@ -16,9 +16,9 @@ this property and prove basic facts about this predicate. open set filter metric open_locale topological_space -variables (𝕜 : Type*) {E F G : Type*} [nontrivially_normed_field 𝕜] [normed_group E] - [normed_group F] [normed_space 𝕜 E] [normed_space 𝕜 F] [normed_group G] [normed_space 𝕜 G] - {f g : E → F} {s t : set E} {x : E} +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] + [normed_space 𝕜 G] {f g : E → F} {s t : set E} {x : E} /-- A predicate saying that a function is differentiable on a set and is continuous on its closure. This is a common assumption in complex analysis. -/ diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 0d4dbf15b1e2c..2f20481ada53d 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -20,7 +20,7 @@ differentiability. open_locale classical topological_space filter open function set filter -variables {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] +variables {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] /-- `dslope f a b` is defined as `slope f a b = (b - a)⁻¹ • (f b - f a)` for `a ≠ b` and `deriv f a` for `a = b`. -/ @@ -33,7 +33,7 @@ variables {f : 𝕜 → E} {a b : 𝕜} {s : set 𝕜} lemma dslope_of_ne (f : 𝕜 → E) (h : b ≠ a) : dslope f a b = slope f a b := update_noteq h _ _ -lemma continuous_linear_map.dslope_comp {F : Type*} [normed_group F] [normed_space 𝕜 F] +lemma continuous_linear_map.dslope_comp {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] (f : E →L[𝕜] F) (g : 𝕜 → E) (a b : 𝕜) (H : a = b → differentiable_at 𝕜 g a) : dslope (f ∘ g) a b = f (dslope g a b) := begin diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 2b1ef3254164b..0d87c20369832 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -20,8 +20,8 @@ of the one-dimensional derivative `deriv ℝ f`. -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] - {F : Type*} [normed_group F] [normed_space ℝ F] +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 diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 39b841409d071..bee2e8b8abe3a 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -122,10 +122,10 @@ noncomputable theory section variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] -variables {G : Type*} [normed_group G] [normed_space 𝕜 G] -variables {G' : Type*} [normed_group G'] [normed_space 𝕜 G'] +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'] /-- A function `f` has the continuous linear map `f'` as derivative along the filter `L` if `f x' = f x + f' (x' - x) + o (x' - x)` when `x'` converges along the filter `L`. This definition @@ -1771,7 +1771,7 @@ theorem: differentiability of `Φ`. -/ -variables {ι : Type*} [fintype ι] {F' : ι → Type*} [Π i, normed_group (F' i)] +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} @@ -2165,7 +2165,7 @@ end bilinear_map section clm_comp_apply /-! ### Derivative of the pointwise composition/application of continuous linear maps -/ -variables {H : Type*} [normed_group H] [normed_space 𝕜 H] {c : E → G →L[𝕜] H} +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} @@ -2887,8 +2887,8 @@ section -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] -variables {F : Type*} [normed_group F] [normed_space ℝ F] +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} : @@ -2914,8 +2914,8 @@ end section tangent_cone variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{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 @@ -2980,9 +2980,9 @@ 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_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] variables [is_scalar_tower 𝕜 𝕜' E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] [normed_space 𝕜' F] +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} @@ -3054,8 +3054,8 @@ end restrict_scalars section support open function -variables (𝕜 : Type*) {E F : Type*} [nontrivially_normed_field 𝕜] -variables [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] {f : E → F} +variables (𝕜 : Type*) {E F : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] + [normed_space 𝕜 E] [normed_add_comm_group F] [normed_space 𝕜 F] {f : E → F} lemma support_fderiv_subset : support (fderiv 𝕜 f) ⊆ tsupport f := begin diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index c960d0386e654..ce82fb6b3c0ca 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -18,8 +18,8 @@ open filter asymptotics open_locale ennreal variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] section fderiv diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index ad43c5975cc4d..192e6c4a9a955 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -81,7 +81,7 @@ open_locale topological_space namespace continuous_linear_map variables {𝕜 E F : Type*} [nontrivially_normed_field 𝕜] - [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] + [normed_add_comm_group E] [normed_space 𝕜 E] [normed_add_comm_group F] [normed_space 𝕜 F] lemma measurable_apply₂ [measurable_space E] [opens_measurable_space E] [second_countable_topology E] [second_countable_topology (E →L[𝕜] F)] @@ -94,8 +94,8 @@ end continuous_linear_map section fderiv variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] +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} (K : set (E →L[𝕜] F)) namespace fderiv_measurable_aux @@ -432,7 +432,7 @@ end fderiv section right_deriv -variables {F : Type*} [normed_group F] [normed_space ℝ F] +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] variables {f : ℝ → F} (K : set F) namespace right_deriv_measurable_aux diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index a4a4dab9b1e29..41b99929f84cc 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -50,8 +50,8 @@ rectangle are contained in `s` by convexity. The general case follows by lineari open asymptotics set open_locale topological_space -variables {E F : Type*} [normed_group E] [normed_space ℝ E] -[normed_group F] [normed_space ℝ F] +variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] +[normed_add_comm_group F] [normed_space ℝ F] {s : set E} (s_conv : convex ℝ s) {f : E → F} {f' : E → (E →L[ℝ] F)} {f'' : E →L[ℝ] (E →L[ℝ] F)} (hf : ∀ x ∈ interior s, has_fderiv_at f (f' x) x) diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index 66f0ee7693468..d08591e5e5873 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -115,9 +115,9 @@ end namespace formal_multilinear_series variables [nontrivially_normed_field 𝕜] - [normed_group E] [normed_space 𝕜 E] - [normed_group F] [normed_space 𝕜 F] - [normed_group G] [normed_space 𝕜 G] + [normed_add_comm_group E] [normed_space 𝕜 E] + [normed_add_comm_group F] [normed_space 𝕜 F] + [normed_add_comm_group G] [normed_space 𝕜 G] variables (p : formal_multilinear_series 𝕜 E F) diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 68558d8d73bb9..dfeac8ee6776b 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -89,9 +89,9 @@ such that * the kernels of the derivatives are complementary subspaces of `E`. -/ @[nolint has_inhabited_instance] structure implicit_function_data (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] [complete_space E] - (F : Type*) [normed_group F] [normed_space 𝕜 F] [complete_space F] - (G : Type*) [normed_group G] [normed_space 𝕜 G] [complete_space G] := + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] [complete_space E] + (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] [complete_space F] + (G : Type*) [normed_add_comm_group G] [normed_space 𝕜 G] [complete_space G] := (left_fun : E → F) (left_deriv : E →L[𝕜] F) (right_fun : E → G) @@ -106,9 +106,9 @@ structure implicit_function_data (𝕜 : Type*) [nontrivially_normed_field 𝕜] namespace implicit_function_data variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] [complete_space F] - {G : Type*} [normed_group G] [normed_space 𝕜 G] [complete_space G] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [complete_space E] + {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] [complete_space F] + {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] [complete_space G] (φ : implicit_function_data 𝕜 E F G) /-- The function given by `x ↦ (left_fun x, right_fun x)`. -/ @@ -204,8 +204,8 @@ complementary to `ker f'` lead to different maps `φ`. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] [complete_space F] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [complete_space E] + {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] [complete_space F] {f : E → F} {f' : E →L[𝕜] F} {a : E} section defs @@ -336,8 +336,8 @@ complementary to `ker f'` lead to different maps `φ`. section finite_dimensional variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] [complete_space E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] [finite_dimensional 𝕜 F] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [complete_space E] + {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] [finite_dimensional 𝕜 F] (f : E → F) (f' : E →L[𝕜] F) {a : E} /-- Given a map `f : E → F` to a finite dimensional space with a surjective derivative `f'`, diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 247b5082505bf..5006b15c09371 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -62,10 +62,10 @@ open_locale topological_space classical nnreal noncomputable theory variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] -variables {G : Type*} [normed_group G] [normed_space 𝕜 G] -variables {G' : Type*} [normed_group G'] [normed_space 𝕜 G'] +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 {ε : ℝ} @@ -481,8 +481,8 @@ omit cs /-- In a real vector space, a function `f` that approximates a linear equivalence on a subset `s` can be extended to a homeomorphism of the whole space. -/ -lemma exists_homeomorph_extension {E : Type*} [normed_group E] [normed_space ℝ E] - {F : Type*} [normed_group F] [normed_space ℝ F] [finite_dimensional ℝ F] +lemma exists_homeomorph_extension {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [finite_dimensional ℝ F] {s : set E} {f : E → F} {f' : E ≃L[ℝ] F} {c : ℝ≥0} (hf : approximates_linear_on f (f' : E →L[ℝ] F) s c) (hc : subsingleton E ∨ lipschitz_extension_constant F * c < (∥(f'.symm : F →L[ℝ] E)∥₊)⁻¹) : @@ -725,8 +725,8 @@ is_open_map_iff_nhds_le.2 $ λ x, ((hf x).map_nhds_eq (h0 x)).ge namespace cont_diff_at variables {𝕂 : Type*} [is_R_or_C 𝕂] -variables {E' : Type*} [normed_group E'] [normed_space 𝕂 E'] -variables {F' : Type*} [normed_group F'] [normed_space 𝕂 F'] +variables {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕂 E'] +variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕂 F'] variables [complete_space E'] (f : E' → F') {f' : E' ≃L[𝕂] F'} {a : E'} /-- Given a `cont_diff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 51be7f72aa90a..10795867616d8 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -46,8 +46,8 @@ open filter asymptotics set variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] /-- The `n`-th iterated derivative of a function from `𝕜` to `F`, as a function from `𝕜` to `F`. -/ def iterated_deriv (n : ℕ) (f : 𝕜 → F) (x : 𝕜) : F := diff --git a/src/analysis/calculus/lagrange_multipliers.lean b/src/analysis/calculus/lagrange_multipliers.lean index 159827200298c..5c06edbf3880e 100644 --- a/src/analysis/calculus/lagrange_multipliers.lean +++ b/src/analysis/calculus/lagrange_multipliers.lean @@ -28,8 +28,8 @@ lagrange multiplier, local extremum open filter set open_locale topological_space filter big_operators -variables {E F : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] - [normed_group F] [normed_space ℝ F] [complete_space F] +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[ℝ] ℝ} /-- Lagrange multipliers theorem: if `φ : E → ℝ` has a local extremum on the set `{x | f x = f x₀}` diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index 0b78a246c8959..5cec762b19290 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -69,7 +69,7 @@ open_locale topological_space classical polynomial section module -variables {E : Type u} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {a : E} +variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {a : E} {f' : E →L[ℝ] ℝ} /-- "Positive" tangent cone to `s` at `x`; the only difference from `tangent_cone_at` diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index cbb4351b87552..3a9e0bc673cae 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -64,8 +64,8 @@ In this file we prove the following facts: -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] - {F : Type*} [normed_group F] [normed_space ℝ F] +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 @@ -251,8 +251,8 @@ Let `f` and `B` be continuous functions on `[a, b]` such that * we have `f' x < B' x` whenever `∥f x∥ = B x`. Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. -/ -lemma image_norm_le_of_liminf_right_slope_norm_lt_deriv_boundary {E : Type*} [normed_group E] - {f : ℝ → E} {f' : ℝ → ℝ} (hf : continuous_on f (Icc a b)) +lemma image_norm_le_of_liminf_right_slope_norm_lt_deriv_boundary {E : Type*} + [normed_add_comm_group E] {f : ℝ → E} {f' : ℝ → ℝ} (hf : continuous_on f (Icc a b)) -- `hf'` actually says `liminf (∥f z∥ - ∥f x∥) / (z - x) ≤ f' x` (hf' : ∀ x ∈ Ico a b, ∀ r, f' x < r → ∃ᶠ z in 𝓝[>] x, slope (norm ∘ f) x z < r) @@ -471,7 +471,8 @@ also assume `[normed_space ℝ E]` to have a notion of a `convex` set. -/ section -variables {𝕜 G : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] [normed_group G] [normed_space 𝕜 G] +variables {𝕜 G : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] [normed_add_comm_group G] + [normed_space 𝕜 G] namespace convex @@ -1367,8 +1368,8 @@ make sense and are enough. Many formulations of the mean value inequality could balls over `ℝ` or `ℂ`. For now, we only include the ones that we need. -/ -variables {𝕜 : Type*} [is_R_or_C 𝕜] {G : Type*} [normed_group G] [normed_space 𝕜 G] - {H : Type*} [normed_group H] [normed_space 𝕜 H] {f : G → H} {f' : G → G →L[𝕜] H} {x : G} +variables {𝕜 : Type*} [is_R_or_C 𝕜] {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] + {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] {f : G → H} {f' : G → G →L[𝕜] H} {x : G} /-- Over the reals or the complexes, a continuously differentiable function is strictly differentiable. -/ diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index 1d52b259e2857..dd7ed194fc022 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -57,9 +57,9 @@ open topological_space measure_theory filter metric open_locale topological_space filter variables {α : Type*} [measurable_space α] {μ : measure α} {𝕜 : Type*} [is_R_or_C 𝕜] - {E : Type*} [normed_group E] [normed_space ℝ E] [normed_space 𝕜 E] + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_space 𝕜 E] [complete_space E] - {H : Type*} [normed_group H] [normed_space 𝕜 H] + {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] /-- Differentiation under integral of `x ↦ ∫ F x a` at a given point `x₀`, assuming `F x₀` is integrable, `∥F x a - F x₀ a∥ ≤ bound a * ∥x - x₀∥` for `x` in a ball around `x₀` for ae `a` with diff --git a/src/analysis/calculus/parametric_interval_integral.lean b/src/analysis/calculus/parametric_interval_integral.lean index b5cfe97ee45fa..1b8e62cef1690 100644 --- a/src/analysis/calculus/parametric_interval_integral.lean +++ b/src/analysis/calculus/parametric_interval_integral.lean @@ -17,9 +17,9 @@ open topological_space measure_theory filter metric open_locale topological_space filter interval variables {𝕜 : Type*} [is_R_or_C 𝕜] {μ : measure ℝ} - {E : Type*} [normed_group E] [normed_space ℝ E] [normed_space 𝕜 E] + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_space 𝕜 E] [complete_space E] - {H : Type*} [normed_group H] [normed_space 𝕜 H] + {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] {a b ε : ℝ} {bound : ℝ → ℝ} namespace interval_integral diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/specific_functions.lean index 8820b4018fe90..d02ba83483161 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/specific_functions.lean @@ -292,7 +292,7 @@ lemma R_pos {c : E} (f : cont_diff_bump_of_inner c) : 0 < f.R := f.r_pos.trans f instance (c : E) : inhabited (cont_diff_bump_of_inner c) := ⟨⟨1, 2, zero_lt_one, one_lt_two⟩⟩ -variables [inner_product_space ℝ E] [normed_group X] [normed_space ℝ X] +variables [inner_product_space ℝ E] [normed_add_comm_group X] [normed_space ℝ X] variables {c : E} (f : cont_diff_bump_of_inner c) {x : E} {n : with_top ℕ} /-- The function defined by `f : cont_diff_bump_of_inner c`. Use automatic coercion to @@ -467,12 +467,13 @@ a bundled smooth function such that 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_group E] [normed_space ℝ E] [finite_dimensional ℝ E] (c : E) +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_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {c x : E} +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 @@ -560,7 +561,7 @@ neighborhood `s` there exists an infinitely smooth function with the following p 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_group E] [normed_space ℝ E] +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 := diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index 7be5a67081f7c..f030aa993bdfd 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -63,9 +63,9 @@ def unique_diff_on (s : set E) : Prop := end tangent_cone -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] -variables {G : Type*} [normed_group G] [normed_space ℝ G] +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 {𝕜} {x y : E} {s t : set E} section tangent_cone @@ -183,7 +183,7 @@ end /-- The tangent cone of a product contains the tangent cone of each factor. -/ lemma maps_to_tangent_cone_pi {ι : Type*} [decidable_eq ι] {E : ι → Type*} - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] {s : Π i, set (E i)} {x : Π i, E i} {i : ι} (hi : ∀ j ≠ i, x j ∈ closure (s j)) : maps_to (linear_map.single i : E i →ₗ[𝕜] Π j, E j) (tangent_cone_at 𝕜 (s i) (x i)) (tangent_cone_at 𝕜 (set.pi univ s) x) := @@ -325,7 +325,7 @@ begin end lemma unique_diff_within_at.univ_pi (ι : Type*) [fintype ι] (E : ι → Type*) - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (x : Π i, E i) (h : ∀ i, unique_diff_within_at 𝕜 (s i) (x i)) : unique_diff_within_at 𝕜 (set.pi univ s) x := begin @@ -339,7 +339,7 @@ begin end lemma unique_diff_within_at.pi (ι : Type*) [fintype ι] (E : ι → Type*) - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (x : Π i, E i) (I : set ι) (h : ∀ i ∈ I, unique_diff_within_at 𝕜 (s i) (x i)) : unique_diff_within_at 𝕜 (set.pi I s) x := @@ -358,7 +358,7 @@ lemma unique_diff_on.prod {t : set F} (hs : unique_diff_on 𝕜 s) (ht : unique_ /-- The finite product of a family of sets of unique differentiability is a set of unique differentiability. -/ lemma unique_diff_on.pi (ι : Type*) [fintype ι] (E : ι → Type*) - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (I : set ι) (h : ∀ i ∈ I, unique_diff_on 𝕜 (s i)) : unique_diff_on 𝕜 (set.pi I s) := λ x hx, unique_diff_within_at.pi _ _ _ _ _ $ λ i hi, h i hi (x i) (hx i hi) @@ -366,7 +366,7 @@ lemma unique_diff_on.pi (ι : Type*) [fintype ι] (E : ι → Type*) /-- The finite product of a family of sets of unique differentiability is a set of unique differentiability. -/ lemma unique_diff_on.univ_pi (ι : Type*) [fintype ι] (E : ι → Type*) - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (h : ∀ i, unique_diff_on 𝕜 (s i)) : unique_diff_on 𝕜 (set.pi univ s) := unique_diff_on.pi _ _ _ _ $ λ i _, h i diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index 9f82194df62d3..9c42071ab16b7 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -81,8 +81,8 @@ open topological_space metric set filter asymptotics function measure_theory aff open_locale topological_space filter nnreal real universes u v w -variables {E : Type u} [normed_group E] [normed_space ℂ E] - {F : Type v} [normed_group F] [normed_space ℂ F] +variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] + {F : Type v} [normed_add_comm_group F] [normed_space ℂ F] local postfix `̂`:100 := uniform_space.completion diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 871f8a70623dd..fd4b701a61d4f 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -30,7 +30,6 @@ We also register the fact that `ℂ` is an `is_R_or_C` field. -/ noncomputable theory - namespace complex open_locale complex_conjugate topological_space @@ -39,8 +38,8 @@ instance : has_norm ℂ := ⟨abs⟩ @[simp] lemma norm_eq_abs (z : ℂ) : ∥z∥ = abs z := rfl -instance : normed_group ℂ := -normed_group.of_core ℂ +instance : normed_add_comm_group ℂ := +normed_add_comm_group.of_core ℂ { norm_eq_zero_iff := λ z, abs_eq_zero, triangle := abs_add, norm_neg := abs_neg } @@ -49,7 +48,7 @@ instance : normed_field ℂ := { norm := abs, dist_eq := λ _ _, rfl, norm_mul' := abs_mul, - .. complex.field, .. complex.normed_group } + .. complex.field, .. complex.normed_add_comm_group } instance : nontrivially_normed_field ℂ := { non_trivial := ⟨2, by simp; norm_num⟩ } @@ -62,10 +61,11 @@ instance {R : Type*} [normed_field R] [normed_algebra R ℝ] : normed_algebra R end, to_algebra := complex.algebra } +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] + /-- The module structure from `module.complex_to_real` is a normed space. -/ @[priority 900] -- see Note [lower instance priority] -instance _root_.normed_space.complex_to_real {E : Type*} [normed_group E] [normed_space ℂ E] : - normed_space ℝ E := +instance _root_.normed_space.complex_to_real : normed_space ℝ E := normed_space.restrict_scalars ℝ ℂ E lemma dist_eq (z w : ℂ) : dist z w = abs (z - w) := rfl @@ -192,7 +192,7 @@ calc 1 = ∥im_clm I∥ : by simp @[simp] lemma im_clm_nnnorm : ∥im_clm∥₊ = 1 := subtype.ext im_clm_norm -lemma restrict_scalars_one_smul_right' {E : Type*} [normed_group E] [normed_space ℂ E] (x : E) : +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 := by { ext ⟨a, b⟩, simp [mk_eq_add_mul_I, add_smul, mul_smul, smul_comm I] } diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 33d2020f48624..d109234451380 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -149,7 +149,7 @@ noncomputable theory universes u -variables {E : Type u} [normed_group E] [normed_space ℂ E] [complete_space E] +variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] namespace complex diff --git a/src/analysis/complex/conformal.lean b/src/analysis/complex/conformal.lean index 9e8d060c0fb90..bc434f9a0eb91 100644 --- a/src/analysis/complex/conformal.lean +++ b/src/analysis/complex/conformal.lean @@ -39,7 +39,7 @@ conj_lie.to_linear_isometry.is_conformal_map section conformal_into_complex_normed -variables {E : Type*} [normed_group E] [normed_space ℝ E] [normed_space ℂ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_space ℂ E] {z : ℂ} {g : ℂ →L[ℝ] E} {f : ℂ → E} lemma is_conformal_map_complex_linear {map : ℂ →L[ℂ] E} (nonzero : map ≠ 0) : diff --git a/src/analysis/complex/liouville.lean b/src/analysis/complex/liouville.lean index 6ec6d0b88aca2..133d8fa3ea95b 100644 --- a/src/analysis/complex/liouville.lean +++ b/src/analysis/complex/liouville.lean @@ -24,8 +24,8 @@ open topological_space metric set filter asymptotics function measure_theory open_locale topological_space filter nnreal real universes u v -variables {E : Type u} [normed_group E] [normed_space ℂ E] - {F : Type v} [normed_group F] [normed_space ℂ F] +variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] + {F : Type v} [normed_add_comm_group F] [normed_space ℂ F] local postfix `̂`:100 := uniform_space.completion diff --git a/src/analysis/complex/phragmen_lindelof.lean b/src/analysis/complex/phragmen_lindelof.lean index f243effb84056..afcff7594f579 100644 --- a/src/analysis/complex/phragmen_lindelof.lean +++ b/src/analysis/complex/phragmen_lindelof.lean @@ -53,7 +53,7 @@ namespace phragmen_lindelof ### Auxiliary lemmas -/ -variables {E : Type*} [normed_group E] +variables {E : Type*} [normed_add_comm_group E] /-- An auxiliary lemma that combines two double exponential estimates into a similar estimate on the difference of the functions. -/ diff --git a/src/analysis/complex/real_deriv.lean b/src/analysis/complex/real_deriv.lean index 6dd836813f070..9f40a3dcf68f6 100644 --- a/src/analysis/complex/real_deriv.lean +++ b/src/analysis/complex/real_deriv.lean @@ -79,7 +79,7 @@ theorem cont_diff.real_of_complex {n : with_top ℕ} (h : cont_diff ℂ n e) : cont_diff_iff_cont_diff_at.2 $ λ x, h.cont_diff_at.real_of_complex -variables {E : Type*} [normed_group E] [normed_space ℂ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] lemma has_strict_deriv_at.complex_to_real_fderiv' {f : ℂ → E} {x : ℂ} {f' : E} (h : has_strict_deriv_at f f' x) : @@ -122,7 +122,7 @@ section conformality open complex continuous_linear_map open_locale complex_conjugate -variables {E : Type*} [normed_group E] [normed_space ℂ E] {z : ℂ} {f : ℂ → E} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {z : ℂ} {f : ℂ → E} /-- A real differentiable function of the complex plane into some complex normed space `E` is conformal at a point `z` if it is holomorphic at that point with a nonvanishing differential. diff --git a/src/analysis/complex/removable_singularity.lean b/src/analysis/complex/removable_singularity.lean index 245b0a0ab5b51..816caed16fc9a 100644 --- a/src/analysis/complex/removable_singularity.lean +++ b/src/analysis/complex/removable_singularity.lean @@ -20,7 +20,7 @@ open topological_space metric set filter asymptotics function open_locale topological_space filter nnreal universe u -variables {E : Type u} [normed_group E] [normed_space ℂ E] [complete_space E] +variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] namespace complex diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index ae1002ae85c86..36e97b00f016a 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -53,7 +53,8 @@ namespace complex section space -variables {E : Type*} [normed_group E] [normed_space ℂ E] {R R₁ R₂ : ℝ} {f : ℂ → E} {c z : ℂ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {R R₁ R₂ : ℝ} {f : ℂ → E} + {c z : ℂ} /-- An auxiliary lemma for `complex.norm_dslope_le_div_of_maps_to_ball`. -/ lemma schwarz_aux {f : ℂ → ℂ} (hd : differentiable_on ℂ f (ball c R₁)) diff --git a/src/analysis/convex/exposed.lean b/src/analysis/convex/exposed.lean index 565f0d6d47336..fe68eb7cfffb2 100644 --- a/src/analysis/convex/exposed.lean +++ b/src/analysis/convex/exposed.lean @@ -45,7 +45,7 @@ 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 𝕜] [normed_group E] +variables (𝕜 : Type*) {E : Type*} [normed_linear_ordered_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] {l : E →L[𝕜] 𝕜} {A B C : set E} {X : finset E} {x : E} /-- A set `B` is exposed with respect to `A` iff it maximizes some functional over `A` (and contains diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 3236b6575c0f5..33907add5532f 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -416,7 +416,7 @@ lemma seminorm.gauge_seminorm_ball (p : seminorm ℝ E) : end add_comm_group section norm -variables [semi_normed_group E] [normed_space ℝ E] {s : set E} {r : ℝ} {x : E} +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 diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index f9abe177f1bfd..bcef36ed9e9f8 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -38,8 +38,8 @@ open measure_theory measure_theory.measure metric set filter topological_space f open_locale topological_space big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} - [normed_group E] [normed_space ℝ E] [complete_space E] - [normed_group F] [normed_space ℝ F] [complete_space F] + [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} {t : set α} {f : α → E} {g : E → ℝ} {C : ℝ} /-! diff --git a/src/analysis/convex/krein_milman.lean b/src/analysis/convex/krein_milman.lean index 9069f54d51ec0..e1f8047f531c4 100644 --- a/src/analysis/convex/krein_milman.lean +++ b/src/analysis/convex/krein_milman.lean @@ -56,7 +56,7 @@ See chapter 8 of [Barry Simon, *Convexity*][simon2011] open set open_locale classical -variables {E : Type*} [normed_group E] [normed_space ℝ E] {s : set E} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {s : set E} /-- **Krein-Milman lemma**: In a LCTVS (currently only in normed `ℝ`-spaces), any nonempty compact set has an extreme point. -/ diff --git a/src/analysis/convex/measure.lean b/src/analysis/convex/measure.lean index cf1486d21fdaa..3a82f562c9094 100644 --- a/src/analysis/convex/measure.lean +++ b/src/analysis/convex/measure.lean @@ -18,8 +18,8 @@ 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 -variables {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] {s : set E} +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} namespace convex diff --git a/src/analysis/convex/strict_convex_space.lean b/src/analysis/convex/strict_convex_space.lean index 293808bc66e2e..bd1f6c3715ac3 100644 --- a/src/analysis/convex/strict_convex_space.lean +++ b/src/analysis/convex/strict_convex_space.lean @@ -61,12 +61,12 @@ require balls of positive radius with center at the origin to be strictly convex then prove that any closed ball is strictly convex in `strict_convex_closed_ball` below. See also `strict_convex_space.of_strict_convex_closed_unit_ball`. -/ -class strict_convex_space (𝕜 E : Type*) [normed_linear_ordered_field 𝕜] [normed_group E] +class strict_convex_space (𝕜 E : Type*) [normed_linear_ordered_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] : Prop := (strict_convex_closed_ball : ∀ r : ℝ, 0 < r → strict_convex 𝕜 (closed_ball (0 : E) r)) variables (𝕜 : Type*) {E : Type*} [normed_linear_ordered_field 𝕜] - [normed_group E] [normed_space 𝕜 E] + [normed_add_comm_group E] [normed_space 𝕜 E] /-- A closed ball in a strictly convex space is strictly convex. -/ lemma strict_convex_closed_ball [strict_convex_space 𝕜 E] (x : E) (r : ℝ) : @@ -246,7 +246,7 @@ by rw [norm_smul, real.norm_of_nonneg (one_div_nonneg.2 zero_le_two), ←inv_eq_ ←div_eq_inv_mul, div_lt_iff (@zero_lt_two ℝ _ _), mul_two, ←not_same_ray_iff_of_norm_eq h, not_same_ray_iff_norm_add_lt, h] -variables {F : Type*} [normed_group F] [normed_space ℝ F] +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] variables {PF : Type*} {PE : Type*} [metric_space PF] [metric_space PE] variables [normed_add_torsor F PF] [normed_add_torsor E PE] diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 5381e8564c8c4..a8cc9a81ebd30 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -300,7 +300,7 @@ end has_continuous_smul /-! ### Normed vector space -/ section normed_space -variables [semi_normed_group E] [normed_space ℝ E] {s t : set E} +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`. -/ diff --git a/src/analysis/convex/uniform.lean b/src/analysis/convex/uniform.lean index e196cce44555d..442e50d6e55a4 100644 --- a/src/analysis/convex/uniform.lean +++ b/src/analysis/convex/uniform.lean @@ -37,14 +37,14 @@ uniform bound. Namely, over the `x` and `y` of norm `1`, `∥x + y∥` is unifor by a constant `< 2` when `∥x - y∥` is uniformly bounded below by a positive constant. See also `uniform_convex_space.of_uniform_convex_closed_unit_ball`. -/ -class uniform_convex_space (E : Type*) [semi_normed_group E] : Prop := +class uniform_convex_space (E : Type*) [seminormed_add_comm_group E] : Prop := (uniform_convex : ∀ ⦃ε : ℝ⦄, 0 < ε → ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ∥x∥ = 1 → ∀ ⦃y⦄, ∥y∥ = 1 → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 - δ) variables {E : Type*} -section semi_normed_group -variables (E) [semi_normed_group E] [uniform_convex_space E] {ε : ℝ} +section seminormed_add_comm_group +variables (E) [seminormed_add_comm_group E] [uniform_convex_space E] {ε : ℝ} lemma exists_forall_sphere_dist_add_le_two_sub (hε : 0 < ε) : ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ∥x∥ = 1 → ∀ ⦃y⦄, ∥y∥ = 1 → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 - δ := @@ -116,9 +116,9 @@ begin exact this hxy, end -end semi_normed_group +end seminormed_add_comm_group -variables [normed_group E] [normed_space ℝ E] [uniform_convex_space E] +variables [normed_add_comm_group E] [normed_space ℝ E] [uniform_convex_space E] @[priority 100] -- See note [lower instance priority] instance uniform_convex_space.to_strict_convex_space : strict_convex_space ℝ E := diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 84fd5935c9834..83e485df2d100 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -81,8 +81,8 @@ open continuous_linear_map metric open_locale pointwise topological_space nnreal variables {𝕜 G E E' E'' F F' F'' : Type*} -variables [normed_group E] [normed_group E'] [normed_group E''] [normed_group F] -variables {f f' : G → E} {g g' : G → E'} {x x' : G} {y y' : E} +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} section nontrivially_normed_field @@ -563,9 +563,9 @@ by { rw [← convolution_flip], exact hcf.continuous_convolution_right_of_integr end comm_group -section normed_group +section normed_add_comm_group -variables [semi_normed_group G] +variables [seminormed_add_comm_group G] /-- Compute `(f ⋆ g) x₀` if the support of the `f` is within `metric.ball 0 R`, and `g` is constant on `metric.ball x₀ R`. @@ -680,7 +680,7 @@ begin .trans_lt (half_lt_self hε) end -end normed_group +end normed_add_comm_group namespace cont_diff_bump_of_inner @@ -725,7 +725,7 @@ lemma convolution_tendsto_right' {ι} {φ : ι → cont_diff_bump_of_inner (0 : begin refine convolution_tendsto_right (λ i, (φ i).nonneg_normed) (λ i, (φ i).integral_normed) _ hmg hcg, - rw [normed_group.tendsto_nhds_zero] at hφ, + rw [normed_add_comm_group.tendsto_nhds_zero] at hφ, rw [tendsto_small_sets_iff], intros t ht, rcases metric.mem_nhds_iff.mp ht with ⟨ε, hε, ht⟩, @@ -764,8 +764,8 @@ variables [measurable_space G] {μ : measure G} variables (L : E →L[𝕜] E' →L[𝕜] F) section assoc -variables [normed_group F'] [normed_space ℝ F'] [normed_space 𝕜 F'] [complete_space F'] -variables [normed_group F''] [normed_space ℝ F''] [normed_space 𝕜 F''] [complete_space F''] +variables [normed_add_comm_group F'] [normed_space ℝ F'] [normed_space 𝕜 F'] [complete_space F'] +variables [normed_add_comm_group F''] [normed_space ℝ F''] [normed_space 𝕜 F''] [complete_space F''] variables {k : G → E''} variables (L₂ : F →L[𝕜] E'' →L[𝕜] F') variables (L₃ : E →L[𝕜] F'' →L[𝕜] F') @@ -795,7 +795,7 @@ end end assoc -variables [normed_group G] [borel_space G] +variables [normed_add_comm_group G] [borel_space G] variables [second_countable_topology G] [sigma_compact_space G] lemma convolution_precompR_apply {g : G → E'' →L[𝕜] E'} diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index cb7d7fb7fe998..1157808b31d9a 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -99,13 +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_group E, normed_space 𝕜 E, has_inner 𝕜 E := + 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) (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_group +attribute [nolint dangerous_instance] inner_product_space.to_normed_add_comm_group -- note [is_R_or_C instance] /-! @@ -321,8 +321,8 @@ begin end /-- Normed group structure constructed from an `inner_product_space.core` structure -/ -def to_normed_group : normed_group F := -normed_group.of_core F +def to_normed_add_comm_group : normed_add_comm_group F := +normed_add_comm_group.of_core F { norm_eq_zero_iff := assume x, begin split, @@ -349,7 +349,7 @@ normed_group.of_core F end, norm_neg := λ x, by simp only [norm, inner_neg_left, neg_neg, inner_neg_right] } -local attribute [instance] to_normed_group +local attribute [instance] to_normed_add_comm_group /-- Normed space structure constructed from a `inner_product_space.core` structure -/ def to_normed_space : normed_space 𝕜 F := @@ -369,7 +369,8 @@ the space into an inner product space, constructing the norm out of the inner pr 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_group F := @inner_product_space.of_core.to_normed_group 𝕜 F _ _ _ c, + 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 @@ -1948,7 +1949,8 @@ lemma orthogonal_family.summable_iff_norm_sq_summable [complete_space E] (f : Π summable (λ i, V i (f i)) ↔ summable (λ i, ∥f i∥ ^ 2) := begin classical, - simp only [summable_iff_cauchy_seq_finset, normed_group.cauchy_seq_iff, real.norm_eq_abs], + simp only [summable_iff_cauchy_seq_finset, normed_add_comm_group.cauchy_seq_iff, + real.norm_eq_abs], split, { intros hf ε hε, obtain ⟨a, H⟩ := hf _ (sqrt_pos.mpr hε), diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index b5882bb467ee2..35910f4e3d76b 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -52,7 +52,7 @@ cont_diff_inner.cont_diff_at lemma differentiable_inner : differentiable ℝ (λ p : E × E, ⟪p.1, p.2⟫) := is_bounded_bilinear_map_inner.differentiable_at -variables {G : Type*} [normed_group G] [normed_space ℝ G] +variables {G : Type*} [normed_add_comm_group G] [normed_space ℝ G] {f g : G → E} {f' g' : G →L[ℝ] E} {s : set G} {x : G} {n : with_top ℕ} include 𝕜 @@ -257,7 +257,7 @@ section pi_like open continuous_linear_map -variables {𝕜 ι H : Type*} [is_R_or_C 𝕜] [normed_group H] [normed_space 𝕜 H] +variables {𝕜 ι H : Type*} [is_R_or_C 𝕜] [normed_add_comm_group H] [normed_space 𝕜 H] [fintype ι] {f : H → euclidean_space 𝕜 ι} {f' : H →L[𝕜] euclidean_space 𝕜 ι} {t : set H} {y : H} lemma differentiable_within_at_euclidean : diff --git a/src/analysis/inner_product_space/euclidean_dist.lean b/src/analysis/inner_product_space/euclidean_dist.lean index 3fbd0fbeb4f4d..bfdd60a7cc2eb 100644 --- a/src/analysis/inner_product_space/euclidean_dist.lean +++ b/src/analysis/inner_product_space/euclidean_dist.lean @@ -20,7 +20,7 @@ distance. This way we hide the usage of `to_euclidean` behind an API. open_locale topological_space open set -variables {E : Type*} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] noncomputable theory @@ -106,7 +106,7 @@ nhds_basis_ball.mem_of_mem hr end euclidean -variables {F : Type*} [normed_group F] [normed_space ℝ F] {f g : F → E} {n : with_top ℕ} +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] {f g : F → E} {n : with_top ℕ} lemma cont_diff.euclidean_dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) (h : ∀ x, f x ≠ g x) : diff --git a/src/analysis/locally_convex/balanced_core_hull.lean b/src/analysis/locally_convex/balanced_core_hull.lean index 2ba576738c4ce..65e6cf2b31a42 100644 --- a/src/analysis/locally_convex/balanced_core_hull.lean +++ b/src/analysis/locally_convex/balanced_core_hull.lean @@ -223,7 +223,8 @@ begin { have h : filter.tendsto (λ (x : 𝕜 × E), x.fst • x.snd) (𝓝 (0,0)) (𝓝 0), from continuous_smul.tendsto' (0, 0) _ (smul_zero _), simpa only [← prod.exists', ← prod.forall', ← and_imp, ← and.assoc, exists_prop] - using h.basis_left (normed_group.nhds_zero_basis_norm_lt.prod_nhds ((𝓝 _).basis_sets)) U hU }, + using h.basis_left (normed_add_comm_group.nhds_zero_basis_norm_lt.prod_nhds + ((𝓝 _).basis_sets)) U hU }, rcases normed_field.exists_norm_lt 𝕜 hr with ⟨y, hy₀, hyr⟩, rw [norm_pos_iff] at hy₀, have : y • V ∈ 𝓝 (0 : E) := (set_smul_mem_nhds_zero_iff hy₀).mpr hV, diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index c47f7268e7070..c5b4ed4bbe977 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -194,7 +194,7 @@ end uniform_add_group section vonN_bornology_eq_metric -variables (𝕜 E) [nontrivially_normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] +variables (𝕜 E) [nontrivially_normed_field 𝕜] [seminormed_add_comm_group E] [normed_space 𝕜 E] namespace normed_space diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index ed3ad0f744078..3f1cba49ad904 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -314,7 +314,7 @@ end topological_add_group section normed_space /-- The topology of a `normed_space 𝕜 E` is induced by the seminorm `norm_seminorm 𝕜 E`. -/ -lemma norm_with_seminorms (𝕜 E) [normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] : +lemma norm_with_seminorms (𝕜 E) [normed_field 𝕜] [seminormed_add_comm_group E] [normed_space 𝕜 E] : with_seminorms (λ (_ : fin 1), norm_seminorm 𝕜 E) := begin let p : seminorm_family 𝕜 E (fin 1) := λ _, norm_seminorm 𝕜 E, @@ -419,7 +419,7 @@ begin rw ball_smul (s₁.sup p) hC, end -lemma cont_with_seminorms_normed_space (F) [semi_normed_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) : @@ -429,7 +429,7 @@ begin exact continuous_from_bounded hp (norm_with_seminorms 𝕜 F) f hf, end -lemma cont_normed_space_to_with_seminorms (E) [semi_normed_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 := @@ -467,7 +467,7 @@ end locally_convex_space section normed_space -variables (𝕜) [normed_field 𝕜] [normed_space ℝ 𝕜] [semi_normed_group E] +variables (𝕜) [normed_field 𝕜] [normed_space ℝ 𝕜] [seminormed_add_comm_group E] /-- Not an instance since `𝕜` can't be inferred. See `normed_space.to_locally_convex_space` for a slightly weaker instance version. -/ diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index 2ce178ba29b1d..7c794ed523c4c 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -14,22 +14,22 @@ In this file we provide the following non-instances for norms on matrices: * The elementwise norm: - * `matrix.semi_normed_group` - * `matrix.normed_group` + * `matrix.seminormed_add_comm_group` + * `matrix.normed_add_comm_group` * `matrix.normed_space` * The Frobenius norm: - * `matrix.frobenius_semi_normed_group` - * `matrix.frobenius_normed_group` + * `matrix.frobenius_seminormed_add_comm_group` + * `matrix.frobenius_normed_add_comm_group` * `matrix.frobenius_normed_space` * `matrix.frobenius_normed_ring` * `matrix.frobenius_normed_algebra` * The $L^\infty$ operator norm: - * `matrix.linfty_op_semi_normed_group` - * `matrix.linfty_op_normed_group` + * `matrix.linfty_op_seminormed_add_comm_group` + * `matrix.linfty_op_normed_add_comm_group` * `matrix.linfty_op_normed_space` * `matrix.linfty_op_non_unital_semi_normed_ring` * `matrix.linfty_op_semi_normed_ring` @@ -53,16 +53,16 @@ variables {R l m n α β : Type*} [fintype l] [fintype m] [fintype n] section linf_linf -section semi_normed_group -variables [semi_normed_group α] [semi_normed_group β] +section seminormed_add_comm_group +variables [seminormed_add_comm_group α] [seminormed_add_comm_group β] /-- Seminormed group instance (using sup norm of sup norm) for matrices over a seminormed group. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ -protected def semi_normed_group : semi_normed_group (matrix m n α) := -pi.semi_normed_group +protected def seminormed_add_comm_group : seminormed_add_comm_group (matrix m n α) := +pi.seminormed_add_comm_group -local attribute [instance] matrix.semi_normed_group +local attribute [instance] matrix.seminormed_add_comm_group lemma norm_le_iff {r : ℝ} (hr : 0 ≤ r) {A : matrix m n α} : ∥A∥ ≤ r ↔ ∀ i j, ∥A i j∥ ≤ r := @@ -136,18 +136,19 @@ instance [nonempty n] [decidable_eq n] [has_one α] [norm_one_class α] : norm_one_class (matrix n n α) := ⟨(norm_diagonal _).trans $ norm_one⟩ -end semi_normed_group +end seminormed_add_comm_group /-- Normed group instance (using sup norm of sup norm) for matrices over a normed group. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ -protected def normed_group [normed_group α] : normed_group (matrix m n α) := -pi.normed_group +protected def normed_add_comm_group [normed_add_comm_group α] : + normed_add_comm_group (matrix m n α) := +pi.normed_add_comm_group section normed_space -local attribute [instance] matrix.semi_normed_group +local attribute [instance] matrix.seminormed_add_comm_group -variables [normed_field R] [semi_normed_group α] [normed_space R α] +variables [normed_field R] [seminormed_add_comm_group α] [normed_space R α] /-- Normed space instance (using sup norm of sup norm) for matrices over a normed space. Not declared as an instance because there are several natural choices for defining the norm of a @@ -172,28 +173,29 @@ section linfty_op declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -protected def linfty_op_semi_normed_group [semi_normed_group α] : - semi_normed_group (matrix m n α) := -(by apply_instance : semi_normed_group (m → pi_Lp 1 (λ j : n, α))) +protected def linfty_op_seminormed_add_comm_group [seminormed_add_comm_group α] : + seminormed_add_comm_group (matrix m n α) := +(by apply_instance : seminormed_add_comm_group (m → pi_Lp 1 (λ j : n, α))) /-- Normed group instance (using sup norm of L1 norm) for matrices over a normed ring. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -protected def linfty_op_normed_group [normed_group α] : - normed_group (matrix m n α) := -(by apply_instance : normed_group (m → pi_Lp 1 (λ j : n, α))) +protected def linfty_op_normed_add_comm_group [normed_add_comm_group α] : + normed_add_comm_group (matrix m n α) := +(by apply_instance : normed_add_comm_group (m → pi_Lp 1 (λ j : n, α))) /-- Normed space instance (using sup norm of L1 norm) for matrices over a normed space. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -protected def linfty_op_normed_space [normed_field R] [semi_normed_group α] [normed_space R α] : +protected def linfty_op_normed_space [normed_field R] [seminormed_add_comm_group α] + [normed_space R α] : normed_space R (matrix m n α) := (by apply_instance : normed_space R (m → pi_Lp 1 (λ j : n, α))) -section semi_normed_group -variables [semi_normed_group α] +section seminormed_add_comm_group +variables [seminormed_add_comm_group α] lemma linfty_op_norm_def (A : matrix m n α) : ∥A∥ = ((finset.univ : finset m).sup (λ i : m, ∑ j : n, ∥A i j∥₊) : ℝ≥0) := @@ -238,7 +240,7 @@ lemma linfty_op_norm_diagonal [decidable_eq m] (v : m → α) : ∥diagonal v∥ = ∥v∥ := congr_arg coe $ linfty_op_nnnorm_diagonal v -end semi_normed_group +end seminormed_add_comm_group section non_unital_semi_normed_ring variables [non_unital_semi_normed_ring α] @@ -280,7 +282,7 @@ local attribute [instance] protected def linfty_op_non_unital_semi_normed_ring [non_unital_semi_normed_ring α] : non_unital_semi_normed_ring (matrix n n α) := { norm_mul := linfty_op_norm_mul, - .. matrix.linfty_op_semi_normed_group, + .. matrix.linfty_op_seminormed_add_comm_group, .. matrix.non_unital_ring } /-- The `L₁-L∞` norm preserves one on non-empty matrices. Note this is safe as an instance, as it @@ -339,28 +341,28 @@ open_locale matrix big_operators declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -def frobenius_semi_normed_group [semi_normed_group α] : - semi_normed_group (matrix m n α) := -(by apply_instance : semi_normed_group (pi_Lp 2 (λ i : m, pi_Lp 2 (λ j : n, α)))) +def frobenius_seminormed_add_comm_group [seminormed_add_comm_group α] : + seminormed_add_comm_group (matrix m n α) := +(by apply_instance : seminormed_add_comm_group (pi_Lp 2 (λ i : m, pi_Lp 2 (λ j : n, α)))) /-- Normed group instance (using frobenius norm) for matrices over a normed group. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -def frobenius_normed_group [normed_group α] : - normed_group (matrix m n α) := -(by apply_instance : normed_group (pi_Lp 2 (λ i : m, pi_Lp 2 (λ j : n, α)))) +def frobenius_normed_add_comm_group [normed_add_comm_group α] : + normed_add_comm_group (matrix m n α) := +(by apply_instance : normed_add_comm_group (pi_Lp 2 (λ i : m, pi_Lp 2 (λ j : n, α)))) /-- Normed space instance (using frobenius norm) for matrices over a normed space. Not declared as an instance because there are several natural choices for defining the norm of a matrix. -/ local attribute [instance] -def frobenius_normed_space [normed_field R] [semi_normed_group α] [normed_space R α] : +def frobenius_normed_space [normed_field R] [seminormed_add_comm_group α] [normed_space R α] : normed_space R (matrix m n α) := (by apply_instance : normed_space R (pi_Lp 2 (λ i : m, pi_Lp 2 (λ j : n, α)))) -section semi_normed_group -variables [semi_normed_group α] [semi_normed_group β] +section seminormed_add_comm_group +variables [seminormed_add_comm_group α] [seminormed_add_comm_group β] lemma frobenius_nnnorm_def (A : matrix m n α) : ∥A∥₊ = (∑ i j, ∥A i j∥₊ ^ (2 : ℝ)) ^ (1/2 : ℝ) := @@ -421,9 +423,9 @@ end ∥diagonal v∥ = ∥(pi_Lp.equiv 2 _).symm v∥ := (congr_arg coe $ frobenius_nnnorm_diagonal v : _).trans rfl -end semi_normed_group +end seminormed_add_comm_group -lemma frobenius_nnnorm_one [decidable_eq n] [semi_normed_group α] [has_one α] : +lemma frobenius_nnnorm_one [decidable_eq n] [seminormed_add_comm_group α] [has_one α] : ∥(1 : matrix n n α)∥₊ = nnreal.sqrt (fintype.card n) * ∥(1 : α)∥₊:= begin refine (frobenius_nnnorm_diagonal _).trans _, @@ -460,7 +462,7 @@ local attribute [instance] def frobenius_normed_ring [decidable_eq m] : normed_ring (matrix m m α) := { norm := has_norm.norm, norm_mul := frobenius_norm_mul, - ..matrix.frobenius_semi_normed_group } + ..matrix.frobenius_seminormed_add_comm_group } /-- Normed algebra instance (using frobenius norm) for matrices over `ℝ` or `ℂ`. Not declared as an instance because there are several natural choices for defining the norm of a diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 38d1510fc40f0..6dc141b00ac35 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -96,7 +96,7 @@ instance normed_comm_ring.to_semi_normed_comm_ring [β : normed_comm_ring α] : instance : normed_comm_ring punit := { norm_mul := λ _ _, by simp, - ..punit.normed_group, + ..punit.normed_add_comm_group, ..punit.comm_ring, } /-- A mixin class with the axiom `∥1∥ = 1`. Many `normed_ring`s and all `normed_field`s satisfy this @@ -108,10 +108,12 @@ export norm_one_class (norm_one) attribute [simp] norm_one -@[simp] lemma nnnorm_one [semi_normed_group α] [has_one α] [norm_one_class α] : ∥(1 : α)∥₊ = 1 := +@[simp] lemma nnnorm_one [seminormed_add_comm_group α] [has_one α] [norm_one_class α] : + ∥(1 : α)∥₊ = 1 := nnreal.eq norm_one -lemma norm_one_class.nontrivial (α : Type*) [semi_normed_group α] [has_one α] [norm_one_class α] : +lemma norm_one_class.nontrivial (α : Type*) [seminormed_add_comm_group α] [has_one α] + [norm_one_class α] : nontrivial α := nontrivial_of_ne 0 1 $ ne_of_apply_ne norm $ by simp @@ -119,23 +121,24 @@ nontrivial_of_ne 0 1 $ ne_of_apply_ne norm $ by simp instance semi_normed_comm_ring.to_comm_ring [β : semi_normed_comm_ring α] : comm_ring α := { ..β } @[priority 100] -- see Note [lower instance priority] -instance non_unital_normed_ring.to_normed_group [β : non_unital_normed_ring α] : normed_group α := +instance non_unital_normed_ring.to_normed_add_comm_group [β : non_unital_normed_ring α] : + normed_add_comm_group α := { ..β } @[priority 100] -- see Note [lower instance priority] -instance non_unital_semi_normed_ring.to_semi_normed_group [β : non_unital_semi_normed_ring α] : - semi_normed_group α := { ..β } +instance non_unital_semi_normed_ring.to_seminormed_add_comm_group [non_unital_semi_normed_ring α] : + seminormed_add_comm_group α := { ..‹non_unital_semi_normed_ring α› } -instance [semi_normed_group α] [has_one α] [norm_one_class α] : norm_one_class (ulift α) := +instance [seminormed_add_comm_group α] [has_one α] [norm_one_class α] : norm_one_class (ulift α) := ⟨by simp [ulift.norm_def]⟩ -instance prod.norm_one_class [semi_normed_group α] [has_one α] [norm_one_class α] - [semi_normed_group β] [has_one β] [norm_one_class β] : +instance prod.norm_one_class [seminormed_add_comm_group α] [has_one α] [norm_one_class α] + [seminormed_add_comm_group β] [has_one β] [norm_one_class β] : norm_one_class (α × β) := ⟨by simp [prod.norm_def]⟩ instance pi.norm_one_class {ι : Type*} {α : ι → Type*} [nonempty ι] [fintype ι] - [Π i, semi_normed_group (α i)] [Π i, has_one (α i)] [∀ i, norm_one_class (α i)] : + [Π i, seminormed_add_comm_group (α i)] [Π i, has_one (α i)] [∀ i, norm_one_class (α i)] : norm_one_class (Π i, α i) := ⟨by simp [pi.norm_def, finset.sup_const finset.univ_nonempty]⟩ @@ -178,7 +181,7 @@ lemma mul_right_bound (x : α) : instance : non_unital_semi_normed_ring (ulift α) := { norm_mul := λ x y, (norm_mul_le x.down y.down : _), - .. ulift.semi_normed_group } + .. ulift.seminormed_add_comm_group } /-- Non-unital seminormed ring structure on the product of two non-unital seminormed rings, using the sup norm. -/ @@ -195,7 +198,7 @@ instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : by apply max_mul_mul_le_max_mul_max; simp [norm_nonneg] ... = (max (∥x.1∥) (∥x.2∥)) * (max (∥y.1∥) (∥y.2∥)) : by simp [max_comm] ... = (∥x∥*∥y∥) : rfl, - ..prod.semi_normed_group } + ..prod.seminormed_add_comm_group } /-- Non-unital seminormed ring structure on the product of finitely many non-unital seminormed rings, using the sup norm. -/ @@ -208,7 +211,7 @@ instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] finset.sup_mono_fun $ λ b hb, norm_mul_le _ _ ... ≤ finset.univ.sup (λ i, ∥x i∥₊) * finset.univ.sup (λ i, ∥y i∥₊) : finset.sup_mul_le_mul_sup_of_nonneg _ (λ i _, zero_le _) (λ i _, zero_le _), - ..pi.semi_normed_group } + ..pi.seminormed_add_comm_group } end non_unital_semi_normed_ring @@ -222,7 +225,7 @@ See note [implicit instance arguments]. -/ instance subalgebra.semi_normed_ring {𝕜 : Type*} {_ : comm_ring 𝕜} {E : Type*} [semi_normed_ring E] {_ : algebra 𝕜 E} (s : subalgebra 𝕜 E) : semi_normed_ring s := { norm_mul := λ a b, norm_mul_le a.1 b.1, - ..s.to_submodule.semi_normed_group } + ..s.to_submodule.seminormed_add_comm_group } /-- A subalgebra of a normed ring is also a normed ring, with the restriction of the norm. @@ -308,21 +311,21 @@ eventually_at_top.mpr ⟨1, λ b h, norm_pow_le' a (nat.succ_le_iff.mp h)⟩ instance : semi_normed_ring (ulift α) := { .. ulift.non_unital_semi_normed_ring, - .. ulift.semi_normed_group } + .. ulift.seminormed_add_comm_group } /-- Seminormed ring structure on the product of two seminormed rings, using the sup norm. -/ instance prod.semi_normed_ring [semi_normed_ring β] : semi_normed_ring (α × β) := { ..prod.non_unital_semi_normed_ring, - ..prod.semi_normed_group, } + ..prod.seminormed_add_comm_group, } /-- Seminormed ring structure on the product of finitely many seminormed rings, using the sup norm. -/ instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, semi_normed_ring (π i)] : semi_normed_ring (Π i, π i) := { ..pi.non_unital_semi_normed_ring, - ..pi.semi_normed_group, } + ..pi.seminormed_add_comm_group, } end semi_normed_ring @@ -331,20 +334,20 @@ variables [non_unital_normed_ring α] instance : non_unital_normed_ring (ulift α) := { .. ulift.non_unital_semi_normed_ring, - .. ulift.semi_normed_group } + .. ulift.seminormed_add_comm_group } /-- Non-unital normed ring structure on the product of two non-unital normed rings, using the sup norm. -/ instance prod.non_unital_normed_ring [non_unital_normed_ring β] : non_unital_normed_ring (α × β) := { norm_mul := norm_mul_le, - ..prod.semi_normed_group } + ..prod.seminormed_add_comm_group } /-- Normed ring structure on the product of finitely many non-unital normed rings, using the sup norm. -/ instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_normed_ring (π i)] : non_unital_normed_ring (Π i, π i) := { norm_mul := norm_mul_le, - ..pi.normed_group } + ..pi.normed_add_comm_group } end non_unital_normed_ring @@ -360,18 +363,18 @@ x.norm_pos instance : normed_ring (ulift α) := { .. ulift.semi_normed_ring, - .. ulift.normed_group } + .. ulift.normed_add_comm_group } /-- Normed ring structure on the product of two normed rings, using the sup norm. -/ instance prod.normed_ring [normed_ring β] : normed_ring (α × β) := { norm_mul := norm_mul_le, - ..prod.normed_group } + ..prod.normed_add_comm_group } /-- Normed ring structure on the product of finitely many normed rings, using the sup norm. -/ instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π i)] : normed_ring (Π i, π i) := { norm_mul := norm_mul_le, - ..pi.normed_group } + ..pi.normed_add_comm_group } end normed_ring @@ -551,7 +554,7 @@ end normed_field instance : normed_field ℝ := { norm_mul' := abs_mul, - .. real.normed_group } + .. real.normed_add_comm_group } instance : nontrivially_normed_field ℝ := { non_trivial := ⟨2, by { unfold norm, rw abs_of_nonneg; norm_num }⟩ } @@ -608,24 +611,24 @@ nnreal.eq $ real.norm_of_nonneg x.2 end nnreal -@[simp] lemma norm_norm [semi_normed_group α] (x : α) : ∥∥x∥∥ = ∥x∥ := +@[simp] lemma norm_norm [seminormed_add_comm_group α] (x : α) : ∥∥x∥∥ = ∥x∥ := real.norm_of_nonneg (norm_nonneg _) -@[simp] lemma nnnorm_norm [semi_normed_group α] (a : α) : ∥∥a∥∥₊ = ∥a∥₊ := +@[simp] lemma nnnorm_norm [seminormed_add_comm_group α] (a : α) : ∥∥a∥∥₊ = ∥a∥₊ := by simpa [real.nnnorm_of_nonneg (norm_nonneg a)] /-- A restatement of `metric_space.tendsto_at_top` in terms of the norm. -/ -lemma normed_group.tendsto_at_top [nonempty α] [semilattice_sup α] {β : Type*} [semi_normed_group β] - {f : α → β} {b : β} : +lemma normed_add_comm_group.tendsto_at_top [nonempty α] [semilattice_sup α] {β : Type*} + [seminormed_add_comm_group β] {f : α → β} {b : β} : tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N ≤ n → ∥f n - b∥ < ε := (at_top_basis.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) /-- -A variant of `normed_group.tendsto_at_top` that +A variant of `normed_add_comm_group.tendsto_at_top` that uses `∃ N, ∀ n > N, ...` rather than `∃ N, ∀ n ≥ N, ...` -/ -lemma normed_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [no_max_order α] - {β : Type*} [semi_normed_group β] +lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [no_max_order α] + {β : Type*} [seminormed_add_comm_group β] {f : α → β} {b : β} : tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ∥f n - b∥ < ε := (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) @@ -672,7 +675,7 @@ 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`. section -variables [semi_normed_group α] +variables [seminormed_add_comm_group α] lemma norm_nsmul_le (n : ℕ) (a : α) : ∥n • a∥ ≤ n * ∥a∥ := begin diff --git a/src/analysis/normed/group/SemiNormedGroup.lean b/src/analysis/normed/group/SemiNormedGroup.lean index 4c1ec7aa17b98..cb87e2c2ba58f 100644 --- a/src/analysis/normed/group/SemiNormedGroup.lean +++ b/src/analysis/normed/group/SemiNormedGroup.lean @@ -22,30 +22,32 @@ universes u open category_theory /-- The category of seminormed abelian groups and bounded group homomorphisms. -/ -def SemiNormedGroup : Type (u+1) := bundled semi_normed_group +def SemiNormedGroup : Type (u+1) := bundled seminormed_add_comm_group namespace SemiNormedGroup -instance bundled_hom : bundled_hom @normed_group_hom := -⟨@normed_group_hom.to_fun, @normed_group_hom.id, @normed_group_hom.comp, @normed_group_hom.coe_inj⟩ +instance bundled_hom : bundled_hom @normed_add_group_hom := +⟨@normed_add_group_hom.to_fun, @normed_add_group_hom.id, @normed_add_group_hom.comp, + @normed_add_group_hom.coe_inj⟩ attribute [derive [large_category, concrete_category]] SemiNormedGroup instance : has_coe_to_sort SemiNormedGroup (Type u) := bundled.has_coe_to_sort /-- Construct a bundled `SemiNormedGroup` from the underlying type and typeclass. -/ -def of (M : Type u) [semi_normed_group M] : SemiNormedGroup := bundled.of M +def of (M : Type u) [seminormed_add_comm_group M] : SemiNormedGroup := bundled.of M -instance (M : SemiNormedGroup) : semi_normed_group M := M.str +instance (M : SemiNormedGroup) : seminormed_add_comm_group M := M.str -@[simp] lemma coe_of (V : Type u) [semi_normed_group V] : (SemiNormedGroup.of V : Type u) = V := rfl +@[simp] lemma coe_of (V : Type u) [seminormed_add_comm_group V] : + (SemiNormedGroup.of V : Type u) = V := rfl @[simp] lemma coe_id (V : SemiNormedGroup) : ⇑(𝟙 V) = id := rfl @[simp] lemma coe_comp {M N K : SemiNormedGroup} (f : M ⟶ N) (g : N ⟶ K) : ((f ≫ g) : M → K) = g ∘ f := rfl instance : inhabited SemiNormedGroup := ⟨of punit⟩ -instance of_unique (V : Type u) [semi_normed_group V] [i : unique V] : +instance of_unique (V : Type u) [seminormed_add_comm_group V] [i : unique V] : unique (SemiNormedGroup.of V) := i instance : limits.has_zero_morphisms.{u (u+1)} SemiNormedGroup := {} @@ -80,20 +82,21 @@ end SemiNormedGroup `SemiNormedGroup₁` is a type synonym for `SemiNormedGroup`, which we shall equip with the category structure consisting only of the norm non-increasing maps. -/ -def SemiNormedGroup₁ : Type (u+1) := bundled semi_normed_group +def SemiNormedGroup₁ : Type (u+1) := bundled seminormed_add_comm_group namespace SemiNormedGroup₁ instance : has_coe_to_sort SemiNormedGroup₁ (Type u) := bundled.has_coe_to_sort instance : large_category.{u} SemiNormedGroup₁ := -{ hom := λ X Y, { f : normed_group_hom X Y // f.norm_noninc }, - id := λ X, ⟨normed_group_hom.id X, normed_group_hom.norm_noninc.id⟩, - comp := λ X Y Z f g, ⟨(g : normed_group_hom Y Z).comp (f : normed_group_hom X Y), g.2.comp f.2⟩, } +{ hom := λ X Y, { f : normed_add_group_hom X Y // f.norm_noninc }, + id := λ X, ⟨normed_add_group_hom.id X, normed_add_group_hom.norm_noninc.id⟩, + comp := λ X Y Z f g, + ⟨(g : normed_add_group_hom Y Z).comp (f : normed_add_group_hom X Y), g.2.comp f.2⟩ } @[ext] lemma hom_ext {M N : SemiNormedGroup₁} (f g : M ⟶ N) (w : (f : M → N) = (g : M → N)) : f = g := -subtype.eq (normed_group_hom.ext (congr_fun w)) +subtype.eq (normed_add_group_hom.ext (congr_fun w)) instance : concrete_category.{u} SemiNormedGroup₁ := { forget := @@ -102,9 +105,9 @@ instance : concrete_category.{u} SemiNormedGroup₁ := forget_faithful := {} } /-- Construct a bundled `SemiNormedGroup₁` from the underlying type and typeclass. -/ -def of (M : Type u) [semi_normed_group M] : SemiNormedGroup₁ := bundled.of M +def of (M : Type u) [seminormed_add_comm_group M] : SemiNormedGroup₁ := bundled.of M -instance (M : SemiNormedGroup₁) : semi_normed_group M := M.str +instance (M : SemiNormedGroup₁) : seminormed_add_comm_group M := M.str /-- Promote a morphism in `SemiNormedGroup` to a morphism in `SemiNormedGroup₁`. -/ def mk_hom {M N : SemiNormedGroup} (f : M ⟶ N) (i : f.norm_noninc) : @@ -128,22 +131,22 @@ instance : has_forget₂ SemiNormedGroup₁ SemiNormedGroup := { obj := λ X, X, map := λ X Y f, f.1, }, } -@[simp] lemma coe_of (V : Type u) [semi_normed_group V] : (SemiNormedGroup₁.of V : Type u) = V := -rfl +@[simp] lemma coe_of (V : Type u) [seminormed_add_comm_group V] : + (SemiNormedGroup₁.of V : Type u) = V := rfl @[simp] lemma coe_id (V : SemiNormedGroup₁) : ⇑(𝟙 V) = id := rfl @[simp] lemma coe_comp {M N K : SemiNormedGroup₁} (f : M ⟶ N) (g : N ⟶ K) : ((f ≫ g) : M → K) = g ∘ f := rfl -- If `coe_fn_coe_base` fires before `coe_comp`, `coe_comp'` puts us back in normal form. @[simp] lemma coe_comp' {M N K : SemiNormedGroup₁} (f : M ⟶ N) (g : N ⟶ K) : - ((f ≫ g) : normed_group_hom M K) = (↑g : normed_group_hom N K).comp ↑f := rfl + ((f ≫ g) : normed_add_group_hom M K) = (↑g : normed_add_group_hom N K).comp ↑f := rfl instance : inhabited SemiNormedGroup₁ := ⟨of punit⟩ -instance of_unique (V : Type u) [semi_normed_group V] [i : unique V] : +instance of_unique (V : Type u) [seminormed_add_comm_group V] [i : unique V] : unique (SemiNormedGroup₁.of V) := i instance : limits.has_zero_morphisms.{u (u+1)} SemiNormedGroup₁ := -{ has_zero := λ X Y, { zero := ⟨0, normed_group_hom.norm_noninc.zero⟩, }, +{ has_zero := λ X Y, { zero := ⟨0, normed_add_group_hom.norm_noninc.zero⟩, }, comp_zero' := λ X Y f Z, by { ext, refl, }, zero_comp' := λ X Y Z f, by { ext, simp [coe_fn_coe_base'] } } diff --git a/src/analysis/normed/group/SemiNormedGroup/completion.lean b/src/analysis/normed/group/SemiNormedGroup/completion.lean index a6d79cf48c914..0baa630b179a1 100644 --- a/src/analysis/normed/group/SemiNormedGroup/completion.lean +++ b/src/analysis/normed/group/SemiNormedGroup/completion.lean @@ -34,7 +34,7 @@ noncomputable theory universe u -open uniform_space mul_opposite category_theory normed_group_hom +open uniform_space mul_opposite category_theory normed_add_group_hom namespace SemiNormedGroup @@ -60,9 +60,9 @@ lemma Completion.norm_incl_eq {V : SemiNormedGroup} {v : V} : ∥Completion.incl lemma Completion.map_norm_noninc {V W : SemiNormedGroup} {f : V ⟶ W} (hf : f.norm_noninc) : (Completion.map f).norm_noninc := -normed_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.2 $ - (normed_group_hom.norm_completion f).le.trans $ - normed_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.1 hf +normed_add_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.2 $ + (normed_add_group_hom.norm_completion f).le.trans $ + normed_add_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.1 hf /-- Given a normed group hom `V ⟶ W`, this defines the associated morphism from the completion of `V` to the completion of `W`. @@ -79,9 +79,9 @@ add_monoid_hom.mk' (category_theory.functor.map Completion) $ λ f g, instance : preadditive SemiNormedGroup.{u} := { hom_group := λ P Q, infer_instance, add_comp' := by { intros, ext, - simp only [normed_group_hom.add_apply, category_theory.comp_apply, map_add] }, + simp only [normed_add_group_hom.add_apply, category_theory.comp_apply, map_add] }, comp_add' := by { intros, ext, - simp only [normed_group_hom.add_apply, category_theory.comp_apply, map_add] } } + simp only [normed_add_group_hom.add_apply, category_theory.comp_apply, map_add] } } instance : functor.additive Completion := { map_add' := λ X Y, (Completion.map_hom _ _).map_add } @@ -97,10 +97,10 @@ def Completion.lift {V W : SemiNormedGroup} [complete_space W] [separated_space lemma Completion.lift_comp_incl {V W : SemiNormedGroup} [complete_space W] [separated_space W] (f : V ⟶ W) : Completion.incl ≫ (Completion.lift f) = f := -by { ext, apply normed_group_hom.extension_coe } +by { ext, apply normed_add_group_hom.extension_coe } lemma Completion.lift_unique {V W : SemiNormedGroup} [complete_space W] [separated_space W] (f : V ⟶ W) (g : Completion.obj V ⟶ W) : Completion.incl ≫ g = f → g = Completion.lift f := -λ h, (normed_group_hom.extension_unique _ (λ v, ((ext_iff.1 h) v).symm)).symm +λ h, (normed_add_group_hom.extension_unique _ (λ v, ((ext_iff.1 h) v).symm)).symm end SemiNormedGroup diff --git a/src/analysis/normed/group/SemiNormedGroup/kernels.lean b/src/analysis/normed/group/SemiNormedGroup/kernels.lean index 29102a8b0c423..7d9c71e257445 100644 --- a/src/analysis/normed/group/SemiNormedGroup/kernels.lean +++ b/src/analysis/normed/group/SemiNormedGroup/kernels.lean @@ -35,13 +35,13 @@ noncomputable theory def cokernel_cocone {X Y : SemiNormedGroup₁.{u}} (f : X ⟶ Y) : cofork f 0 := cofork.of_π (@SemiNormedGroup₁.mk_hom - _ (SemiNormedGroup.of (Y ⧸ (normed_group_hom.range f.1))) + _ (SemiNormedGroup.of (Y ⧸ (normed_add_group_hom.range f.1))) f.1.range.normed_mk - (normed_group_hom.is_quotient_quotient _).norm_le) + (normed_add_group_hom.is_quotient_quotient _).norm_le) begin ext, - simp only [comp_apply, limits.zero_comp, normed_group_hom.zero_apply, - SemiNormedGroup₁.mk_hom_apply, SemiNormedGroup₁.zero_apply, ←normed_group_hom.mem_ker, + simp only [comp_apply, limits.zero_comp, normed_add_group_hom.zero_apply, + SemiNormedGroup₁.mk_hom_apply, SemiNormedGroup₁.zero_apply, ←normed_add_group_hom.mem_ker, f.1.range.ker_normed_mk, f.1.mem_range], use x, refl, @@ -53,12 +53,12 @@ def cokernel_lift {X Y : SemiNormedGroup₁.{u}} (f : X ⟶ Y) (s : cokernel_cof begin fsplit, -- The lift itself: - { apply normed_group_hom.lift _ s.π.1, + { apply normed_add_group_hom.lift _ s.π.1, rintro _ ⟨b, rfl⟩, change (f ≫ s.π) b = 0, simp, }, -- The lift has norm at most one: - exact normed_group_hom.lift_norm_noninc _ _ _ s.π.2, + exact normed_add_group_hom.lift_norm_noninc _ _ _ s.π.2, end instance : has_cokernels SemiNormedGroup₁.{u} := @@ -68,13 +68,13 @@ instance : has_cokernels SemiNormedGroup₁.{u} := (cokernel_lift f) (λ s, begin ext, - apply normed_group_hom.lift_mk f.1.range, + apply normed_add_group_hom.lift_mk f.1.range, rintro _ ⟨b, rfl⟩, change (f ≫ s.π) b = 0, simp, end) (λ s m w, subtype.eq - (normed_group_hom.lift_unique f.1.range _ _ _ (congr_arg subtype.val w : _))), } } + (normed_add_group_hom.lift_unique f.1.range _ _ _ (congr_arg subtype.val w : _))), } } -- Sanity check example : has_cokernels SemiNormedGroup₁ := by apply_instance @@ -87,13 +87,13 @@ section equalizers_and_kernels /-- The equalizer cone for a parallel pair of morphisms of seminormed groups. -/ def fork {V W : SemiNormedGroup.{u}} (f g : V ⟶ W) : fork f g := -@fork.of_ι _ _ _ _ _ _ (of (f - g).ker) (normed_group_hom.incl (f - g).ker) $ +@fork.of_ι _ _ _ _ _ _ (of (f - g).ker) (normed_add_group_hom.incl (f - g).ker) $ begin ext v, have : v.1 ∈ (f - g).ker := v.2, - simpa only [normed_group_hom.incl_apply, pi.zero_apply, coe_comp, normed_group_hom.coe_zero, - subtype.val_eq_coe, normed_group_hom.mem_ker, - normed_group_hom.coe_sub, pi.sub_apply, sub_eq_zero] using this + simpa only [normed_add_group_hom.incl_apply, pi.zero_apply, coe_comp, + normed_add_group_hom.coe_zero, subtype.val_eq_coe, normed_add_group_hom.mem_ker, + normed_add_group_hom.coe_sub, pi.sub_apply, sub_eq_zero] using this end instance has_limit_parallel_pair {V W : SemiNormedGroup.{u}} (f g : V ⟶ W) : @@ -101,10 +101,10 @@ instance has_limit_parallel_pair {V W : SemiNormedGroup.{u}} (f g : V ⟶ W) : { exists_limit := nonempty.intro { cone := fork f g, is_limit := fork.is_limit.mk _ - (λ c, normed_group_hom.ker.lift (fork.ι c) _ $ - show normed_group_hom.comp_hom (f - g) c.ι = 0, + (λ c, normed_add_group_hom.ker.lift (fork.ι c) _ $ + show normed_add_group_hom.comp_hom (f - g) c.ι = 0, by { rw [add_monoid_hom.map_sub, add_monoid_hom.sub_apply, sub_eq_zero], exact c.condition }) - (λ c, normed_group_hom.ker.incl_comp_lift _ _ _) + (λ c, normed_add_group_hom.ker.incl_comp_lift _ _ _) (λ c g h, by { ext x, dsimp, rw ← h, refl }) } } instance : limits.has_equalizers.{u (u+1)} SemiNormedGroup := @@ -121,17 +121,17 @@ section cokernel /-- Auxiliary definition for `has_cokernels SemiNormedGroup`. -/ def cokernel_cocone {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) : cofork f 0 := @cofork.of_π _ _ _ _ _ _ - (SemiNormedGroup.of (Y ⧸ (normed_group_hom.range f))) + (SemiNormedGroup.of (Y ⧸ (normed_add_group_hom.range f))) f.range.normed_mk begin ext, - simp only [comp_apply, limits.zero_comp, normed_group_hom.zero_apply, - ←normed_group_hom.mem_ker, f.range.ker_normed_mk, f.mem_range, exists_apply_eq_apply], + simp only [comp_apply, limits.zero_comp, normed_add_group_hom.zero_apply, + ←normed_add_group_hom.mem_ker, f.range.ker_normed_mk, f.mem_range, exists_apply_eq_apply], end /-- Auxiliary definition for `has_cokernels SemiNormedGroup`. -/ def cokernel_lift {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) (s : cokernel_cofork f) : - (cokernel_cocone f).X ⟶ s.X := normed_group_hom.lift _ s.π + (cokernel_cocone f).X ⟶ s.X := normed_add_group_hom.lift _ s.π begin rintro _ ⟨b, rfl⟩, change (f ≫ s.π) b = 0, @@ -144,12 +144,12 @@ def is_colimit_cokernel_cocone {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) : is_colimit_aux _ (cokernel_lift f) (λ s, begin ext, - apply normed_group_hom.lift_mk f.range, + apply normed_add_group_hom.lift_mk f.range, rintro _ ⟨b, rfl⟩, change (f ≫ s.π) b = 0, simp, end) -(λ s m w, normed_group_hom.lift_unique f.range _ _ _ w) +(λ s m w, normed_add_group_hom.lift_unique f.range _ _ _ w) instance : has_cokernels SemiNormedGroup.{u} := { has_colimit := λ X Y f, has_colimit.mk @@ -252,8 +252,8 @@ begin end lemma is_quotient_explicit_cokernel_π {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) : -normed_group_hom.is_quotient (explicit_cokernel_π f) := -normed_group_hom.is_quotient_quotient _ +normed_add_group_hom.is_quotient (explicit_cokernel_π f) := +normed_add_group_hom.is_quotient_quotient _ lemma norm_noninc_explicit_cokernel_π {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) : (explicit_cokernel_π f).norm_noninc := @@ -264,16 +264,16 @@ open_locale nnreal lemma explicit_cokernel_desc_norm_le_of_norm_le {X Y Z : SemiNormedGroup.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) (c : ℝ≥0) (h : ∥ g ∥ ≤ c) : ∥ explicit_cokernel_desc w ∥ ≤ c := -normed_group_hom.lift_norm_le _ _ _ h +normed_add_group_hom.lift_norm_le _ _ _ h lemma explicit_cokernel_desc_norm_noninc {X Y Z : SemiNormedGroup.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} {cond : f ≫ g = 0} (hg : g.norm_noninc) : (explicit_cokernel_desc cond).norm_noninc := begin - refine normed_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.2 _, + refine normed_add_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.2 _, rw [← nnreal.coe_one], exact explicit_cokernel_desc_norm_le_of_norm_le cond 1 - (normed_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.1 hg) + (normed_add_group_hom.norm_noninc.norm_noninc_iff_norm_le_one.1 hg) end lemma explicit_cokernel_desc_comp_eq_zero {X Y Z W : SemiNormedGroup.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index ca0d5166bca5e..77b484b3d4dc4 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -19,22 +19,22 @@ open_locale nnreal topological_space open filter /-- A `normed_add_torsor V P` is a torsor of an additive seminormed group -action by a `semi_normed_group V` on points `P`. We bundle the pseudometric space +action by a `seminormed_add_comm_group V` on points `P`. We bundle the pseudometric space structure and require the distance to be the same as results from the norm (which in fact implies the distance yields a pseudometric space, but bundling just the distance and using an instance for the pseudometric space results in type class problems). -/ class normed_add_torsor (V : out_param $ Type*) (P : Type*) - [out_param $ semi_normed_group V] [pseudo_metric_space P] + [out_param $ seminormed_add_comm_group V] [pseudo_metric_space P] extends add_torsor V P := (dist_eq_norm' : ∀ (x y : P), dist x y = ∥(x -ᵥ y : V)∥) -variables {α V P : Type*} [semi_normed_group V] [pseudo_metric_space P] [normed_add_torsor V P] -variables {W Q : Type*} [normed_group W] [metric_space Q] [normed_add_torsor W Q] +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] -/-- A `semi_normed_group` is a `normed_add_torsor` over itself. -/ +/-- A `seminormed_add_comm_group` is a `normed_add_torsor` over itself. -/ @[priority 100] -instance semi_normed_group.to_normed_add_torsor : normed_add_torsor V V := +instance seminormed_add_comm_group.to_normed_add_torsor : normed_add_torsor V V := { dist_eq_norm' := dist_eq_norm } include V @@ -147,8 +147,8 @@ omit V /-- The pseudodistance defines a pseudometric space structure on the torsor. This is not an instance because it depends on `V` to define a `metric_space P`. -/ -def pseudo_metric_space_of_normed_group_of_add_torsor (V P : Type*) [semi_normed_group V] - [add_torsor V P] : pseudo_metric_space P := +def pseudo_metric_space_of_normed_add_comm_group_of_add_torsor (V P : Type*) + [seminormed_add_comm_group V] [add_torsor V P] : pseudo_metric_space P := { dist := λ x y, ∥(x -ᵥ y : V)∥, dist_self := λ x, by simp, dist_comm := λ x y, by simp only [←neg_vsub_eq_vsub_rev y x, norm_neg], @@ -162,7 +162,8 @@ def pseudo_metric_space_of_normed_group_of_add_torsor (V P : Type*) [semi_normed /-- The distance defines a metric space structure on the torsor. This is not an instance because it depends on `V` to define a `metric_space P`. -/ -def metric_space_of_normed_group_of_add_torsor (V P : Type*) [normed_group V] [add_torsor V P] : +def metric_space_of_normed_add_comm_group_of_add_torsor (V P : Type*) + [normed_add_comm_group V] [add_torsor V P] : metric_space P := { dist := λ x y, ∥(x -ᵥ y : V)∥, dist_self := λ x, by simp, diff --git a/src/analysis/normed/group/ball_sphere.lean b/src/analysis/normed/group/ball_sphere.lean index da815442fa753..6ff647c77a400 100644 --- a/src/analysis/normed/group/ball_sphere.lean +++ b/src/analysis/normed/group/ball_sphere.lean @@ -14,7 +14,7 @@ semi normed group. open metric set -variables {E : Type*} [semi_normed_group E] {r : ℝ} +variables {E : Type*} [seminormed_add_comm_group E] {r : ℝ} /-- We equip the sphere, in a seminormed group, with a formal operation of negation, namely the antipodal map. -/ diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index b597da76db301..50a3481fe0ff7 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -17,10 +17,10 @@ In this file we define four classes: * `has_norm`, `has_nnnorm`: auxiliary classes endowing a type `α` with a function `norm : α → ℝ` (notation: `∥x∥`) and `nnnorm : α → ℝ≥0` (notation: `∥x∥₊`), respectively; -* `semi_normed_group`: a seminormed group is an additive group with a norm and a pseudo metric space - structures that agree with each other: `∀ x y, dist x y = ∥x - y∥`; -* `normed_group`: a normed group is an additive group with a norm and a metric space structures that - agree with each other: `∀ x y, dist x y = ∥x - y∥`. +* `seminormed_add_comm_group`: a seminormed group is an additive group with a norm and a + pseudo-metric space structures that agree with each other: `∀ x y, dist x y = ∥x - y∥`; +* `normed_add_comm_group`: a normed group is an additive group with a norm and a metric space + structures that agree with each other: `∀ x y, dist x y = ∥x - y∥`. We also prove basic properties of (semi)normed groups and provide some instances. @@ -45,23 +45,25 @@ notation `∥` e `∥` := norm e /-- A seminormed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a pseudometric space structure. -/ -class semi_normed_group (E : Type*) extends has_norm E, add_comm_group E, pseudo_metric_space E := +class seminormed_add_comm_group (E : Type*) + extends has_norm E, add_comm_group E, pseudo_metric_space E := (dist_eq : ∀ x y : E, dist x y = norm (x - y)) /-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a metric space structure. -/ -class normed_group (E : Type*) extends has_norm E, add_comm_group E, metric_space E := +class normed_add_comm_group (E : Type*) extends has_norm E, add_comm_group E, metric_space E := (dist_eq : ∀ x y : E, dist x y = norm (x - y)) /-- A normed group is a seminormed group. -/ @[priority 100] -- see Note [lower instance priority] -instance normed_group.to_semi_normed_group [h : normed_group E] : semi_normed_group E := +instance normed_add_comm_group.to_seminormed_add_comm_group [h : normed_add_comm_group E] : + seminormed_add_comm_group E := { .. h } /-- Construct a seminormed group from a translation invariant pseudodistance. -/ -def semi_normed_group.of_add_dist [has_norm E] [add_comm_group E] [pseudo_metric_space E] +def seminormed_add_comm_group.of_add_dist [has_norm E] [add_comm_group E] [pseudo_metric_space E] (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : semi_normed_group E := + (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : seminormed_add_comm_group E := { dist_eq := λ x y, begin rw H1, apply le_antisymm, { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 }, @@ -69,9 +71,9 @@ def semi_normed_group.of_add_dist [has_norm E] [add_comm_group E] [pseudo_metric end } /-- Construct a seminormed group from a translation invariant pseudodistance -/ -def semi_normed_group.of_add_dist' [has_norm E] [add_comm_group E] [pseudo_metric_space E] +def seminormed_add_comm_group.of_add_dist' [has_norm E] [add_comm_group E] [pseudo_metric_space E] (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist (x + z) (y + z) ≤ dist x y) : semi_normed_group E := + (H2 : ∀ x y z : E, dist (x + z) (y + z) ≤ dist x y) : seminormed_add_comm_group E := { dist_eq := λ x y, begin rw H1, apply le_antisymm, { have := H2 (x - y) 0 y, rwa [sub_add_cancel, zero_add] at this }, @@ -80,7 +82,7 @@ def semi_normed_group.of_add_dist' [has_norm E] [add_comm_group E] [pseudo_metri /-- A seminormed group can be built from a seminorm that satisfies algebraic properties. This is formalised in this structure. -/ -structure semi_normed_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := +structure seminormed_add_comm_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := (norm_zero : ∥(0 : E)∥ = 0) (triangle : ∀ x y : E, ∥x + y∥ ≤ ∥x∥ + ∥y∥) (norm_neg : ∀ x : E, ∥-x∥ = ∥x∥) @@ -89,8 +91,8 @@ structure semi_normed_group.core (E : Type*) [add_comm_group E] [has_norm E] : P pseudodistance and the pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on `E`). -/ -def semi_normed_group.of_core (E : Type*) [add_comm_group E] [has_norm E] - (C : semi_normed_group.core E) : semi_normed_group E := +def seminormed_add_comm_group.of_core (E : Type*) [add_comm_group E] [has_norm E] + (C : seminormed_add_comm_group.core E) : seminormed_add_comm_group E := { dist := λ x y, ∥x - y∥, dist_eq := assume x y, by refl, dist_self := assume x, by simp [C.norm_zero], @@ -101,23 +103,23 @@ def semi_normed_group.of_core (E : Type*) [add_comm_group E] [has_norm E] calc ∥x - y∥ = ∥ -(y - x)∥ : by simp ... = ∥y - x∥ : by { rw [C.norm_neg] } } -instance : normed_group punit := +instance : normed_add_comm_group punit := { norm := function.const _ 0, dist_eq := λ _ _, rfl, } @[simp] lemma punit.norm_eq_zero (r : punit) : ∥r∥ = 0 := rfl -noncomputable instance : normed_group ℝ := +noncomputable instance : normed_add_comm_group ℝ := { norm := λ x, |x|, dist_eq := assume x y, rfl } @[simp] lemma real.norm_eq_abs (r : ℝ) : ∥r∥ = |r| := rfl -section semi_normed_group -variables [semi_normed_group E] [semi_normed_group F] [semi_normed_group G] +section seminormed_add_comm_group +variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] lemma dist_eq_norm (g h : E) : dist g h = ∥g - h∥ := -semi_normed_group.dist_eq _ _ +seminormed_add_comm_group.dist_eq _ _ lemma dist_eq_norm' (g h : E) : dist g h = ∥h - g∥ := by rw [dist_comm, dist_eq_norm] @@ -423,33 +425,33 @@ variable {E} end isometric -theorem normed_group.tendsto_nhds_zero {f : α → E} {l : filter α} : +theorem normed_add_comm_group.tendsto_nhds_zero {f : α → E} {l : filter α} : tendsto f l (𝓝 0) ↔ ∀ ε > 0, ∀ᶠ x in l, ∥ f x ∥ < ε := metric.tendsto_nhds.trans $ by simp only [dist_zero_right] -lemma normed_group.tendsto_nhds_nhds {f : E → F} {x : E} {y : F} : +lemma normed_add_comm_group.tendsto_nhds_nhds {f : E → F} {x : E} {y : F} : tendsto f (𝓝 x) (𝓝 y) ↔ ∀ ε > 0, ∃ δ > 0, ∀ x', ∥x' - x∥ < δ → ∥f x' - y∥ < ε := by simp_rw [metric.tendsto_nhds_nhds, dist_eq_norm] -lemma normed_group.cauchy_seq_iff [nonempty α] [semilattice_sup α] {u : α → E} : +lemma normed_add_comm_group.cauchy_seq_iff [nonempty α] [semilattice_sup α] {u : α → E} : cauchy_seq u ↔ ∀ ε > 0, ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → ∥u m - u n∥ < ε := by simp [metric.cauchy_seq_iff, dist_eq_norm] -lemma normed_group.nhds_basis_norm_lt (x : E) : +lemma normed_add_comm_group.nhds_basis_norm_lt (x : E) : (𝓝 x).has_basis (λ (ε : ℝ), 0 < ε) (λ (ε : ℝ), { y | ∥y - x∥ < ε }) := begin simp_rw ← ball_eq, exact metric.nhds_basis_ball, end -lemma normed_group.nhds_zero_basis_norm_lt : +lemma normed_add_comm_group.nhds_zero_basis_norm_lt : (𝓝 (0 : E)).has_basis (λ (ε : ℝ), 0 < ε) (λ (ε : ℝ), { y | ∥y∥ < ε }) := begin - convert normed_group.nhds_basis_norm_lt (0 : E), + convert normed_add_comm_group.nhds_basis_norm_lt (0 : E), simp, end -lemma normed_group.uniformity_basis_dist : +lemma normed_add_comm_group.uniformity_basis_dist : (𝓤 E).has_basis (λ (ε : ℝ), 0 < ε) (λ ε, {p : E × E | ∥p.fst - p.snd∥ < ε}) := begin convert metric.uniformity_basis_dist, @@ -581,7 +583,7 @@ export has_nnnorm (nnnorm) notation `∥`e`∥₊` := nnnorm e @[priority 100] -- see Note [lower instance priority] -instance semi_normed_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨norm a, norm_nonneg a⟩⟩ +instance seminormed_add_comm_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨norm a, norm_nonneg a⟩⟩ @[simp, norm_cast] lemma coe_nnnorm (a : E) : (∥a∥₊ : ℝ) = norm a := rfl @@ -724,24 +726,26 @@ by simp [← dist_eq_norm, hf.le_mul_dist x y] end antilipschitz_with -/-- A group homomorphism from an `add_comm_group` to a `semi_normed_group` induces a -`semi_normed_group` structure on the domain. +/-- A group homomorphism from an `add_comm_group` to a `seminormed_add_comm_group` induces a +`seminormed_add_comm_group` structure on the domain. See note [reducible non-instances] -/ @[reducible] -def semi_normed_group.induced {E} [add_comm_group E] (f : E →+ F) : semi_normed_group E := +def seminormed_add_comm_group.induced {E} [add_comm_group E] (f : E →+ F) : + seminormed_add_comm_group E := { norm := λ x, ∥f x∥, dist_eq := λ x y, by simpa only [add_monoid_hom.map_sub, ← dist_eq_norm], - .. pseudo_metric_space.induced f semi_normed_group.to_pseudo_metric_space, } + .. pseudo_metric_space.induced f seminormed_add_comm_group.to_pseudo_metric_space, } /-- A subgroup of a seminormed group is also a seminormed group, with the restriction of the norm. -/ -instance add_subgroup.semi_normed_group (s : add_subgroup E) : semi_normed_group s := -semi_normed_group.induced s.subtype +instance add_subgroup.seminormed_add_comm_group (s : add_subgroup E) : + seminormed_add_comm_group s := +seminormed_add_comm_group.induced s.subtype /-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to its norm in `E`. -/ -@[simp] lemma add_subgroup.coe_norm {E : Type*} [semi_normed_group E] +@[simp] lemma add_subgroup.coe_norm {E : Type*} [seminormed_add_comm_group E] {s : add_subgroup E} (x : s) : ∥(x : s)∥ = ∥(x:E)∥ := rfl @@ -752,16 +756,17 @@ its norm in `E`. This is a reversed version of the `simp` lemma `add_subgroup.coe_norm` for use by `norm_cast`. -/ -@[norm_cast] lemma add_subgroup.norm_coe {E : Type*} [semi_normed_group E] {s : add_subgroup E} - (x : s) : +@[norm_cast] lemma add_subgroup.norm_coe {E : Type*} [seminormed_add_comm_group E] + {s : add_subgroup E} (x : s) : ∥(x : E)∥ = ∥(x : s)∥ := rfl /-- A submodule of a seminormed group is also a seminormed group, with the restriction of the norm. See note [implicit instance arguments]. -/ -instance submodule.semi_normed_group {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [semi_normed_group E] {_ : module 𝕜 E} (s : submodule 𝕜 E) : semi_normed_group s := +instance submodule.seminormed_add_comm_group {𝕜 E : Type*} {_ : ring 𝕜} + [seminormed_add_comm_group E] {_ : module 𝕜 E} (s : submodule 𝕜 E) : + seminormed_add_comm_group s := { norm := λx, norm (x : E), dist_eq := λx y, dist_eq_norm (x : E) (y : E) } @@ -770,7 +775,7 @@ norm in `E`. See note [implicit instance arguments]. -/ @[simp] lemma submodule.coe_norm {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [semi_normed_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : + {E : Type*} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : ∥(x : s)∥ = ∥(x : E)∥ := rfl @@ -781,12 +786,12 @@ This is a reversed version of the `simp` lemma `submodule.coe_norm` for use by ` See note [implicit instance arguments]. -/ @[norm_cast] lemma submodule.norm_coe {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [semi_normed_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : + {E : Type*} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : ∥(x : E)∥ = ∥(x : s)∥ := rfl -instance ulift.semi_normed_group : semi_normed_group (ulift E) := -semi_normed_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ +instance ulift.seminormed_add_comm_group : seminormed_add_comm_group (ulift E) := +seminormed_add_comm_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ lemma ulift.norm_def (x : ulift E) : ∥x∥ = ∥x.down∥ := rfl lemma ulift.nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl @@ -795,7 +800,7 @@ lemma ulift.nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl @[simp] lemma ulift.nnnorm_up (x : E) : ∥ulift.up x∥₊ = ∥x∥₊ := rfl /-- seminormed group instance on the product of two seminormed groups, using the sup norm. -/ -noncomputable instance prod.semi_normed_group : semi_normed_group (E × F) := +noncomputable instance prod.seminormed_add_comm_group : seminormed_add_comm_group (E × F) := { norm := λx, max ∥x.1∥ ∥x.2∥, dist_eq := assume (x y : E × F), show max (dist x.1 y.1) (dist x.2 y.2) = (max ∥(x - y).1∥ ∥(x - y).2∥), by simp [dist_eq_norm] } @@ -815,68 +820,55 @@ lemma norm_prod_le_iff {x : E × F} {r : ℝ} : ∥x∥ ≤ r ↔ ∥x.1∥ ≤ r ∧ ∥x.2∥ ≤ r := max_le_iff +section pi +variables {π : ι → Type*} [fintype ι] [Π i, seminormed_add_comm_group (π i)] (f : Π i, π i) + /-- seminormed group instance on the product of finitely many seminormed groups, using the sup norm. -/ -noncomputable instance pi.semi_normed_group {π : ι → Type*} [fintype ι] - [Π i, semi_normed_group (π i)] : semi_normed_group (Π i, π i) := +noncomputable instance pi.seminormed_add_comm_group : seminormed_add_comm_group (Π i, π i) := { norm := λ f, ↑(finset.univ.sup (λ b, ∥f b∥₊)), dist_eq := assume x y, congr_arg (coe : ℝ≥0 → ℝ) $ congr_arg (finset.sup finset.univ) $ funext $ assume a, show nndist (x a) (y a) = ∥x a - y a∥₊, from nndist_eq_nnnorm _ _ } -lemma pi.norm_def {π : ι → Type*} [fintype ι] [Π i, semi_normed_group (π i)] (f : Π i, π i) : - ∥f∥ = ↑(finset.univ.sup (λ b, ∥f b∥₊)) := rfl - -lemma pi.nnnorm_def {π : ι → Type*} [fintype ι] [Π i, semi_normed_group (π i)] (f : Π i, π i) : - ∥f∥₊ = finset.univ.sup (λ b, ∥f b∥₊) := subtype.eta _ _ +lemma pi.norm_def : ∥f∥ = ↑(finset.univ.sup (λ b, ∥f b∥₊)) := rfl +lemma pi.nnnorm_def : ∥f∥₊ = finset.univ.sup (λ b, ∥f b∥₊) := subtype.eta _ _ /-- The seminorm of an element in a product space is `≤ r` if and only if the norm of each component is. -/ -lemma pi_norm_le_iff {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] {r : ℝ} - (hr : 0 ≤ r) {x : Πi, π i} : ∥x∥ ≤ r ↔ ∀i, ∥x i∥ ≤ r := +lemma pi_norm_le_iff {r : ℝ} (hr : 0 ≤ r) {x : Π i, π i} : ∥x∥ ≤ r ↔ ∀ i, ∥x i∥ ≤ r := by simp only [← dist_zero_right, dist_pi_le_iff hr, pi.zero_apply] -lemma pi_nnnorm_le_iff {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] {r : ℝ≥0} - {x : Πi, π i} : ∥x∥₊ ≤ r ↔ ∀i, ∥x i∥₊ ≤ r := +lemma pi_nnnorm_le_iff {r : ℝ≥0} {x : Π i, π i} : ∥x∥₊ ≤ r ↔ ∀i, ∥x i∥₊ ≤ r := pi_norm_le_iff r.coe_nonneg /-- The seminorm of an element in a product space is `< r` if and only if the norm of each component is. -/ -lemma pi_norm_lt_iff {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] {r : ℝ} - (hr : 0 < r) {x : Πi, π i} : ∥x∥ < r ↔ ∀i, ∥x i∥ < r := +lemma pi_norm_lt_iff {r : ℝ} (hr : 0 < r) {x : Π i, π i} : ∥x∥ < r ↔ ∀ i, ∥x i∥ < r := by simp only [← dist_zero_right, dist_pi_lt_iff hr, pi.zero_apply] -lemma pi_nnnorm_lt_iff {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] {r : ℝ≥0} - (hr : 0 < r) {x : Πi, π i} : ∥x∥₊ < r ↔ ∀i, ∥x i∥₊ < r := +lemma pi_nnnorm_lt_iff {r : ℝ≥0} (hr : 0 < r) {x : Π i, π i} : ∥x∥₊ < r ↔ ∀ i, ∥x i∥₊ < r := pi_norm_lt_iff hr -lemma norm_le_pi_norm {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] (x : Πi, π i) - (i : ι) : ∥x i∥ ≤ ∥x∥ := -(pi_norm_le_iff (norm_nonneg x)).1 le_rfl i +lemma norm_le_pi_norm (i : ι) : ∥f i∥ ≤ ∥f∥ := (pi_norm_le_iff $ norm_nonneg _).1 le_rfl i -lemma nnnorm_le_pi_nnnorm {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] (x : Πi, π i) - (i : ι) : ∥x i∥₊ ≤ ∥x∥₊ := -norm_le_pi_norm x i +lemma nnnorm_le_pi_nnnorm (i : ι) : ∥f i∥₊ ≤ ∥f∥₊ := norm_le_pi_norm _ i -@[simp] lemma pi_norm_const [nonempty ι] [fintype ι] (a : E) : ∥(λ i : ι, a)∥ = ∥a∥ := +@[simp] lemma pi_norm_const [nonempty ι] (a : E) : ∥(λ i : ι, a)∥ = ∥a∥ := by simpa only [← dist_zero_right] using dist_pi_const a 0 -@[simp] lemma pi_nnnorm_const [nonempty ι] [fintype ι] (a : E) : - ∥(λ i : ι, a)∥₊ = ∥a∥₊ := +@[simp] lemma pi_nnnorm_const [nonempty ι] (a : E) : ∥(λ i : ι, a)∥₊ = ∥a∥₊ := nnreal.eq $ pi_norm_const a /-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ -lemma pi.sum_norm_apply_le_norm {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] - (x : Π i, π i) : - ∑ i, ∥x i∥ ≤ fintype.card ι • ∥x∥ := -calc ∑ i, ∥x i∥ ≤ ∑ i : ι, ∥x∥ : finset.sum_le_sum $ λ i hi, norm_le_pi_norm x i - ... = fintype.card ι • ∥x∥ : finset.sum_const _ +lemma pi.sum_norm_apply_le_norm : ∑ i, ∥f i∥ ≤ fintype.card ι • ∥f∥ := +finset.sum_le_card_nsmul _ _ _ $ λ i hi, norm_le_pi_norm _ i /-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ -lemma pi.sum_nnnorm_apply_le_nnnorm {π : ι → Type*} [fintype ι] [∀i, semi_normed_group (π i)] - (x : Π i, π i) : - ∑ i, ∥x i∥₊ ≤ fintype.card ι • ∥x∥₊ := -nnreal.coe_sum.trans_le $ pi.sum_norm_apply_le_norm x +lemma pi.sum_nnnorm_apply_le_nnnorm : ∑ i, ∥f i∥₊ ≤ fintype.card ι • ∥f∥₊ := +nnreal.coe_sum.trans_le $ pi.sum_norm_apply_le_norm _ + +end pi lemma tendsto_iff_norm_tendsto_zero {f : α → E} {a : filter α} {b : E} : tendsto f a (𝓝 b) ↔ tendsto (λ e, ∥f e - b∥) a (𝓝 0) := @@ -947,7 +939,7 @@ lemma filter.tendsto.op_zero_is_bounded_under_le' {f : α → E} {g : α → F} begin cases h_op with A h_op, rcases hg with ⟨C, hC⟩, rw eventually_map at hC, - rw normed_group.tendsto_nhds_zero at hf ⊢, + rw normed_add_comm_group.tendsto_nhds_zero at hf ⊢, intros ε ε₀, rcases exists_pos_mul_lt ε₀ (A * C) with ⟨δ, δ₀, hδ⟩, filter_upwards [hf δ δ₀, hC] with i hf hg, @@ -1020,7 +1012,7 @@ lemma eventually_ne_of_tendsto_norm_at_top {l : filter α} {f : α → E} (h.eventually_ne_at_top _).mono $ λ x, ne_of_apply_ne norm @[priority 100] -- see Note [lower instance priority] -instance semi_normed_group.has_lipschitz_add : has_lipschitz_add E := +instance seminormed_add_comm_group.has_lipschitz_add : has_lipschitz_add E := { lipschitz_add := ⟨2, lipschitz_with.prod_fst.add lipschitz_with.prod_snd⟩ } /-- A seminormed group is a uniform additive group, i.e., addition and subtraction are uniformly @@ -1033,15 +1025,15 @@ instance normed_uniform_group : uniform_add_group E := instance normed_top_group : topological_add_group E := by apply_instance -- short-circuit type class inference -lemma semi_normed_group.mem_closure_iff {s : set E} {x : E} : +lemma seminormed_add_comm_group.mem_closure_iff {s : set E} {x : E} : x ∈ closure s ↔ ∀ ε > 0, ∃ y ∈ s, ∥x - y∥ < ε := by simp [metric.mem_closure_iff, dist_eq_norm] lemma norm_le_zero_iff' [t0_space E] {g : E} : ∥g∥ ≤ 0 ↔ g = 0 := begin - letI : normed_group E := { to_metric_space := metric.of_t0_pseudo_metric_space E, - .. ‹semi_normed_group E› }, + letI : normed_add_comm_group E := { to_metric_space := metric.of_t0_pseudo_metric_space E, + .. ‹seminormed_add_comm_group E› }, rw [← dist_zero_right], exact dist_le_zero end @@ -1066,14 +1058,14 @@ begin exact (tendsto_at_top_of_eventually_const this).cauchy_seq.add hv end -end semi_normed_group +end seminormed_add_comm_group -section normed_group +section normed_add_comm_group /-- Construct a normed group from a translation invariant distance -/ -def normed_group.of_add_dist [has_norm E] [add_comm_group E] [metric_space E] +def normed_add_comm_group.of_add_dist [has_norm E] [add_comm_group E] [metric_space E] (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : normed_group E := + (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : normed_add_comm_group E := { dist_eq := λ x y, begin rw H1, apply le_antisymm, { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 }, @@ -1082,30 +1074,32 @@ def normed_group.of_add_dist [has_norm E] [add_comm_group E] [metric_space E] /-- A normed group can be built from a norm that satisfies algebraic properties. This is formalised in this structure. -/ -structure normed_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := +structure normed_add_comm_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := (norm_eq_zero_iff : ∀ x : E, ∥x∥ = 0 ↔ x = 0) (triangle : ∀ x y : E, ∥x + y∥ ≤ ∥x∥ + ∥y∥) (norm_neg : ∀ x : E, ∥-x∥ = ∥x∥) -/-- The `semi_normed_group.core` induced by a `normed_group.core`. -/ -lemma normed_group.core.to_semi_normed_group.core {E : Type*} [add_comm_group E] [has_norm E] - (C : normed_group.core E) : semi_normed_group.core E := +/-- The `seminormed_add_comm_group.core` induced by a `normed_add_comm_group.core`. -/ +lemma normed_add_comm_group.core.to_seminormed_add_comm_group.core {E : Type*} [add_comm_group E] + [has_norm E] + (C : normed_add_comm_group.core E) : seminormed_add_comm_group.core E := { norm_zero := (C.norm_eq_zero_iff 0).2 rfl, triangle := C.triangle, norm_neg := C.norm_neg } /-- Constructing a normed group from core properties of a norm, i.e., registering the distance and the metric space structure from the norm properties. -/ -def normed_group.of_core (E : Type*) [add_comm_group E] [has_norm E] - (C : normed_group.core E) : normed_group E := +def normed_add_comm_group.of_core (E : Type*) [add_comm_group E] [has_norm E] + (C : normed_add_comm_group.core E) : normed_add_comm_group E := { eq_of_dist_eq_zero := λ x y h, begin rw [dist_eq_norm] at h, exact sub_eq_zero.mp ((C.norm_eq_zero_iff _).1 h) end - ..semi_normed_group.of_core E (normed_group.core.to_semi_normed_group.core C) } + ..seminormed_add_comm_group.of_core E + (normed_add_comm_group.core.to_seminormed_add_comm_group.core C) } -variables [normed_group E] [normed_group F] {x y : E} +variables [normed_add_comm_group E] [normed_add_comm_group F] {x y : E} @[simp] lemma norm_eq_zero {g : E} : ∥g∥ = 0 ↔ g = 0 := norm_eq_zero_iff' @@ -1132,35 +1126,38 @@ by rw [← nnreal.coe_eq_zero, coe_nnnorm, norm_eq_zero] lemma nnnorm_ne_zero_iff {g : E} : ∥g∥₊ ≠ 0 ↔ g ≠ 0 := not_congr nnnorm_eq_zero -/-- An injective group homomorphism from an `add_comm_group` to a `normed_group` induces a -`normed_group` structure on the domain. +/-- An injective group homomorphism from an `add_comm_group` to a `normed_add_comm_group` induces a +`normed_add_comm_group` structure on the domain. See note [reducible non-instances]. -/ @[reducible] -def normed_group.induced {E} [add_comm_group E] - (f : E →+ F) (h : function.injective f) : normed_group E := -{ .. semi_normed_group.induced f, - .. metric_space.induced f h normed_group.to_metric_space, } +def normed_add_comm_group.induced {E} [add_comm_group E] + (f : E →+ F) (h : function.injective f) : normed_add_comm_group E := +{ .. seminormed_add_comm_group.induced f, + .. metric_space.induced f h normed_add_comm_group.to_metric_space, } /-- A subgroup of a normed group is also a normed group, with the restriction of the norm. -/ -instance add_subgroup.normed_group (s : add_subgroup E) : normed_group s := -normed_group.induced s.subtype subtype.coe_injective +instance add_subgroup.normed_add_comm_group (s : add_subgroup E) : normed_add_comm_group s := +normed_add_comm_group.induced s.subtype subtype.coe_injective /-- A submodule of a normed group is also a normed group, with the restriction of the norm. See note [implicit instance arguments]. -/ -instance submodule.normed_group {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [normed_group E] {_ : module 𝕜 E} (s : submodule 𝕜 E) : normed_group s := -{ ..submodule.semi_normed_group s } +instance submodule.normed_add_comm_group {𝕜 E : Type*} {_ : ring 𝕜} [normed_add_comm_group E] + {_ : module 𝕜 E} (s : submodule 𝕜 E) : normed_add_comm_group s := +{ ..submodule.seminormed_add_comm_group s } -instance ulift.normed_group : normed_group (ulift E) := { ..ulift.semi_normed_group } +instance ulift.normed_add_comm_group : normed_add_comm_group (ulift E) := +{ ..ulift.seminormed_add_comm_group } /-- normed group instance on the product of two normed groups, using the sup norm. -/ -noncomputable instance prod.normed_group : normed_group (E × F) := { ..prod.semi_normed_group } +noncomputable instance prod.normed_add_comm_group : normed_add_comm_group (E × F) := +{ ..prod.seminormed_add_comm_group } /-- normed group instance on the product of finitely many normed groups, using the sup norm. -/ -noncomputable instance pi.normed_group {π : ι → Type*} [fintype ι] [∀i, normed_group (π i)] : - normed_group (Πi, π i) := { ..pi.semi_normed_group } +noncomputable instance pi.normed_add_comm_group {π : ι → Type*} [fintype ι] + [Π i, normed_add_comm_group (π i)] : + normed_add_comm_group (Πi, π i) := { ..pi.seminormed_add_comm_group } lemma tendsto_norm_sub_self_punctured_nhds (a : E) : tendsto (λ x, ∥x - a∥) (𝓝[≠] a) (𝓝[>] 0) := (tendsto_norm_sub_self a).inf $ tendsto_principal_principal.2 $ λ x hx, @@ -1182,4 +1179,4 @@ lemma continuous.bounded_above_of_compact_support [topological_space α] {f : α by simpa [bdd_above_def] using hf.norm.bdd_above_range_of_has_compact_support hsupp.norm -end normed_group +end normed_add_comm_group diff --git a/src/analysis/normed/group/completion.lean b/src/analysis/normed/group/completion.lean index a5feedcfe3f58..de99d3ecfcde7 100644 --- a/src/analysis/normed/group/completion.lean +++ b/src/analysis/normed/group/completion.lean @@ -28,11 +28,11 @@ instance [uniform_space E] [has_norm E] : has_norm (completion E) := { norm := completion.extension has_norm.norm } -@[simp] lemma norm_coe {E} [semi_normed_group E] (x : E) : +@[simp] lemma norm_coe {E} [seminormed_add_comm_group E] (x : E) : ∥(x : completion E)∥ = ∥x∥ := completion.extension_coe uniform_continuous_norm x -instance [semi_normed_group E] : normed_group (completion E) := +instance [seminormed_add_comm_group E] : normed_add_comm_group (completion E) := { dist_eq := begin intros x y, diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index bc4e4aaf7d305..89bb35bccb42a 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -21,59 +21,63 @@ simple constructions for normed group homs, like kernel, range and equalizer. Some easy other constructions are related to subgroups of normed groups. Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the -theory of `semi_normed_group_hom` and we specialize to `normed_group_hom` when needed. +theory of `seminormed_add_group_hom` and we specialize to `normed_add_group_hom` when needed. -/ noncomputable theory open_locale nnreal big_operators /-- A morphism of seminormed abelian groups is a bounded group homomorphism. -/ -structure normed_group_hom (V W : Type*) [semi_normed_group V] [semi_normed_group W] := +structure normed_add_group_hom (V W : Type*) [seminormed_add_comm_group V] + [seminormed_add_comm_group W] := (to_fun : V → W) (map_add' : ∀ v₁ v₂, to_fun (v₁ + v₂) = to_fun v₁ + to_fun v₂) (bound' : ∃ C, ∀ v, ∥to_fun v∥ ≤ C * ∥v∥) namespace add_monoid_hom -variables {V W : Type*} [semi_normed_group V] [semi_normed_group W] {f g : normed_group_hom V W} +variables {V W : Type*} [seminormed_add_comm_group V] [seminormed_add_comm_group W] + {f g : normed_add_group_hom V W} /-- Associate to a group homomorphism a bounded group homomorphism under a norm control condition. -See `add_monoid_hom.mk_normed_group_hom'` for a version that uses `ℝ≥0` for the bound. -/ -def mk_normed_group_hom (f : V →+ W) - (C : ℝ) (h : ∀ v, ∥f v∥ ≤ C * ∥v∥) : normed_group_hom V W := +See `add_monoid_hom.mk_normed_add_group_hom'` for a version that uses `ℝ≥0` for the bound. -/ +def mk_normed_add_group_hom (f : V →+ W) + (C : ℝ) (h : ∀ v, ∥f v∥ ≤ C * ∥v∥) : normed_add_group_hom V W := { bound' := ⟨C, h⟩, ..f } /-- Associate to a group homomorphism a bounded group homomorphism under a norm control condition. -See `add_monoid_hom.mk_normed_group_hom` for a version that uses `ℝ` for the bound. -/ -def mk_normed_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : - normed_group_hom V W := +See `add_monoid_hom.mk_normed_add_group_hom` for a version that uses `ℝ` for the bound. -/ +def mk_normed_add_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : + normed_add_group_hom V W := { bound' := ⟨C, hC⟩ .. f} end add_monoid_hom -lemma exists_pos_bound_of_bound {V W : Type*} [semi_normed_group V] [semi_normed_group W] +lemma exists_pos_bound_of_bound {V W : Type*} [seminormed_add_comm_group V] + [seminormed_add_comm_group W] {f : V → W} (M : ℝ) (h : ∀x, ∥f x∥ ≤ M * ∥x∥) : ∃ N, 0 < N ∧ ∀x, ∥f x∥ ≤ N * ∥x∥ := ⟨max M 1, lt_of_lt_of_le zero_lt_one (le_max_right _ _), λx, calc ∥f x∥ ≤ M * ∥x∥ : h x ... ≤ max M 1 * ∥x∥ : mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg _) ⟩ -namespace normed_group_hom +namespace normed_add_group_hom -variables {V V₁ V₂ V₃ : Type*} -variables [semi_normed_group V] [semi_normed_group V₁] [semi_normed_group V₂] [semi_normed_group V₃] -variables {f g : normed_group_hom V₁ V₂} +variables {V V₁ V₂ V₃ : Type*} [seminormed_add_comm_group V] [seminormed_add_comm_group V₁] + [seminormed_add_comm_group V₂] [seminormed_add_comm_group V₃] +variables {f g : normed_add_group_hom V₁ V₂} -instance : has_coe_to_fun (normed_group_hom V₁ V₂) (λ _, V₁ → V₂) := ⟨normed_group_hom.to_fun⟩ +instance : has_coe_to_fun (normed_add_group_hom V₁ V₂) (λ _, V₁ → V₂) := +⟨normed_add_group_hom.to_fun⟩ -initialize_simps_projections normed_group_hom (to_fun → apply) +initialize_simps_projections normed_add_group_hom (to_fun → apply) lemma coe_inj (H : (f : V₁ → V₂) = g) : f = g := by cases f; cases g; congr'; exact funext H -lemma coe_injective : @function.injective (normed_group_hom V₁ V₂) (V₁ → V₂) coe_fn := +lemma coe_injective : @function.injective (normed_add_group_hom V₁ V₂) (V₁ → V₂) coe_fn := by apply coe_inj lemma coe_inj_iff : f = g ↔ (f : V₁ → V₂) = g := ⟨congr_arg _, coe_inj⟩ @@ -86,28 +90,28 @@ variables (f g) @[simp] lemma to_fun_eq_coe : f.to_fun = f := rfl -@[simp] lemma coe_mk (f) (h₁) (h₂) (h₃) : ⇑(⟨f, h₁, h₂, h₃⟩ : normed_group_hom V₁ V₂) = f := rfl +@[simp] lemma coe_mk (f) (h₁) (h₂) (h₃) : ⇑(⟨f, h₁, h₂, h₃⟩ : normed_add_group_hom V₁ V₂) = f := rfl -@[simp] lemma coe_mk_normed_group_hom (f : V₁ →+ V₂) (C) (hC) : - ⇑(f.mk_normed_group_hom C hC) = f := rfl +@[simp] lemma coe_mk_normed_add_group_hom (f : V₁ →+ V₂) (C) (hC) : + ⇑(f.mk_normed_add_group_hom C hC) = f := rfl -@[simp] lemma coe_mk_normed_group_hom' (f : V₁ →+ V₂) (C) (hC) : - ⇑(f.mk_normed_group_hom' C hC) = f := rfl +@[simp] lemma coe_mk_normed_add_group_hom' (f : V₁ →+ V₂) (C) (hC) : + ⇑(f.mk_normed_add_group_hom' C hC) = f := rfl /-- The group homomorphism underlying a bounded group homomorphism. -/ -def to_add_monoid_hom (f : normed_group_hom V₁ V₂) : V₁ →+ V₂ := +def to_add_monoid_hom (f : normed_add_group_hom V₁ V₂) : V₁ →+ V₂ := add_monoid_hom.mk' f f.map_add' @[simp] lemma coe_to_add_monoid_hom : ⇑f.to_add_monoid_hom = f := rfl lemma to_add_monoid_hom_injective : - function.injective (@normed_group_hom.to_add_monoid_hom V₁ V₂ _ _) := + function.injective (@normed_add_group_hom.to_add_monoid_hom V₁ V₂ _ _) := λ f g h, coe_inj $ show ⇑f.to_add_monoid_hom = g, by { rw h, refl } @[simp] lemma mk_to_add_monoid_hom (f) (h₁) (h₂) : - (⟨f, h₁, h₂⟩ : normed_group_hom V₁ V₂).to_add_monoid_hom = add_monoid_hom.mk' f h₁ := rfl + (⟨f, h₁, h₂⟩ : normed_add_group_hom V₁ V₂).to_add_monoid_hom = add_monoid_hom.mk' f h₁ := rfl -instance : add_monoid_hom_class (normed_group_hom V₁ V₂) V₁ V₂ := +instance : add_monoid_hom_class (normed_add_group_hom V₁ V₂) V₁ V₂ := { coe := coe_fn, coe_injective' := coe_injective, map_add := λ f, f.to_add_monoid_hom.map_add, @@ -125,10 +129,10 @@ antilipschitz_with.of_le_mul_dist $ `x` of `K` has a preimage whose norm is bounded above by `C*∥x∥`. This is a more abstract version of `f` having a right inverse defined on `K` with operator norm at most `C`. -/ -def surjective_on_with (f : normed_group_hom V₁ V₂) (K : add_subgroup V₂) (C : ℝ) : Prop := +def surjective_on_with (f : normed_add_group_hom V₁ V₂) (K : add_subgroup V₂) (C : ℝ) : Prop := ∀ h ∈ K, ∃ g, f g = h ∧ ∥g∥ ≤ C*∥h∥ -lemma surjective_on_with.mono {f : normed_group_hom V₁ V₂} {K : add_subgroup V₂} {C C' : ℝ} +lemma surjective_on_with.mono {f : normed_add_group_hom V₁ V₂} {K : add_subgroup V₂} {C C' : ℝ} (h : f.surjective_on_with K C) (H : C ≤ C') : f.surjective_on_with K C' := begin intros k k_in, @@ -139,7 +143,7 @@ begin { exact hg.trans ((mul_le_mul_right $ (ne.symm Hg).le_iff_lt.mp (norm_nonneg _)).mpr H) } end -lemma surjective_on_with.exists_pos {f : normed_group_hom V₁ V₂} {K : add_subgroup V₂} {C : ℝ} +lemma surjective_on_with.exists_pos {f : normed_add_group_hom V₁ V₂} {K : add_subgroup V₂} {C : ℝ} (h : f.surjective_on_with K C) : ∃ C' > 0, f.surjective_on_with K C' := begin refine ⟨|C| + 1, _, _⟩, @@ -148,25 +152,25 @@ begin linarith [le_abs_self C] } end -lemma surjective_on_with.surj_on {f : normed_group_hom V₁ V₂} {K : add_subgroup V₂} {C : ℝ} +lemma surjective_on_with.surj_on {f : normed_add_group_hom V₁ V₂} {K : add_subgroup V₂} {C : ℝ} (h : f.surjective_on_with K C) : set.surj_on f set.univ K := λ x hx, (h x hx).imp $ λ a ⟨ha, _⟩, ⟨set.mem_univ _, ha⟩ /-! ### The operator norm -/ /-- The operator norm of a seminormed group homomorphism is the inf of all its bounds. -/ -def op_norm (f : normed_group_hom V₁ V₂) := Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} -instance has_op_norm : has_norm (normed_group_hom V₁ V₂) := ⟨op_norm⟩ +def op_norm (f : normed_add_group_hom V₁ V₂) := Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} +instance has_op_norm : has_norm (normed_add_group_hom V₁ V₂) := ⟨op_norm⟩ lemma norm_def : ∥f∥ = Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} := rfl -- So that invocations of `le_cInf` make sense: we show that the set of -- bounds is nonempty and bounded below. -lemma bounds_nonempty {f : normed_group_hom V₁ V₂} : +lemma bounds_nonempty {f : normed_add_group_hom V₁ V₂} : ∃ c, c ∈ { c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥ } := let ⟨M, hMp, hMb⟩ := f.bound in ⟨M, le_of_lt hMp, hMb⟩ -lemma bounds_bdd_below {f : normed_group_hom V₁ V₂} : +lemma bounds_bdd_below {f : normed_add_group_hom V₁ V₂} : bdd_below {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} := ⟨0, λ _ ⟨hn, _⟩, hn⟩ @@ -196,11 +200,11 @@ theorem lipschitz : lipschitz_with ⟨∥f∥, op_norm_nonneg f⟩ f := lipschitz_with.of_dist_le_mul $ λ x y, by { rw [dist_eq_norm, dist_eq_norm, ←map_sub], apply le_op_norm } -protected lemma uniform_continuous (f : normed_group_hom V₁ V₂) : +protected lemma uniform_continuous (f : normed_add_group_hom V₁ V₂) : uniform_continuous f := f.lipschitz.uniform_continuous @[continuity] -protected lemma continuous (f : normed_group_hom V₁ V₂) : continuous f := +protected lemma continuous (f : normed_add_group_hom V₁ V₂) : continuous f := f.uniform_continuous.continuous lemma ratio_le_op_norm (x : V₁) : ∥f x∥ / ∥x∥ ≤ ∥f∥ := @@ -215,36 +219,38 @@ lemma op_norm_eq_of_bounds {M : ℝ} (M_nonneg : 0 ≤ M) (h_above : ∀ x, ∥f x∥ ≤ M*∥x∥) (h_below : ∀ N ≥ 0, (∀ x, ∥f x∥ ≤ N*∥x∥) → M ≤ N) : ∥f∥ = M := le_antisymm (f.op_norm_le_bound M_nonneg h_above) - ((le_cInf_iff normed_group_hom.bounds_bdd_below ⟨M, M_nonneg, h_above⟩).mpr $ + ((le_cInf_iff normed_add_group_hom.bounds_bdd_below ⟨M, M_nonneg, h_above⟩).mpr $ λ N ⟨N_nonneg, hN⟩, h_below N N_nonneg hN) -theorem op_norm_le_of_lipschitz {f : normed_group_hom V₁ V₂} {K : ℝ≥0} (hf : lipschitz_with K f) : +theorem op_norm_le_of_lipschitz {f : normed_add_group_hom V₁ V₂} {K : ℝ≥0} + (hf : lipschitz_with K f) : ∥f∥ ≤ K := f.op_norm_le_bound K.2 $ λ x, by simpa only [dist_zero_right, map_zero] using hf.dist_le_mul x 0 /-- If a bounded group homomorphism map is constructed from a group homomorphism via the constructor -`mk_normed_group_hom`, then its norm is bounded by the bound given to the constructor if it is +`mk_normed_add_group_hom`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ -lemma mk_normed_group_hom_norm_le (f : V₁ →+ V₂) {C : ℝ} (hC : 0 ≤ C) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_normed_group_hom C h∥ ≤ C := +lemma mk_normed_add_group_hom_norm_le (f : V₁ →+ V₂) {C : ℝ} (hC : 0 ≤ C) + (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : + ∥f.mk_normed_add_group_hom C h∥ ≤ C := op_norm_le_bound _ hC h /-- If a bounded group homomorphism map is constructed from a group homomorphism -via the constructor `mk_normed_group_hom`, then its norm is bounded by the bound +via the constructor `mk_normed_add_group_hom`, then its norm is bounded by the bound given to the constructor or zero if this bound is negative. -/ -lemma mk_normed_group_hom_norm_le' (f : V₁ →+ V₂) {C : ℝ} (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_normed_group_hom C h∥ ≤ max C 0 := +lemma mk_normed_add_group_hom_norm_le' (f : V₁ →+ V₂) {C : ℝ} (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : + ∥f.mk_normed_add_group_hom C h∥ ≤ max C 0 := op_norm_le_bound _ (le_max_right _ _) $ λ x, (h x).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg x) -alias mk_normed_group_hom_norm_le ← _root_.add_monoid_hom.mk_normed_group_hom_norm_le -alias mk_normed_group_hom_norm_le' ← _root_.add_monoid_hom.mk_normed_group_hom_norm_le' +alias mk_normed_add_group_hom_norm_le ← _root_.add_monoid_hom.mk_normed_add_group_hom_norm_le +alias mk_normed_add_group_hom_norm_le' ← _root_.add_monoid_hom.mk_normed_add_group_hom_norm_le' /-! ### Addition of normed group homs -/ /-- Addition of normed group homs. -/ -instance : has_add (normed_group_hom V₁ V₂) := -⟨λ f g, (f.to_add_monoid_hom + g.to_add_monoid_hom).mk_normed_group_hom (∥f∥ + ∥g∥) $ λ v, calc +instance : has_add (normed_add_group_hom V₁ V₂) := +⟨λ f g, (f.to_add_monoid_hom + g.to_add_monoid_hom).mk_normed_add_group_hom (∥f∥ + ∥g∥) $ λ v, calc ∥f v + g v∥ ≤ ∥f v∥ + ∥g v∥ : norm_add_le _ _ ... ≤ ∥f∥ * ∥v∥ + ∥g∥ * ∥v∥ : add_le_add (le_op_norm f v) (le_op_norm g v) @@ -252,7 +258,7 @@ instance : has_add (normed_group_hom V₁ V₂) := /-- The operator norm satisfies the triangle inequality. -/ theorem op_norm_add_le : ∥f + g∥ ≤ ∥f∥ + ∥g∥ := -mk_normed_group_hom_norm_le _ (add_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ +mk_normed_add_group_hom_norm_le _ (add_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ /-- Terms containing `@has_add.add (has_coe_to_fun.F ...) pi.has_add` @@ -263,26 +269,26 @@ As a workaround, we add a type annotation: `(f + g : V₁ → V₂)` library_note "addition on function coercions" -- see Note [addition on function coercions] -@[simp] lemma coe_add (f g : normed_group_hom V₁ V₂) : ⇑(f + g) = (f + g : V₁ → V₂) := rfl -@[simp] lemma add_apply (f g : normed_group_hom V₁ V₂) (v : V₁) : - (f + g : normed_group_hom V₁ V₂) v = f v + g v := rfl +@[simp] lemma coe_add (f g : normed_add_group_hom V₁ V₂) : ⇑(f + g) = (f + g : V₁ → V₂) := rfl +@[simp] lemma add_apply (f g : normed_add_group_hom V₁ V₂) (v : V₁) : + (f + g : normed_add_group_hom V₁ V₂) v = f v + g v := rfl /-! ### The zero normed group hom -/ -instance : has_zero (normed_group_hom V₁ V₂) := -⟨(0 : V₁ →+ V₂).mk_normed_group_hom 0 (by simp)⟩ +instance : has_zero (normed_add_group_hom V₁ V₂) := +⟨(0 : V₁ →+ V₂).mk_normed_add_group_hom 0 (by simp)⟩ -instance : inhabited (normed_group_hom V₁ V₂) := ⟨0⟩ +instance : inhabited (normed_add_group_hom V₁ V₂) := ⟨0⟩ /-- The norm of the `0` operator is `0`. -/ -theorem op_norm_zero : ∥(0 : normed_group_hom V₁ V₂)∥ = 0 := +theorem op_norm_zero : ∥(0 : normed_add_group_hom V₁ V₂)∥ = 0 := le_antisymm (cInf_le bounds_bdd_below ⟨ge_of_eq rfl, λ _, le_of_eq (by { rw [zero_mul], exact norm_zero })⟩) (op_norm_nonneg _) /-- For normed groups, an operator is zero iff its norm vanishes. -/ -theorem op_norm_zero_iff {V₁ V₂ : Type*} [normed_group V₁] [normed_group V₂] - {f : normed_group_hom V₁ V₂} : ∥f∥ = 0 ↔ f = 0 := +theorem op_norm_zero_iff {V₁ V₂ : Type*} [normed_add_comm_group V₁] [normed_add_comm_group V₂] + {f : normed_add_group_hom V₁ V₂} : ∥f∥ = 0 ↔ f = 0 := iff.intro (λ hn, ext (λ x, norm_le_zero_iff.1 (calc _ ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ @@ -290,8 +296,8 @@ iff.intro (λ hf, by rw [hf, op_norm_zero] ) -- see Note [addition on function coercions] -@[simp] lemma coe_zero : ⇑(0 : normed_group_hom V₁ V₂) = (0 : V₁ → V₂) := rfl -@[simp] lemma zero_apply (v : V₁) : (0 : normed_group_hom V₁ V₂) v = 0 := rfl +@[simp] lemma coe_zero : ⇑(0 : normed_add_group_hom V₁ V₂) = (0 : V₁ → V₂) := rfl +@[simp] lemma zero_apply (v : V₁) : (0 : normed_add_group_hom V₁ V₂) v = 0 := rfl variables {f g} @@ -301,13 +307,13 @@ variable (V) /-- The identity as a continuous normed group hom. -/ @[simps] -def id : normed_group_hom V V := -(add_monoid_hom.id V).mk_normed_group_hom 1 (by simp [le_refl]) +def id : normed_add_group_hom V V := +(add_monoid_hom.id V).mk_normed_add_group_hom 1 (by simp [le_refl]) /-- The norm of the identity is at most `1`. It is in fact `1`, except when the norm of every element vanishes, where it is `0`. (Since we are working with seminorms this can happen even if the space is non-trivial.) It means that one can not do better than an inequality in general. -/ -lemma norm_id_le : ∥(id V : normed_group_hom V V)∥ ≤ 1 := +lemma norm_id_le : ∥(id V : normed_add_group_hom V V)∥ ≤ 1 := op_norm_le_bound _ zero_le_one (λx, by simp) /-- If there is an element with norm different from `0`, then the norm of the identity equals `1`. @@ -319,33 +325,33 @@ have _ := (id V).ratio_le_op_norm x, by rwa [id_apply, div_self hx] at this /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ -lemma norm_id {V : Type*} [normed_group V] [nontrivial V] : ∥(id V)∥ = 1 := +lemma norm_id {V : Type*} [normed_add_comm_group V] [nontrivial V] : ∥(id V)∥ = 1 := begin refine norm_id_of_nontrivial_seminorm V _, obtain ⟨x, hx⟩ := exists_ne (0 : V), exact ⟨x, ne_of_gt (norm_pos_iff.2 hx)⟩, end -lemma coe_id : ((normed_group_hom.id V) : V → V) = (_root_.id : V → V) := rfl +lemma coe_id : ((normed_add_group_hom.id V) : V → V) = (_root_.id : V → V) := rfl /-! ### The negation of a normed group hom -/ /-- Opposite of a normed group hom. -/ -instance : has_neg (normed_group_hom V₁ V₂) := -⟨λ f, (-f.to_add_monoid_hom).mk_normed_group_hom (∥f∥) (λ v, by simp [le_op_norm f v])⟩ +instance : has_neg (normed_add_group_hom V₁ V₂) := +⟨λ f, (-f.to_add_monoid_hom).mk_normed_add_group_hom (∥f∥) (λ v, by simp [le_op_norm f v])⟩ -- see Note [addition on function coercions] -@[simp] lemma coe_neg (f : normed_group_hom V₁ V₂) : ⇑(-f) = (-f : V₁ → V₂) := rfl -@[simp] lemma neg_apply (f : normed_group_hom V₁ V₂) (v : V₁) : - (-f : normed_group_hom V₁ V₂) v = - (f v) := rfl +@[simp] lemma coe_neg (f : normed_add_group_hom V₁ V₂) : ⇑(-f) = (-f : V₁ → V₂) := rfl +@[simp] lemma neg_apply (f : normed_add_group_hom V₁ V₂) (v : V₁) : + (-f : normed_add_group_hom V₁ V₂) v = - (f v) := rfl -lemma op_norm_neg (f : normed_group_hom V₁ V₂) : ∥-f∥ = ∥f∥ := +lemma op_norm_neg (f : normed_add_group_hom V₁ V₂) : ∥-f∥ = ∥f∥ := by simp only [norm_def, coe_neg, norm_neg, pi.neg_apply] /-! ### Subtraction of normed group homs -/ /-- Subtraction of normed group homs. -/ -instance : has_sub (normed_group_hom V₁ V₂) := +instance : has_sub (normed_add_group_hom V₁ V₂) := ⟨λ f g, { bound' := begin @@ -355,9 +361,9 @@ instance : has_sub (normed_group_hom V₁ V₂) := .. (f.to_add_monoid_hom - g.to_add_monoid_hom) }⟩ -- see Note [addition on function coercions] -@[simp] lemma coe_sub (f g : normed_group_hom V₁ V₂) : ⇑(f - g) = (f - g : V₁ → V₂) := rfl -@[simp] lemma sub_apply (f g : normed_group_hom V₁ V₂) (v : V₁) : - (f - g : normed_group_hom V₁ V₂) v = f v - g v := rfl +@[simp] lemma coe_sub (f g : normed_add_group_hom V₁ V₂) : ⇑(f - g) = (f - g : V₁ → V₂) := rfl +@[simp] lemma sub_apply (f g : normed_add_group_hom V₁ V₂) (v : V₁) : + (f - g : normed_add_group_hom V₁ V₂) v = f v - g v := rfl /-! ### Scalar actions on normed group homs -/ @@ -367,7 +373,7 @@ variables {R R' : Type*} [monoid_with_zero R] [distrib_mul_action R V₂] [pseudo_metric_space R] [has_bounded_smul R V₂] [monoid_with_zero R'] [distrib_mul_action R' V₂] [pseudo_metric_space R'] [has_bounded_smul R' V₂] -instance : has_smul R (normed_group_hom V₁ V₂) := +instance : has_smul R (normed_add_group_hom V₁ V₂) := { smul := λ r f, { to_fun := r • f, map_add' := (r • f.to_add_monoid_hom).map_add', @@ -380,23 +386,24 @@ instance : has_smul R (normed_group_hom V₁ V₂) := exact hb x end⟩ } } -@[simp] lemma coe_smul (r : R) (f : normed_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl -@[simp] lemma smul_apply (r : R) (f : normed_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := rfl +@[simp] lemma coe_smul (r : R) (f : normed_add_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl +@[simp] lemma smul_apply (r : R) (f : normed_add_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := +rfl -instance [smul_comm_class R R' V₂] : smul_comm_class R R' (normed_group_hom V₁ V₂) := +instance [smul_comm_class R R' V₂] : smul_comm_class R R' (normed_add_group_hom V₁ V₂) := { smul_comm := λ r r' f, ext $ λ v, smul_comm _ _ _ } instance [has_smul R R'] [is_scalar_tower R R' V₂] : - is_scalar_tower R R' (normed_group_hom V₁ V₂) := + is_scalar_tower R R' (normed_add_group_hom V₁ V₂) := { smul_assoc := λ r r' f, ext $ λ v, smul_assoc _ _ _ } instance [distrib_mul_action Rᵐᵒᵖ V₂] [is_central_scalar R V₂] : - is_central_scalar R (normed_group_hom V₁ V₂) := + is_central_scalar R (normed_add_group_hom V₁ V₂) := { op_smul_eq_smul := λ r f, ext $ λ v, op_smul_eq_smul _ _ } end has_smul -instance has_nat_scalar : has_smul ℕ (normed_group_hom V₁ V₂) := +instance has_nat_scalar : has_smul ℕ (normed_add_group_hom V₁ V₂) := { smul := λ n f, { to_fun := n • f, map_add' := (n • f.to_add_monoid_hom).map_add', @@ -405,10 +412,11 @@ instance has_nat_scalar : has_smul ℕ (normed_group_hom V₁ V₂) := exact (norm_nsmul_le _ _).trans (mul_le_mul_of_nonneg_left (hb _) (nat.cast_nonneg _)), end⟩ } } -@[simp] lemma coe_nsmul (r : ℕ) (f : normed_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl -@[simp] lemma nsmul_apply (r : ℕ) (f : normed_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := rfl +@[simp] lemma coe_nsmul (r : ℕ) (f : normed_add_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl +@[simp] lemma nsmul_apply (r : ℕ) (f : normed_add_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := +rfl -instance has_int_scalar : has_smul ℤ (normed_group_hom V₁ V₂) := +instance has_int_scalar : has_smul ℤ (normed_add_group_hom V₁ V₂) := { smul := λ z f, { to_fun := z • f, map_add' := (z • f.to_add_monoid_hom).map_add', @@ -417,36 +425,39 @@ instance has_int_scalar : has_smul ℤ (normed_group_hom V₁ V₂) := exact (norm_zsmul_le _ _).trans (mul_le_mul_of_nonneg_left (hb _) $ norm_nonneg _), end⟩ } } -@[simp] lemma coe_zsmul (r : ℤ) (f : normed_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl -@[simp] lemma zsmul_apply (r : ℤ) (f : normed_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := rfl +@[simp] lemma coe_zsmul (r : ℤ) (f : normed_add_group_hom V₁ V₂) : ⇑(r • f) = r • f := rfl +@[simp] lemma zsmul_apply (r : ℤ) (f : normed_add_group_hom V₁ V₂) (v : V₁) : (r • f) v = r • f v := +rfl /-! ### Normed group structure on normed group homs -/ /-- Homs between two given normed groups form a commutative additive group. -/ -instance : add_comm_group (normed_group_hom V₁ V₂) := +instance : add_comm_group (normed_add_group_hom V₁ V₂) := coe_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) /-- Normed group homomorphisms themselves form a seminormed group with respect to the operator norm. -/ -instance to_semi_normed_group : semi_normed_group (normed_group_hom V₁ V₂) := -semi_normed_group.of_core _ ⟨op_norm_zero, op_norm_add_le, op_norm_neg⟩ +instance to_seminormed_add_comm_group : seminormed_add_comm_group (normed_add_group_hom V₁ V₂) := +seminormed_add_comm_group.of_core _ ⟨op_norm_zero, op_norm_add_le, op_norm_neg⟩ /-- Normed group homomorphisms themselves form a normed group with respect to the operator norm. -/ -instance to_normed_group {V₁ V₂ : Type*} [normed_group V₁] [normed_group V₂] : - normed_group (normed_group_hom V₁ V₂) := -normed_group.of_core _ ⟨λ f, op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ +instance to_normed_add_comm_group {V₁ V₂ : Type*} [normed_add_comm_group V₁] + [normed_add_comm_group V₂] : + normed_add_comm_group (normed_add_group_hom V₁ V₂) := +normed_add_comm_group.of_core _ ⟨λ f, op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ -/-- Coercion of a `normed_group_hom` is an `add_monoid_hom`. Similar to `add_monoid_hom.coe_fn` -/ +/-- Coercion of a `normed_add_group_hom` is an `add_monoid_hom`. Similar to `add_monoid_hom.coe_fn`. +-/ @[simps] -def coe_fn_add_hom : normed_group_hom V₁ V₂ →+ (V₁ → V₂) := +def coe_fn_add_hom : normed_add_group_hom V₁ V₂ →+ (V₁ → V₂) := { to_fun := coe_fn, map_zero' := coe_zero, map_add' := coe_add} -@[simp] lemma coe_sum {ι : Type*} (s : finset ι) (f : ι → normed_group_hom V₁ V₂) : +@[simp] lemma coe_sum {ι : Type*} (s : finset ι) (f : ι → normed_add_group_hom V₁ V₂) : ⇑(∑ i in s, f i) = ∑ i in s, (f i) := (coe_fn_add_hom : _ →+ (V₁ → V₂)).map_sum f s -lemma sum_apply {ι : Type*} (s : finset ι) (f : ι → normed_group_hom V₁ V₂) (v : V₁) : +lemma sum_apply {ι : Type*} (s : finset ι) (f : ι → normed_add_group_hom V₁ V₂) (v : V₁) : (∑ i in s, f i) v = ∑ i in s, (f i v) := by simp only [coe_sum, finset.sum_apply] @@ -454,68 +465,71 @@ by simp only [coe_sum, finset.sum_apply] instance {R : Type*} [monoid_with_zero R] [distrib_mul_action R V₂] [pseudo_metric_space R] [has_bounded_smul R V₂] : - distrib_mul_action R (normed_group_hom V₁ V₂) := + distrib_mul_action R (normed_add_group_hom V₁ V₂) := function.injective.distrib_mul_action coe_fn_add_hom coe_injective coe_smul instance {R : Type*} [semiring R] [module R V₂] [pseudo_metric_space R] [has_bounded_smul R V₂] : - module R (normed_group_hom V₁ V₂) := + module R (normed_add_group_hom V₁ V₂) := function.injective.module _ coe_fn_add_hom coe_injective coe_smul /-! ### Composition of normed group homs -/ /-- The composition of continuous normed group homs. -/ @[simps] -protected def comp (g : normed_group_hom V₂ V₃) (f : normed_group_hom V₁ V₂) : - normed_group_hom V₁ V₃ := -(g.to_add_monoid_hom.comp f.to_add_monoid_hom).mk_normed_group_hom (∥g∥ * ∥f∥) $ λ v, calc +protected def comp (g : normed_add_group_hom V₂ V₃) (f : normed_add_group_hom V₁ V₂) : + normed_add_group_hom V₁ V₃ := +(g.to_add_monoid_hom.comp f.to_add_monoid_hom).mk_normed_add_group_hom (∥g∥ * ∥f∥) $ λ v, calc ∥g (f v)∥ ≤ ∥g∥ * ∥f v∥ : le_op_norm _ _ ... ≤ ∥g∥ * (∥f∥ * ∥v∥) : mul_le_mul_of_nonneg_left (le_op_norm _ _) (op_norm_nonneg _) ... = ∥g∥ * ∥f∥ * ∥v∥ : by rw mul_assoc -lemma norm_comp_le (g : normed_group_hom V₂ V₃) (f : normed_group_hom V₁ V₂) : +lemma norm_comp_le (g : normed_add_group_hom V₂ V₃) (f : normed_add_group_hom V₁ V₂) : ∥g.comp f∥ ≤ ∥g∥ * ∥f∥ := -mk_normed_group_hom_norm_le _ (mul_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ +mk_normed_add_group_hom_norm_le _ (mul_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ -lemma norm_comp_le_of_le {g : normed_group_hom V₂ V₃} {C₁ C₂ : ℝ} (hg : ∥g∥ ≤ C₂) (hf : ∥f∥ ≤ C₁) : +lemma norm_comp_le_of_le {g : normed_add_group_hom V₂ V₃} {C₁ C₂ : ℝ} (hg : ∥g∥ ≤ C₂) + (hf : ∥f∥ ≤ C₁) : ∥g.comp f∥ ≤ C₂ * C₁ := le_trans (norm_comp_le g f) $ mul_le_mul hg hf (norm_nonneg _) (le_trans (norm_nonneg _) hg) -lemma norm_comp_le_of_le' {g : normed_group_hom V₂ V₃} (C₁ C₂ C₃ : ℝ) (h : C₃ = C₂ * C₁) +lemma norm_comp_le_of_le' {g : normed_add_group_hom V₂ V₃} (C₁ C₂ C₃ : ℝ) (h : C₃ = C₂ * C₁) (hg : ∥g∥ ≤ C₂) (hf : ∥f∥ ≤ C₁) : ∥g.comp f∥ ≤ C₃ := by { rw h, exact norm_comp_le_of_le hg hf } /-- Composition of normed groups hom as an additive group morphism. -/ -def comp_hom : (normed_group_hom V₂ V₃) →+ (normed_group_hom V₁ V₂) →+ (normed_group_hom V₁ V₃) := +def comp_hom : + normed_add_group_hom V₂ V₃ →+ normed_add_group_hom V₁ V₂ →+ normed_add_group_hom V₁ V₃ := add_monoid_hom.mk' (λ g, add_monoid_hom.mk' (λ f, g.comp f) (by { intros, ext, exact map_add g _ _ })) (by { intros, ext, simp only [comp_apply, pi.add_apply, function.comp_app, add_monoid_hom.add_apply, add_monoid_hom.mk'_apply, coe_add] }) -@[simp] lemma comp_zero (f : normed_group_hom V₂ V₃) : f.comp (0 : normed_group_hom V₁ V₂) = 0 := +@[simp] lemma comp_zero (f : normed_add_group_hom V₂ V₃) : + f.comp (0 : normed_add_group_hom V₁ V₂) = 0 := by { ext, exact map_zero f } -@[simp] lemma zero_comp (f : normed_group_hom V₁ V₂) : (0 : normed_group_hom V₂ V₃).comp f = 0 := +@[simp] lemma zero_comp (f : normed_add_group_hom V₁ V₂) : + (0 : normed_add_group_hom V₂ V₃).comp f = 0 := by { ext, refl } -lemma comp_assoc {V₄: Type* } [semi_normed_group V₄] (h : normed_group_hom V₃ V₄) - (g : normed_group_hom V₂ V₃) (f : normed_group_hom V₁ V₂) : +lemma comp_assoc {V₄: Type* } [seminormed_add_comm_group V₄] (h : normed_add_group_hom V₃ V₄) + (g : normed_add_group_hom V₂ V₃) (f : normed_add_group_hom V₁ V₂) : (h.comp g).comp f = h.comp (g.comp f) := by { ext, refl } -lemma coe_comp (f : normed_group_hom V₁ V₂) (g : normed_group_hom V₂ V₃) : +lemma coe_comp (f : normed_add_group_hom V₁ V₂) (g : normed_add_group_hom V₂ V₃) : (g.comp f : V₁ → V₃) = (g : V₂ → V₃) ∘ (f : V₁ → V₂) := rfl -end normed_group_hom +end normed_add_group_hom -namespace normed_group_hom +namespace normed_add_group_hom -variables {V W V₁ V₂ V₃ : Type*} -variables [semi_normed_group V] [semi_normed_group W] [semi_normed_group V₁] [semi_normed_group V₂] -[semi_normed_group V₃] +variables {V W V₁ V₂ V₃ : Type*} [seminormed_add_comm_group V] [seminormed_add_comm_group W] + [seminormed_add_comm_group V₁] [seminormed_add_comm_group V₂] [seminormed_add_comm_group V₃] /-- The inclusion of an `add_subgroup`, as bounded group homomorphism. -/ -@[simps] def incl (s : add_subgroup V) : normed_group_hom s V := +@[simps] def incl (s : add_subgroup V) : normed_add_group_hom s V := { to_fun := (coe : s → V), map_add' := λ v w, add_subgroup.coe_add _ _ _, bound' := ⟨1, λ v, by { rw [one_mul], refl }⟩ } @@ -525,10 +539,10 @@ rfl /-!### Kernel -/ section kernels -variables (f : normed_group_hom V₁ V₂) (g : normed_group_hom V₂ V₃) +variables (f : normed_add_group_hom V₁ V₂) (g : normed_add_group_hom V₂ V₃) /-- The kernel of a bounded group homomorphism. Naturally endowed with a -`semi_normed_group` instance. -/ +`seminormed_add_comm_group` instance. -/ def ker : add_subgroup V₁ := f.to_add_monoid_hom.ker lemma mem_ker (v : V₁) : v ∈ f.ker ↔ f v = 0 := @@ -537,7 +551,7 @@ by { erw f.to_add_monoid_hom.mem_ker, refl } /-- Given a normed group hom `f : V₁ → V₂` satisfying `g.comp f = 0` for some `g : V₂ → V₃`, the corestriction of `f` to the kernel of `g`. -/ @[simps] def ker.lift (h : g.comp f = 0) : - normed_group_hom V₁ g.ker := + normed_add_group_hom V₁ g.ker := { to_fun := λ v, ⟨f v, by { erw g.mem_ker, show (g.comp f) v = 0, rw h, refl }⟩, map_add' := λ v w, by { simp only [map_add], refl }, bound' := f.bound' } @@ -547,12 +561,12 @@ by { erw f.to_add_monoid_hom.mem_ker, refl } by { ext, refl } @[simp] -lemma ker_zero : (0 : normed_group_hom V₁ V₂).ker = ⊤ := +lemma ker_zero : (0 : normed_add_group_hom V₁ V₂).ker = ⊤ := by { ext, simp [mem_ker] } lemma coe_ker : (f.ker : set V₁) = (f : V₁ → V₂) ⁻¹' {0} := rfl -lemma is_closed_ker {V₂ : Type*} [normed_group V₂] (f : normed_group_hom V₁ V₂) : +lemma is_closed_ker {V₂ : Type*} [normed_add_comm_group V₂] (f : normed_add_group_hom V₁ V₂) : is_closed (f.ker : set V₁) := f.coe_ker ▸ is_closed.preimage f.continuous (t1_space.t1 0) @@ -561,10 +575,10 @@ end kernels /-! ### Range -/ section range -variables (f : normed_group_hom V₁ V₂) (g : normed_group_hom V₂ V₃) +variables (f : normed_add_group_hom V₁ V₂) (g : normed_add_group_hom V₂ V₃) /-- The image of a bounded group homomorphism. Naturally endowed with a -`semi_normed_group` instance. -/ +`seminormed_add_comm_group` instance. -/ def range : add_subgroup V₂ := f.to_add_monoid_hom.range lemma mem_range (v : V₂) : v ∈ f.range ↔ ∃ w, f w = v := @@ -586,10 +600,10 @@ by simpa [comp_range, incl_range, ← add_monoid_hom.range_eq_map] end range -variables {f : normed_group_hom V W} +variables {f : normed_add_group_hom V W} -/-- A `normed_group_hom` is *norm-nonincreasing* if `∥f v∥ ≤ ∥v∥` for all `v`. -/ -def norm_noninc (f : normed_group_hom V W) : Prop := +/-- A `normed_add_group_hom` is *norm-nonincreasing* if `∥f v∥ ≤ ∥v∥` for all `v`. -/ +def norm_noninc (f : normed_add_group_hom V W) : Prop := ∀ v, ∥f v∥ ≤ ∥v∥ namespace norm_noninc @@ -602,32 +616,32 @@ begin { simpa using le_of_op_norm_le f h v } end -lemma zero : (0 : normed_group_hom V₁ V₂).norm_noninc := +lemma zero : (0 : normed_add_group_hom V₁ V₂).norm_noninc := λ v, by simp lemma id : (id V).norm_noninc := λ v, le_rfl -lemma comp {g : normed_group_hom V₂ V₃} {f : normed_group_hom V₁ V₂} +lemma comp {g : normed_add_group_hom V₂ V₃} {f : normed_add_group_hom V₁ V₂} (hg : g.norm_noninc) (hf : f.norm_noninc) : (g.comp f).norm_noninc := λ v, (hg (f v)).trans (hf v) -@[simp] lemma neg_iff {f : normed_group_hom V₁ V₂} : (-f).norm_noninc ↔ f.norm_noninc := +@[simp] lemma neg_iff {f : normed_add_group_hom V₁ V₂} : (-f).norm_noninc ↔ f.norm_noninc := ⟨λ h x, by { simpa using h x }, λ h x, (norm_neg (f x)).le.trans (h x)⟩ end norm_noninc section isometry -lemma norm_eq_of_isometry {f : normed_group_hom V W} (hf : isometry f) (v : V) : +lemma norm_eq_of_isometry {f : normed_add_group_hom V W} (hf : isometry f) (v : V) : ∥f v∥ = ∥v∥ := (add_monoid_hom_class.isometry_iff_norm f).mp hf v lemma isometry_id : @isometry V V _ _ (id V) := isometry_id -lemma isometry_comp {g : normed_group_hom V₂ V₃} {f : normed_group_hom V₁ V₂} +lemma isometry_comp {g : normed_add_group_hom V₂ V₃} {f : normed_add_group_hom V₁ V₂} (hg : isometry g) (hf : isometry f) : isometry (g.comp f) := hg.comp hf @@ -637,59 +651,60 @@ lemma norm_noninc_of_isometry (hf : isometry f) : f.norm_noninc := end isometry -variables {W₁ W₂ W₃ : Type*} [semi_normed_group W₁] [semi_normed_group W₂] [semi_normed_group W₃] -variables (f) (g : normed_group_hom V W) -variables {f₁ g₁ : normed_group_hom V₁ W₁} -variables {f₂ g₂ : normed_group_hom V₂ W₂} -variables {f₃ g₃ : normed_group_hom V₃ W₃} +variables {W₁ W₂ W₃ : Type*} [seminormed_add_comm_group W₁] [seminormed_add_comm_group W₂] + [seminormed_add_comm_group W₃] +variables (f) (g : normed_add_group_hom V W) +variables {f₁ g₁ : normed_add_group_hom V₁ W₁} +variables {f₂ g₂ : normed_add_group_hom V₂ W₂} +variables {f₃ g₃ : normed_add_group_hom V₃ W₃} -/-- The equalizer of two morphisms `f g : normed_group_hom V W`. -/ +/-- The equalizer of two morphisms `f g : normed_add_group_hom V W`. -/ def equalizer := (f - g).ker namespace equalizer -/-- The inclusion of `f.equalizer g` as a `normed_group_hom`. -/ -def ι : normed_group_hom (f.equalizer g) V := incl _ +/-- The inclusion of `f.equalizer g` as a `normed_add_group_hom`. -/ +def ι : normed_add_group_hom (f.equalizer g) V := incl _ lemma comp_ι_eq : f.comp (ι f g) = g.comp (ι f g) := -by { ext, rw [comp_apply, comp_apply, ← sub_eq_zero, ← normed_group_hom.sub_apply], exact x.2 } +by { ext, rw [comp_apply, comp_apply, ← sub_eq_zero, ← normed_add_group_hom.sub_apply], exact x.2 } variables {f g} -/-- If `φ : normed_group_hom V₁ V` is such that `f.comp φ = g.comp φ`, the induced morphism -`normed_group_hom V₁ (f.equalizer g)`. -/ +/-- If `φ : normed_add_group_hom V₁ V` is such that `f.comp φ = g.comp φ`, the induced morphism +`normed_add_group_hom V₁ (f.equalizer g)`. -/ @[simps] -def lift (φ : normed_group_hom V₁ V) (h : f.comp φ = g.comp φ) : - normed_group_hom V₁ (f.equalizer g) := +def lift (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.comp φ) : + normed_add_group_hom V₁ (f.equalizer g) := { to_fun := λ v, ⟨φ v, show (f - g) (φ v) = 0, - by rw [normed_group_hom.sub_apply, sub_eq_zero, ← comp_apply, h, comp_apply]⟩, + by rw [normed_add_group_hom.sub_apply, sub_eq_zero, ← comp_apply, h, comp_apply]⟩, map_add' := λ v₁ v₂, by { ext, simp only [map_add, add_subgroup.coe_add, subtype.coe_mk] }, bound' := by { obtain ⟨C, C_pos, hC⟩ := φ.bound, exact ⟨C, hC⟩ } } -@[simp] lemma ι_comp_lift (φ : normed_group_hom V₁ V) (h : f.comp φ = g.comp φ) : +@[simp] lemma ι_comp_lift (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.comp φ) : (ι _ _).comp (lift φ h) = φ := by { ext, refl } /-- The lifting property of the equalizer as an equivalence. -/ @[simps] -def lift_equiv : {φ : normed_group_hom V₁ V // f.comp φ = g.comp φ} ≃ - normed_group_hom V₁ (f.equalizer g) := +def lift_equiv : {φ : normed_add_group_hom V₁ V // f.comp φ = g.comp φ} ≃ + normed_add_group_hom V₁ (f.equalizer g) := { to_fun := λ φ, lift φ φ.prop, inv_fun := λ ψ, ⟨(ι f g).comp ψ, by { rw [← comp_assoc, ← comp_assoc, comp_ι_eq] }⟩, left_inv := λ φ, by simp, right_inv := λ ψ, by { ext, refl } } -/-- Given `φ : normed_group_hom V₁ V₂` and `ψ : normed_group_hom W₁ W₂` such that +/-- Given `φ : normed_add_group_hom V₁ V₂` and `ψ : normed_add_group_hom W₁ W₂` such that `ψ.comp f₁ = f₂.comp φ` and `ψ.comp g₁ = g₂.comp φ`, the induced morphism -`normed_group_hom (f₁.equalizer g₁) (f₂.equalizer g₂)`. -/ -def map (φ : normed_group_hom V₁ V₂) (ψ : normed_group_hom W₁ W₂) +`normed_add_group_hom (f₁.equalizer g₁) (f₂.equalizer g₂)`. -/ +def map (φ : normed_add_group_hom V₁ V₂) (ψ : normed_add_group_hom W₁ W₂) (hf : ψ.comp f₁ = f₂.comp φ) (hg : ψ.comp g₁ = g₂.comp φ) : - normed_group_hom (f₁.equalizer g₁) (f₂.equalizer g₂) := + normed_add_group_hom (f₁.equalizer g₁) (f₂.equalizer g₂) := lift (φ.comp $ ι _ _) $ by { simp only [← comp_assoc, ← hf, ← hg], simp only [comp_assoc, comp_ι_eq] } -variables {φ : normed_group_hom V₁ V₂} {ψ : normed_group_hom W₁ W₂} -variables {φ' : normed_group_hom V₂ V₃} {ψ' : normed_group_hom W₂ W₃} +variables {φ : normed_add_group_hom V₁ V₂} {ψ : normed_add_group_hom W₁ W₂} +variables {φ' : normed_add_group_hom V₂ V₃} {ψ' : normed_add_group_hom W₂ W₃} @[simp] lemma ι_comp_map (hf : ψ.comp f₁ = f₂.comp φ) (hg : ψ.comp g₁ = g₂.comp φ) : (ι f₂ g₂).comp (map φ ψ hf hg) = φ.comp (ι _ _) := @@ -711,12 +726,13 @@ by { ext, refl } lemma ι_norm_noninc : (ι f g).norm_noninc := λ v, le_rfl /-- The lifting of a norm nonincreasing morphism is norm nonincreasing. -/ -lemma lift_norm_noninc (φ : normed_group_hom V₁ V) (h : f.comp φ = g.comp φ) (hφ : φ.norm_noninc) : +lemma lift_norm_noninc (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.comp φ) + (hφ : φ.norm_noninc) : (lift φ h).norm_noninc := hφ /-- If `φ` satisfies `∥φ∥ ≤ C`, then the same is true for the lifted morphism. -/ -lemma norm_lift_le (φ : normed_group_hom V₁ V) (h : f.comp φ = g.comp φ) +lemma norm_lift_le (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.comp φ) (C : ℝ) (hφ : ∥φ∥ ≤ C) : ∥(lift φ h)∥ ≤ C := hφ lemma map_norm_noninc (hf : ψ.comp f₁ = f₂.comp φ) (hg : ψ.comp g₁ = g₂.comp φ) @@ -729,20 +745,20 @@ norm_lift_le _ _ _ hφ end equalizer -end normed_group_hom +end normed_add_group_hom section controlled_closure open filter finset open_locale topological_space -variables {G : Type*} [normed_group G] [complete_space G] -variables {H : Type*} [normed_group H] +variables {G : Type*} [normed_add_comm_group G] [complete_space G] +variables {H : Type*} [normed_add_comm_group H] -/-- Given `f : normed_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every +/-- 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_group_hom G H} {K : add_subgroup H} +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 @@ -773,7 +789,7 @@ begin `b` ensures `s` is Cauchy. -/ set s : ℕ → G := λ n, ∑ k in range (n+1), u k, have : cauchy_seq s, - { apply normed_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, + { 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 @@ -817,15 +833,15 @@ begin ... ≤ (C+ε)*∥h∥ : by { rw [add_comm, add_mul], apply add_le_add_left this } } end -/-- Given `f : normed_group_hom G H` for some complete `G`, if every element `x` of the image of -an isometric immersion `j : normed_group_hom K H` has a preimage under `f` whose norm is at most +/-- 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_group_hom G H} - {K : Type*} [semi_normed_group K] {j : normed_group_hom K H} (hj : ∀ x, ∥j x∥ = ∥x∥) +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 diff --git a/src/analysis/normed/group/hom_completion.lean b/src/analysis/normed/group/hom_completion.lean index 912a430f67e0f..2e3227117b81f 100644 --- a/src/analysis/normed/group/hom_completion.lean +++ b/src/analysis/normed/group/hom_completion.lean @@ -9,9 +9,9 @@ import analysis.normed.group.completion /-! # Completion of normed group homs -Given two (semi) normed groups `G` and `H` and a normed group hom `f : normed_group_hom G H`, +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_group_hom (completion G) (completion H)` such that the diagram +`f.completion : normed_add_group_hom (completion G) (completion H)` such that the diagram ``` f @@ -33,32 +33,32 @@ The vertical maps in the above diagrams are also normed group homs constructed i ## Main definitions and results: -* `normed_group_hom.completion`: see the discussion above. -* `normed_group.to_compl : normed_group_hom G (completion G)`: the canonical map from `G` to its - completion, as a normed group hom -* `normed_group_hom.completion_to_compl`: the above diagram indeed commutes. -* `normed_group_hom.norm_completion`: `∥f.completion∥ = ∥f∥` -* `normed_group_hom.ker_le_ker_completion`: the kernel of `f.completion` contains the image of the - kernel of `f`. -* `normed_group_hom.ker_completion`: the kernel of `f.completion` is the closure of the image of the - kernel of `f` under an assumption that `f` is quantitatively surjective onto its image. -* `normed_group_hom.extension` : if `H` is complete, the extension of `f : normed_group_hom G H` - to a `normed_group_hom (completion G) H`. +* `normed_add_group_hom.completion`: see the discussion above. +* `normed_add_comm_group.to_compl : normed_add_group_hom G (completion G)`: the canonical map from + `G` to its completion, as a normed group hom +* `normed_add_group_hom.completion_to_compl`: the above diagram indeed commutes. +* `normed_add_group_hom.norm_completion`: `∥f.completion∥ = ∥f∥` +* `normed_add_group_hom.ker_le_ker_completion`: the kernel of `f.completion` contains the image of + the kernel of `f`. +* `normed_add_group_hom.ker_completion`: the kernel of `f.completion` is the closure of the image of + the kernel of `f` under an assumption that `f` is quantitatively surjective onto its image. +* `normed_add_group_hom.extension` : if `H` is complete, the extension of + `f : normed_add_group_hom G H` to a `normed_add_group_hom (completion G) H`. -/ noncomputable theory -open set normed_group_hom uniform_space +open set normed_add_group_hom uniform_space section completion -variables {G : Type*} [semi_normed_group G] -variables {H : Type*} [semi_normed_group H] -variables {K : Type*} [semi_normed_group K] +variables {G : Type*} [seminormed_add_comm_group G] +variables {H : Type*} [seminormed_add_comm_group H] +variables {K : Type*} [seminormed_add_comm_group K] /-- The normed group hom induced between completions. -/ -def normed_group_hom.completion (f : normed_group_hom G H) : - normed_group_hom (completion G) (completion H) := +def normed_add_group_hom.completion (f : normed_add_group_hom G H) : + normed_add_group_hom (completion G) (completion H) := { bound' := begin use ∥f∥, intro y, @@ -74,22 +74,23 @@ def normed_group_hom.completion (f : normed_group_hom G H) : end, ..f.to_add_monoid_hom.completion f.continuous } -lemma normed_group_hom.completion_def (f : normed_group_hom G H) (x : completion G) : +lemma normed_add_group_hom.completion_def (f : normed_add_group_hom G H) (x : completion G) : f.completion x = completion.map f x := rfl @[simp] -lemma normed_group_hom.completion_coe_to_fun (f : normed_group_hom G H) : +lemma normed_add_group_hom.completion_coe_to_fun (f : normed_add_group_hom G H) : (f.completion : completion G → completion H) = completion.map f := -by { ext x, exact normed_group_hom.completion_def f x } +by { ext x, exact normed_add_group_hom.completion_def f x } @[simp] -lemma normed_group_hom.completion_coe (f : normed_group_hom G H) (g : G) : f.completion g = f g := +lemma normed_add_group_hom.completion_coe (f : normed_add_group_hom G H) (g : G) : + f.completion g = f g := completion.map_coe f.uniform_continuous _ /-- Completion of normed group homs as a normed group hom. -/ -@[simps] def normed_group_hom_completion_hom : - normed_group_hom G H →+ normed_group_hom (completion G) (completion H) := -{ to_fun := normed_group_hom.completion, +@[simps] def normed_add_group_hom_completion_hom : + normed_add_group_hom G H →+ normed_add_group_hom (completion G) (completion H) := +{ to_fun := normed_add_group_hom.completion, map_zero' := begin apply to_add_monoid_hom_injective, exact add_monoid_hom.completion_zero @@ -100,57 +101,58 @@ completion.map_coe f.uniform_continuous _ end } @[simp] -lemma normed_group_hom.completion_id : - (normed_group_hom.id G).completion = normed_group_hom.id (completion G) := +lemma normed_add_group_hom.completion_id : + (normed_add_group_hom.id G).completion = normed_add_group_hom.id (completion G) := begin ext x, - rw [normed_group_hom.completion_def, normed_group_hom.coe_id, completion.map_id], + rw [normed_add_group_hom.completion_def, normed_add_group_hom.coe_id, completion.map_id], refl end -lemma normed_group_hom.completion_comp (f : normed_group_hom G H) (g : normed_group_hom H K) : +lemma normed_add_group_hom.completion_comp (f : normed_add_group_hom G H) + (g : normed_add_group_hom H K) : g.completion.comp f.completion = (g.comp f).completion := begin ext x, - rw [normed_group_hom.coe_comp, normed_group_hom.completion_def, - normed_group_hom.completion_coe_to_fun, normed_group_hom.completion_coe_to_fun, - completion.map_comp (normed_group_hom.uniform_continuous _) - (normed_group_hom.uniform_continuous _)], + rw [normed_add_group_hom.coe_comp, normed_add_group_hom.completion_def, + normed_add_group_hom.completion_coe_to_fun, normed_add_group_hom.completion_coe_to_fun, + completion.map_comp (normed_add_group_hom.uniform_continuous _) + (normed_add_group_hom.uniform_continuous _)], refl end -lemma normed_group_hom.completion_neg (f : normed_group_hom G H) : +lemma normed_add_group_hom.completion_neg (f : normed_add_group_hom G H) : (-f).completion = -f.completion := -map_neg (normed_group_hom_completion_hom : normed_group_hom G H →+ _) f +map_neg (normed_add_group_hom_completion_hom : normed_add_group_hom G H →+ _) f -lemma normed_group_hom.completion_add (f g : normed_group_hom G H) : +lemma normed_add_group_hom.completion_add (f g : normed_add_group_hom G H) : (f + g).completion = f.completion + g.completion := -normed_group_hom_completion_hom.map_add f g +normed_add_group_hom_completion_hom.map_add f g -lemma normed_group_hom.completion_sub (f g : normed_group_hom G H) : +lemma normed_add_group_hom.completion_sub (f g : normed_add_group_hom G H) : (f - g).completion = f.completion - g.completion := -map_sub (normed_group_hom_completion_hom : normed_group_hom G H →+ _) f g +map_sub (normed_add_group_hom_completion_hom : normed_add_group_hom G H →+ _) f g @[simp] -lemma normed_group_hom.zero_completion : (0 : normed_group_hom G H).completion = 0 := -normed_group_hom_completion_hom.map_zero +lemma normed_add_group_hom.zero_completion : (0 : normed_add_group_hom G H).completion = 0 := +normed_add_group_hom_completion_hom.map_zero /-- The map from a normed group to its completion, as a normed group hom. -/ -def normed_group.to_compl : normed_group_hom G (completion G) := +def normed_add_comm_group.to_compl : normed_add_group_hom G (completion G) := { to_fun := coe, map_add' := completion.to_compl.map_add, bound' := ⟨1, by simp [le_refl]⟩ } -open normed_group +open normed_add_comm_group -lemma normed_group.norm_to_compl (x : G) : ∥to_compl x∥ = ∥x∥ := +lemma normed_add_comm_group.norm_to_compl (x : G) : ∥to_compl x∥ = ∥x∥ := completion.norm_coe x -lemma normed_group.dense_range_to_compl : dense_range (to_compl : G → completion G) := +lemma normed_add_comm_group.dense_range_to_compl : dense_range (to_compl : G → completion G) := completion.dense_inducing_coe.dense @[simp] -lemma normed_group_hom.completion_to_compl (f : normed_group_hom G H) : +lemma normed_add_group_hom.completion_to_compl (f : normed_add_group_hom G H) : f.completion.comp to_compl = to_compl.comp f := begin ext x, @@ -158,7 +160,8 @@ begin simpa end -@[simp] lemma normed_group_hom.norm_completion (f : normed_group_hom G H) : ∥f.completion∥ = ∥f∥ := +@[simp] lemma normed_add_group_hom.norm_completion (f : normed_add_group_hom G H) : + ∥f.completion∥ = ∥f∥ := begin apply f.completion.op_norm_eq_of_bounds (norm_nonneg _), { intro x, @@ -173,7 +176,7 @@ begin simpa using hN x }, end -lemma normed_group_hom.ker_le_ker_completion (f : normed_group_hom G H) : +lemma normed_add_group_hom.ker_le_ker_completion (f : normed_add_group_hom G H) : (to_compl.comp $ incl f.ker).range ≤ f.completion.ker := begin intros a h, @@ -181,24 +184,24 @@ begin rcases h with ⟨⟨g, g_in : g ∈ f.ker⟩, rfl⟩, rw f.mem_ker at g_in, change f.completion (g : completion G) = 0, - simp [normed_group_hom.mem_ker, f.completion_coe g, g_in, completion.coe_zero], + simp [normed_add_group_hom.mem_ker, f.completion_coe g, g_in, completion.coe_zero], end -lemma normed_group_hom.ker_completion {f : normed_group_hom G H} {C : ℝ} +lemma normed_add_group_hom.ker_completion {f : normed_add_group_hom G H} {C : ℝ} (h : f.surjective_on_with f.range C) : (f.completion.ker : set $ completion G) = closure (to_compl.comp $ incl f.ker).range := begin rcases h.exists_pos with ⟨C', C'_pos, hC'⟩, apply le_antisymm, { intros hatg hatg_in, - rw semi_normed_group.mem_closure_iff, + rw seminormed_add_comm_group.mem_closure_iff, intros ε ε_pos, have hCf : 0 ≤ C'*∥f∥ := (zero_le_mul_left C'_pos).mpr (norm_nonneg f), have ineq : 0 < 1 + C'*∥f∥, by linarith, set δ := ε/(1 + C'*∥f∥), have δ_pos : δ > 0, from div_pos ε_pos ineq, - obtain ⟨_, ⟨g : G, rfl⟩, hg : ∥hatg - g∥ < δ⟩ := - semi_normed_group.mem_closure_iff.mp (completion.dense_inducing_coe.dense hatg) δ δ_pos, + obtain ⟨_, ⟨g : G, rfl⟩, hg : ∥hatg - g∥ < δ⟩ := seminormed_add_comm_group.mem_closure_iff.mp + (completion.dense_inducing_coe.dense hatg) δ δ_pos, obtain ⟨g' : G, hgg' : f g' = f g, hfg : ∥g'∥ ≤ C' * ∥f g∥⟩ := hC' (f g) (mem_range_self g), have mem_ker : g - g' ∈ f.ker, @@ -219,7 +222,7 @@ begin exact (mul_le_mul_left C'_pos).mpr this }, refine ⟨g - g', _, _⟩, { norm_cast, - rw normed_group_hom.comp_range, + rw normed_add_group_hom.comp_range, apply add_subgroup.mem_map_of_mem, simp only [incl_range, mem_ker] }, { calc ∥hatg - (g - g')∥ = ∥hatg - g + g'∥ : by abel @@ -236,12 +239,13 @@ end completion section extension -variables {G : Type*} [semi_normed_group G] -variables {H : Type*} [semi_normed_group H] [separated_space H] [complete_space H] +variables {G : Type*} [seminormed_add_comm_group G] +variables {H : Type*} [seminormed_add_comm_group H] [separated_space H] [complete_space H] -/-- If `H` is complete, the extension of `f : normed_group_hom G H` to a -`normed_group_hom (completion G) H`. -/ -def normed_group_hom.extension (f : normed_group_hom G H) : normed_group_hom (completion G) H := +/-- If `H` is complete, the extension of `f : normed_add_group_hom G H` to a +`normed_add_group_hom (completion G) H`. -/ +def normed_add_group_hom.extension (f : normed_add_group_hom G H) : + normed_add_group_hom (completion G) H := { bound' := begin refine ⟨∥f∥, λ v, completion.induction_on v (is_closed_le _ _) (λ a, _)⟩, { exact continuous.comp continuous_norm completion.continuous_extension }, @@ -251,20 +255,20 @@ def normed_group_hom.extension (f : normed_group_hom G H) : normed_group_hom (co end, ..f.to_add_monoid_hom.extension f.continuous } -lemma normed_group_hom.extension_def (f : normed_group_hom G H) (v : G) : +lemma normed_add_group_hom.extension_def (f : normed_add_group_hom G H) (v : G) : f.extension v = completion.extension f v := rfl -@[simp] lemma normed_group_hom.extension_coe (f : normed_group_hom G H) (v : G) : +@[simp] lemma normed_add_group_hom.extension_coe (f : normed_add_group_hom G H) (v : G) : f.extension v = f v := add_monoid_hom.extension_coe _ f.continuous _ -lemma normed_group_hom.extension_coe_to_fun (f : normed_group_hom G H) : +lemma normed_add_group_hom.extension_coe_to_fun (f : normed_add_group_hom G H) : (f.extension : (completion G) → H) = completion.extension f := rfl -lemma normed_group_hom.extension_unique (f : normed_group_hom G H) - {g : normed_group_hom (completion G) H} (hg : ∀ v, f v = g v) : f.extension = g := +lemma normed_add_group_hom.extension_unique (f : normed_add_group_hom G H) + {g : normed_add_group_hom (completion G) H} (hg : ∀ v, f v = g v) : f.extension = g := begin ext v, - rw [normed_group_hom.extension_coe_to_fun, completion.extension_unique f.uniform_continuous + rw [normed_add_group_hom.extension_coe_to_fun, completion.extension_unique f.uniform_continuous g.uniform_continuous (λ a, hg a)] end diff --git a/src/analysis/normed/group/infinite_sum.lean b/src/analysis/normed/group/infinite_sum.lean index e6a9535f972f4..79d5aa4ad15c5 100644 --- a/src/analysis/normed/group/infinite_sum.lean +++ b/src/analysis/normed/group/infinite_sum.lean @@ -30,7 +30,7 @@ infinite series, absolute convergence, normed group open_locale classical big_operators topological_space nnreal open finset filter metric -variables {ι α E F : Type*} [semi_normed_group E] [semi_normed_group F] +variables {ι α E F : Type*} [seminormed_add_comm_group E] [seminormed_add_comm_group F] lemma cauchy_seq_finset_iff_vanishing_norm {f : ι → E} : cauchy_seq (λ s : finset ι, ∑ i in s, f i) ↔ diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index d355e95c1835c..2b71c2037ff92 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -16,9 +16,9 @@ Notably, we show that the sum of bounded sets remain bounded. open metric set open_locale pointwise topological_space -section semi_normed_group +section seminormed_add_comm_group -variables {E : Type*} [semi_normed_group E] {ε δ : ℝ} {s t : set E} {x y : E} +variables {E : Type*} [seminormed_add_comm_group E] {ε δ : ℝ} {s t : set E} {x y : E} lemma bounded_iff_exists_norm_le : bounded s ↔ ∃ R, ∀ x ∈ s, ∥x∥ ≤ R := by simp [subset_def, bounded_iff_subset_ball (0 : E)] @@ -170,4 +170,4 @@ lemma is_compact.closed_ball_sub (hs : is_compact s) (hδ : 0 ≤ δ) (x : E) : closed_ball x δ + s = x +ᵥ cthickening δ s := by simp [sub_eq_add_neg, add_comm, hs.closed_ball_add hδ] -end semi_normed_group +end seminormed_add_comm_group diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index ded8b1adfc0f6..d93fc7cb894a7 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -8,9 +8,9 @@ import analysis.normed.group.hom /-! # Quotients of seminormed groups -For any `semi_normed_group M` and any `S : add_subgroup M`, we provide a `semi_normed_group` -the group quotient `M ⧸ S`. -If `S` is closed, we provide `normed_group (M ⧸ S)` (regardless of whether `M` itself is +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 separated). The two main properties of these structures are the underlying topology is the quotient topology and the projection is a normed group homomorphism which is norm non-increasing (better, it has operator norm exactly one unless `S` is dense in `M`). The corresponding @@ -28,19 +28,19 @@ We use `M` and `N` to denote seminormed groups and `S : add_subgroup M`. All the following definitions are in the `add_subgroup` namespace. Hence we can access `add_subgroup.normed_mk S` as `S.normed_mk`. -* `semi_normed_group_quotient` : The seminormed group structure on the quotient by +* `seminormed_add_comm_group_quotient` : The seminormed group structure on the quotient by an additive subgroup. This is an instance so there is no need to explictly use it. -* `normed_group_quotient` : The normed group structure on the quotient by +* `normed_add_comm_group_quotient` : The normed group structure on the quotient by a closed additive subgroup. This is an instance so there is no need to explictly use it. * `normed_mk S` : the normed group hom from `M` to `M ⧸ S`. * `lift S f hf`: implements the universal property of `M ⧸ S`. Here - `(f : normed_group_hom M N)`, `(hf : ∀ s ∈ S, f s = 0)` and - `lift S f hf : normed_group_hom (M ⧸ S) N`. + `(f : normed_add_group_hom M N)`, `(hf : ∀ s ∈ S, f s = 0)` and + `lift S f hf : normed_add_group_hom (M ⧸ S) N`. -* `is_quotient`: given `f : normed_group_hom M N`, `is_quotient f` means `N` is isomorphic +* `is_quotient`: given `f : normed_add_group_hom M N`, `is_quotient f` means `N` is isomorphic to a quotient of `M` by a subgroup, with projection `f`. Technically it asserts `f` is surjective and the norm of `f x` is the infimum of the norms of `x + m` for `m` in `f.ker`. @@ -54,12 +54,12 @@ All the following definitions are in the `add_subgroup` namespace. Hence we can ## Implementation details -For any `semi_normed_group M` and any `S : add_subgroup M` we define a norm on `M ⧸ S` by +For any `seminormed_add_comm_group M` and any `S : add_subgroup M` we define a norm on `M ⧸ S` by `∥x∥ = Inf (norm '' {m | mk' S m = x})`. This formula is really an implementation detail, it shouldn't be needed outside of this file setting up the theory. Since `M ⧸ S` is automatically a topological space (as any quotient of a topological space), -one needs to be careful while defining the `semi_normed_group` instance to avoid having two +one needs to be careful while defining the `seminormed_add_comm_group` instance to avoid having two different topologies on this quotient. This is not purely a technological issue. Mathematically there is something to prove. The main point is proved in the auxiliary lemma `quotient_nhd_basis` that has no use beyond this verification and states that zero in the quotient @@ -70,7 +70,7 @@ is not good enough for the type class system. As usual we ensure *definitional* using forgetful inheritance, see Note [forgetful inheritance]. A (semi)-normed group structure includes a uniform space structure which includes a topological space structure, together with propositional fields asserting compatibility conditions. -The usual way to define a `semi_normed_group` is to let Lean build a uniform space structure +The usual way to define a `seminormed_add_comm_group` is to let Lean build a uniform space structure using the provided norm, and then trivially build a proof that the norm and uniform structure are compatible. Here the uniform structure is provided using `topological_add_group.to_uniform_space` which uses the topological structure and the group structure to build the uniform structure. This @@ -86,7 +86,7 @@ noncomputable theory open quotient_add_group metric set open_locale topological_space nnreal -variables {M N : Type*} [semi_normed_group M] [semi_normed_group N] +variables {M N : Type*} [seminormed_add_comm_group M] [seminormed_add_comm_group N] /-- The definition of the norm on the quotient by an additive subgroup. -/ noncomputable @@ -281,8 +281,8 @@ end⟩ /-- The seminormed group structure on the quotient by an additive subgroup. -/ noncomputable -instance add_subgroup.semi_normed_group_quotient (S : add_subgroup M) : - semi_normed_group (M ⧸ S) := +instance add_subgroup.seminormed_add_comm_group_quotient (S : add_subgroup M) : + seminormed_add_comm_group (M ⧸ S) := { dist := λ x y, ∥x - y∥, dist_self := λ x, by simp only [norm_mk_zero, sub_self], dist_comm := quotient_norm_sub_rev, @@ -318,35 +318,35 @@ instance add_subgroup.semi_normed_group_quotient (S : add_subgroup M) : -- This is a sanity check left here on purpose to ensure that potential refactors won't destroy -- this important property. example (S : add_subgroup M) : (quotient.topological_space : topological_space $ M ⧸ S) = -S.semi_normed_group_quotient.to_uniform_space.to_topological_space := +S.seminormed_add_comm_group_quotient.to_uniform_space.to_topological_space := rfl /-- The quotient in the category of normed groups. -/ noncomputable -instance add_subgroup.normed_group_quotient (S : add_subgroup M) [hS : is_closed (S : set M)] : - normed_group (M ⧸ S) := +instance add_subgroup.normed_add_comm_group_quotient (S : add_subgroup M) [is_closed (S : set M)] : + normed_add_comm_group (M ⧸ S) := { eq_of_dist_eq_zero := begin rintros ⟨m⟩ ⟨m'⟩ (h : ∥mk' S m - mk' S m'∥ = 0), - erw [← (mk' S).map_sub, quotient_norm_eq_zero_iff, hS.closure_eq, + erw [← (mk' S).map_sub, quotient_norm_eq_zero_iff, ‹is_closed _›.closure_eq, ← quotient_add_group.eq_iff_sub_mem] at h, exact h end, - .. add_subgroup.semi_normed_group_quotient S } + .. add_subgroup.seminormed_add_comm_group_quotient S } -- This is a sanity check left here on purpose to ensure that potential refactors won't destroy -- this important property. example (S : add_subgroup M) [is_closed (S : set M)] : - S.semi_normed_group_quotient = normed_group.to_semi_normed_group := rfl + S.seminormed_add_comm_group_quotient = normed_add_comm_group.to_seminormed_add_comm_group := rfl namespace add_subgroup -open normed_group_hom +open normed_add_group_hom /-- The morphism from a seminormed group to the quotient by a subgroup. -/ noncomputable -def normed_mk (S : add_subgroup M) : normed_group_hom M (M ⧸ S) := +def normed_mk (S : add_subgroup M) : normed_add_group_hom M (M ⧸ S) := { bound' := ⟨1, λ m, by simpa [one_mul] using quotient_norm_mk_le _ m⟩, .. quotient_add_group.mk' S } @@ -365,7 +365,7 @@ quotient_add_group.ker_mk _ /-- The operator norm of the projection is at most `1`. -/ lemma norm_normed_mk_le (S : add_subgroup M) : ∥S.normed_mk∥ ≤ 1 := -normed_group_hom.op_norm_le_bound _ zero_le_one (λ m, by simp [quotient_norm_mk_le']) +normed_add_group_hom.op_norm_le_bound _ zero_le_one (λ m, by simp [quotient_norm_mk_le']) /-- The operator norm of the projection is `1` if the subspace is not dense. -/ lemma norm_normed_mk (S : add_subgroup M) (h : (S.topological_closure : set M) ≠ univ) : @@ -420,20 +420,20 @@ end end add_subgroup -namespace normed_group_hom +namespace normed_add_group_hom /-- `is_quotient f`, for `f : M ⟶ N` means that `N` is isomorphic to the quotient of `M` by the kernel of `f`. -/ -structure is_quotient (f : normed_group_hom M N) : Prop := +structure is_quotient (f : normed_add_group_hom M N) : Prop := (surjective : function.surjective f) (norm : ∀ x, ∥f x∥ = Inf ((λ m, ∥x + m∥) '' f.ker)) -/-- Given `f : normed_group_hom M N` such that `f s = 0` for all `s ∈ S`, where, -`S : add_subgroup M` is closed, the induced morphism `normed_group_hom (M ⧸ S) N`. -/ +/-- Given `f : normed_add_group_hom M N` such that `f s = 0` for all `s ∈ S`, where, +`S : add_subgroup M` is closed, the induced morphism `normed_add_group_hom (M ⧸ S) N`. -/ noncomputable -def lift {N : Type*} [semi_normed_group N] (S : add_subgroup M) - (f : normed_group_hom M N) (hf : ∀ s ∈ S, f s = 0) : - normed_group_hom (M ⧸ S) N := +def lift {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) + (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) : + normed_add_group_hom (M ⧸ S) N := { bound' := begin obtain ⟨c : ℝ, hcpos : (0 : ℝ) < c, hc : ∀ x, ∥f x∥ ≤ c * ∥x∥⟩ := f.bound, @@ -446,13 +446,13 @@ def lift {N : Type*} [semi_normed_group N] (S : add_subgroup M) end, .. quotient_add_group.lift S f.to_add_monoid_hom hf } -lemma lift_mk {N : Type*} [semi_normed_group N] (S : add_subgroup M) - (f : normed_group_hom M N) (hf : ∀ s ∈ S, f s = 0) (m : M) : +lemma lift_mk {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) + (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) (m : M) : lift S f hf (S.normed_mk m) = f m := rfl -lemma lift_unique {N : Type*} [semi_normed_group N] (S : add_subgroup M) - (f : normed_group_hom M N) (hf : ∀ s ∈ S, f s = 0) - (g : normed_group_hom (M ⧸ S) N) : +lemma lift_unique {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) + (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) + (g : normed_add_group_hom (M ⧸ S) N) : g.comp (S.normed_mk) = f → g = lift S f hf := begin intro h, @@ -466,8 +466,8 @@ end lemma is_quotient_quotient (S : add_subgroup M) : is_quotient (S.normed_mk) := ⟨S.surjective_normed_mk, λ m, by simpa [S.ker_normed_mk] using quotient_norm_mk_eq _ m⟩ -lemma is_quotient.norm_lift {f : normed_group_hom M N} (hquot : is_quotient f) {ε : ℝ} (hε : 0 < ε) - (n : N) : ∃ (m : M), f m = n ∧ ∥m∥ < ∥n∥ + ε := +lemma is_quotient.norm_lift {f : normed_add_group_hom M N} (hquot : is_quotient f) {ε : ℝ} + (hε : 0 < ε) (n : N) : ∃ (m : M), f m = n ∧ ∥m∥ < ∥n∥ + ε := begin obtain ⟨m, rfl⟩ := hquot.surjective n, have nonemp : ((λ m', ∥m + m'∥) '' f.ker).nonempty, @@ -475,11 +475,11 @@ begin exact ⟨0, f.ker.zero_mem⟩ }, rcases real.lt_Inf_add_pos nonemp hε with ⟨_, ⟨⟨x, hx, rfl⟩, H : ∥m + x∥ < Inf ((λ (m' : M), ∥m + m'∥) '' f.ker) + ε⟩⟩, - exact ⟨m+x, by rw [map_add,(normed_group_hom.mem_ker f x).mp hx, add_zero], + exact ⟨m+x, by rw [map_add,(normed_add_group_hom.mem_ker f x).mp hx, add_zero], by rwa hquot.norm⟩, end -lemma is_quotient.norm_le {f : normed_group_hom M N} (hquot : is_quotient f) (m : M) : +lemma is_quotient.norm_le {f : normed_add_group_hom M N} (hquot : is_quotient f) (m : M) : ∥f m∥ ≤ ∥m∥ := begin rw hquot.norm, @@ -490,8 +490,8 @@ begin { exact ⟨0, f.ker.zero_mem, by simp⟩ } end -lemma lift_norm_le {N : Type*} [semi_normed_group N] (S : add_subgroup M) - (f : normed_group_hom M N) (hf : ∀ s ∈ S, f s = 0) +lemma lift_norm_le {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) + (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) {c : ℝ≥0} (fb : ∥f∥ ≤ c) : ∥lift S f hf∥ ≤ c := begin @@ -517,8 +517,8 @@ begin { rw [mul_add, mul_div_cancel'], exact_mod_cast hc.ne' } }, end -lemma lift_norm_noninc {N : Type*} [semi_normed_group N] (S : add_subgroup M) - (f : normed_group_hom M N) (hf : ∀ s ∈ S, f s = 0) +lemma lift_norm_noninc {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) + (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) (fb : f.norm_noninc) : (lift S f hf).norm_noninc := λ x, @@ -527,4 +527,4 @@ begin simpa using le_of_op_norm_le _ (f.lift_norm_le _ _ fb') _, end -end normed_group_hom +end normed_add_group_hom diff --git a/src/analysis/normed_space/M_structure.lean b/src/analysis/normed_space/M_structure.lean index 71a47bc0dad31..83daa7c1b1d29 100644 --- a/src/analysis/normed_space/M_structure.lean +++ b/src/analysis/normed_space/M_structure.lean @@ -58,7 +58,7 @@ M-summand, M-projection, L-summand, L-projection, M-ideal, M-structure -/ -variables (X : Type*) [normed_group X] +variables (X : Type*) [normed_add_comm_group 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`, diff --git a/src/analysis/normed_space/add_torsor.lean b/src/analysis/normed_space/add_torsor.lean index 340c15bbe84f4..6beb40bab3293 100644 --- a/src/analysis/normed_space/add_torsor.lean +++ b/src/analysis/normed_space/add_torsor.lean @@ -19,8 +19,8 @@ noncomputable theory open_locale nnreal topological_space open filter -variables {α V P : Type*} [semi_normed_group V] [pseudo_metric_space P] [normed_add_torsor V P] -variables {W Q : Type*} [normed_group W] [metric_space Q] [normed_add_torsor W Q] +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] section normed_space @@ -128,7 +128,7 @@ variables (𝕜) lemma eventually_homothety_mem_of_mem_interior (x : Q) {s : set Q} {y : Q} (hy : y ∈ interior s) : ∀ᶠ δ in 𝓝 (1 : 𝕜), homothety x δ y ∈ s := begin - rw (normed_group.nhds_basis_norm_lt (1 : 𝕜)).eventually_iff, + rw (normed_add_comm_group.nhds_basis_norm_lt (1 : 𝕜)).eventually_iff, cases eq_or_ne y x with h h, { use 1, simp [h.symm, interior_subset hy], }, have hxy : 0 < ∥y -ᵥ x∥, { rwa [norm_pos_iff, vsub_ne_zero], }, obtain ⟨u, hu₁, hu₂, hu₃⟩ := mem_interior.mp hy, diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index bd01b1fc1a814..f219e92a065d0 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -27,7 +27,7 @@ This file contains results about bases in normed affine spaces. section barycentric variables {ι 𝕜 E P : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] -variables [normed_group E] [normed_space 𝕜 E] [finite_dimensional 𝕜 E] +variables [normed_add_comm_group E] [normed_space 𝕜 E] [finite_dimensional 𝕜 E] variables [metric_space P] [normed_add_torsor E P] variables (b : affine_basis ι 𝕜 P) @@ -54,8 +54,8 @@ to this basis. TODO Restate this result for affine spaces (instead of vector spaces) once the definition of convexity is generalised to this setting. -/ -lemma interior_convex_hull_aff_basis {ι E : Type*} [fintype ι] [normed_group E] [normed_space ℝ E] - (b : affine_basis ι ℝ E) : +lemma interior_convex_hull_aff_basis {ι E : Type*} [fintype ι] [normed_add_comm_group E] + [normed_space ℝ E] (b : affine_basis ι ℝ E) : interior (convex_hull ℝ (range b.points)) = { x | ∀ i, 0 < b.coord i x } := begin casesI subsingleton_or_nontrivial ι, @@ -77,7 +77,8 @@ begin interior_Ici, mem_Inter, mem_set_of_eq, mem_Ioi, mem_preimage], }, end -variables {V P : Type*} [normed_group V] [normed_space ℝ V] [metric_space P] [normed_add_torsor V P] +variables {V P : Type*} [normed_add_comm_group V] [normed_space ℝ V] [metric_space P] + [normed_add_torsor V P] include V open affine_map diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index 859a7d258b3b7..269c09455175d 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -17,7 +17,7 @@ We also prove basic lemmas and provide convenience constructors. The choice of constructors is closely modelled on those for the `linear_isometry` and `affine_map` theories. Since many elementary properties don't require `∥x∥ = 0 → x = 0` we initially set up the theory for -`semi_normed_group` and specialize to `normed_group` only when needed. +`seminormed_add_comm_group` and specialize to `normed_add_comm_group` only when needed. ## Notation @@ -31,9 +31,9 @@ algebra-homomorphisms.) open function set variables (𝕜 : Type*) {V V₁ V₂ V₃ V₄ : Type*} {P₁ : Type*} (P P₂ : Type*) {P₃ P₄ : Type*} - [normed_field 𝕜] - [semi_normed_group V] [semi_normed_group V₁] [semi_normed_group V₂] [semi_normed_group V₃] - [semi_normed_group V₄] + [normed_field 𝕜] [seminormed_add_comm_group V] [seminormed_add_comm_group V₁] + [seminormed_add_comm_group V₂] [seminormed_add_comm_group V₃] + [seminormed_add_comm_group V₄] [normed_space 𝕜 V] [normed_space 𝕜 V₁] [normed_space 𝕜 V₂] [normed_space 𝕜 V₃] [normed_space 𝕜 V₄] [pseudo_metric_space P] [metric_space P₁] [pseudo_metric_space P₂] [pseudo_metric_space P₃] diff --git a/src/analysis/normed_space/ball_action.lean b/src/analysis/normed_space/ball_action.lean index b151fffb38294..be948c62c24df 100644 --- a/src/analysis/normed_space/ball_action.lean +++ b/src/analysis/normed_space/ball_action.lean @@ -17,7 +17,7 @@ multiplicative actions. -/ open metric set variables {𝕜 𝕜' E : Type*} [normed_field 𝕜] [normed_field 𝕜'] - [semi_normed_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] {r : ℝ} + [seminormed_add_comm_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] {r : ℝ} section closed_ball diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index 0a2b5bb45e8ef..9d551801527ca 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -18,8 +18,8 @@ open function metric set filter finset open_locale classical topological_space big_operators nnreal variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] (f : E →L[𝕜] F) include 𝕜 diff --git a/src/analysis/normed_space/banach_steinhaus.lean b/src/analysis/normed_space/banach_steinhaus.lean index ea1ac98c6d01a..be992f565df5c 100644 --- a/src/analysis/normed_space/banach_steinhaus.lean +++ b/src/analysis/normed_space/banach_steinhaus.lean @@ -23,7 +23,7 @@ open set variables {E F 𝕜 𝕜₂ : Type*} -[semi_normed_group E] [semi_normed_group F] +[seminormed_add_comm_group E] [seminormed_add_comm_group F] [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index c03a6a5809524..3119b32f9174d 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -21,7 +21,7 @@ noncomputable theory open filter metric function set open_locale topological_space big_operators nnreal ennreal uniformity pointwise -section semi_normed_group +section seminormed_add_comm_group section prio set_option extends_priority 920 @@ -32,14 +32,15 @@ set_option extends_priority 920 equality `∥c • x∥ = ∥c∥ ∥x∥`. We require only `∥c • x∥ ≤ ∥c∥ ∥x∥` in the definition, then prove `∥c • x∥ = ∥c∥ ∥x∥` in `norm_smul`. -Note that since this requires `semi_normed_group` and not `normed_group`, this typeclass can be -used for "semi normed spaces" too, just as `module` can be used for "semi modules". -/ -class normed_space (α : Type*) (β : Type*) [normed_field α] [semi_normed_group β] +Note that since this requires `seminormed_add_comm_group` and not `normed_add_comm_group`, this +typeclass can be used for "semi normed spaces" too, just as `module` can be used for +"semi modules". -/ +class normed_space (α : Type*) (β : Type*) [normed_field α] [seminormed_add_comm_group β] extends module α β := (norm_smul_le : ∀ (a:α) (b:β), ∥a • b∥ ≤ ∥a∥ * ∥b∥) end prio -variables [normed_field α] [semi_normed_group β] +variables [normed_field α] [seminormed_add_comm_group β] @[priority 100] -- see Note [lower instance priority] instance normed_space.has_bounded_smul [normed_space α β] : has_bounded_smul α β := @@ -87,8 +88,8 @@ 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] -variables {E : Type*} [semi_normed_group E] [normed_space α E] -variables {F : Type*} [semi_normed_group F] [normed_space α F] +variables {E : Type*} [seminormed_add_comm_group E] [normed_space α E] +variables {F : Type*} [seminormed_add_comm_group F] [normed_space α F] theorem eventually_nhds_norm_smul_sub_lt (c : α) (x : E) {ε : ℝ} (h : 0 < ε) : ∀ᶠ y in 𝓝 x, ∥c • (y - x)∥ < ε := @@ -162,7 +163,7 @@ This homeomorphism sends `x : E` to `(1 + ∥x∥)⁻¹ • x`. In many cases the actual implementation is not important, so we don't mark the projection lemmas `homeomorph_unit_ball_apply_coe` and `homeomorph_unit_ball_symm_apply` as `@[simp]`. -/ @[simps { attrs := [] }] -def homeomorph_unit_ball {E : Type*} [semi_normed_group E] [normed_space ℝ E] : +def homeomorph_unit_ball {E : Type*} [seminormed_add_comm_group E] [normed_space ℝ E] : E ≃ₜ ball (0 : E) 1 := { to_fun := λ x, ⟨(1 + ∥x∥)⁻¹ • x, begin have : ∥x∥ < |1 + ∥x∥| := (lt_one_add _).trans_le (le_abs_self _), @@ -191,17 +192,17 @@ open normed_field instance : normed_space α (ulift E) := { norm_smul_le := λ s x, (normed_space.norm_smul_le s x.down : _), - ..ulift.normed_group, + ..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], - ..prod.normed_group, + ..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, semi_normed_group (E i)] +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∥₊)) : ℝ) = @@ -210,7 +211,7 @@ instance pi.normed_space {E : ι → Type*} [fintype ι] [∀i, semi_normed_grou /-- 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*} [semi_normed_group E] [normed_space 𝕜 E] [module R E] + {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) } @@ -242,24 +243,25 @@ begin exact mul_le_mul_of_nonneg_right hn.1 (norm_nonneg _) } end -end semi_normed_group +end seminormed_add_comm_group -section normed_group +section normed_add_comm_group variables [normed_field α] -variables {E : Type*} [normed_group E] [normed_space α E] -variables {F : Type*} [normed_group F] [normed_space α F] +variables {E : Type*} [normed_add_comm_group E] [normed_space α E] +variables {F : Type*} [normed_add_comm_group F] [normed_space α F] open normed_field /-- While this may appear identical to `normed_space.to_module`, it contains an implicit argument -involving `normed_group.to_semi_normed_group` that typeclass inference has trouble inferring. +involving `normed_add_comm_group.to_seminormed_add_comm_group` that typeclass inference has trouble +inferring. Specifically, the following instance cannot be found without this `normed_space.to_module'`: ```lean example (𝕜 ι : Type*) (E : ι → Type*) - [normed_field 𝕜] [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] : + [normed_field 𝕜] [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] : Π i, module 𝕜 (E i) := by apply_instance ``` @@ -312,11 +314,11 @@ lemma rescale_to_shell {c : α} (hc : 1 < ∥c∥) {ε : ℝ} (εpos : 0 < ε) { ∃d:α, d ≠ 0 ∧ ∥d • x∥ < ε ∧ (ε/∥c∥ ≤ ∥d • x∥) ∧ (∥d∥⁻¹ ≤ ε⁻¹ * ∥c∥ * ∥x∥) := rescale_to_shell_semi_normed hc εpos (ne_of_lt (norm_pos_iff.2 hx)).symm -end normed_group +end normed_add_comm_group section nontrivially_normed_space -variables (𝕜 E : Type*) [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] +variables (𝕜 E : Type*) [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] [nontrivial E] include 𝕜 @@ -472,13 +474,13 @@ end normed_algebra section restrict_scalars variables (𝕜 : Type*) (𝕜' : Type*) [normed_field 𝕜] [normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] -(E : Type*) [semi_normed_group E] [normed_space 𝕜' E] +(E : Type*) [seminormed_add_comm_group E] [normed_space 𝕜' E] -instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : semi_normed_group E] : - semi_normed_group (restrict_scalars 𝕜 𝕜' E) := I +instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : seminormed_add_comm_group E] : + seminormed_add_comm_group (restrict_scalars 𝕜 𝕜' E) := I -instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : normed_group E] : - normed_group (restrict_scalars 𝕜 𝕜' E) := I +instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : normed_add_comm_group E] : + normed_add_comm_group (restrict_scalars 𝕜 𝕜' E) := I /-- If `E` is a normed space over `𝕜'` and `𝕜` is a normed algebra over `𝕜'`, then `restrict_scalars.module` is additionally a `normed_space`. -/ @@ -494,7 +496,7 @@ This is not an instance as it would be contrary to the purpose of `restrict_scal -- If you think you need this, consider instead reproducing `restrict_scalars.lsmul` -- appropriately modified here. def module.restrict_scalars.normed_space_orig {𝕜 : Type*} {𝕜' : Type*} {E : Type*} - [normed_field 𝕜'] [semi_normed_group E] [I : normed_space 𝕜' E] : + [normed_field 𝕜'] [seminormed_add_comm_group E] [I : normed_space 𝕜' E] : normed_space 𝕜' (restrict_scalars 𝕜 𝕜' E) := I /-- Warning: This declaration should be used judiciously. diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 505fd5d8af797..67d91463d48ba 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -54,15 +54,15 @@ open_locale classical big_operators topological_space open filter (tendsto) metric continuous_linear_map variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] - {G : Type*} [normed_group G] [normed_space 𝕜 G] + {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] /-- A function `f` satisfies `is_bounded_linear_map 𝕜 f` if it is linear and satisfies the inequality `∥f x∥ ≤ M * ∥x∥` for some positive constant `M`. -/ structure is_bounded_linear_map (𝕜 : Type*) [normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] (f : E → F) + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] + {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] (f : E → F) extends is_linear_map 𝕜 f : Prop := (bound : ∃ M, 0 < M ∧ ∀ x : E, ∥f x∥ ≤ M * ∥x∥) @@ -188,7 +188,7 @@ variables {ι : Type*} [decidable_eq ι] [fintype ι] /-- Taking the cartesian product of two continuous multilinear maps is a bounded linear operation. -/ lemma is_bounded_linear_map_prod_multilinear - {E : ι → Type*} [∀ i, normed_group (E i)] [∀ i, normed_space 𝕜 (E i)] : + {E : ι → Type*} [∀ i, normed_add_comm_group (E i)] [∀ i, normed_space 𝕜 (E i)] : is_bounded_linear_map 𝕜 (λ p : (continuous_multilinear_map 𝕜 E F) × (continuous_multilinear_map 𝕜 E G), p.1.prod p.2) := { map_add := λ p₁ p₂, by { ext1 m, refl }, @@ -234,16 +234,16 @@ namespace continuous_linear_map If `f` is a continuuous bilinear map, to use the corresponding rules for the second argument, use `(f _).map_add` and similar. - We have to assume that `F` and `G` are normed spaces in this section, to use - `continuous_linear_map.to_normed_group`, but we don't need to assume this for the first argument - of `f`. +We have to assume that `F` and `G` are normed spaces in this section, to use +`continuous_linear_map.to_normed_add_comm_group`, but we don't need to assume this for the first +argument of `f`. -/ variables {R : Type*} variables {𝕜₂ 𝕜' : Type*} [nontrivially_normed_field 𝕜'] [nontrivially_normed_field 𝕜₂] variables {M : Type*} [topological_space M] variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] -variables {G' : Type*} [normed_group G'] [normed_space 𝕜₂ G'] [normed_space 𝕜' G'] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜₂ G'] [normed_space 𝕜' G'] variables [smul_comm_class 𝕜₂ 𝕜' G'] section semiring @@ -404,7 +404,7 @@ lemma is_bounded_bilinear_map.is_bounded_linear_map_right end } lemma is_bounded_bilinear_map_smul {𝕜' : Type*} [normed_field 𝕜'] - [normed_algebra 𝕜 𝕜'] {E : Type*} [normed_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] + [normed_algebra 𝕜 𝕜'] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] [is_scalar_tower 𝕜 𝕜' E] : is_bounded_bilinear_map 𝕜 (λ (p : 𝕜' × E), p.1 • p.2) := (lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E).is_bounded_bilinear_map @@ -440,7 +440,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_group (E i)] [∀ i, normed_space 𝕜 (E i)] : +[decidable_eq ι] [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/complemented.lean b/src/analysis/normed_space/complemented.lean index 893dd4efe9ea2..e666e595040c4 100644 --- a/src/analysis/normed_space/complemented.lean +++ b/src/analysis/normed_space/complemented.lean @@ -20,8 +20,9 @@ is always a complemented subspace. complemented subspace, normed vector space -/ -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] +variables {𝕜 E F G : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] + [normed_space 𝕜 E] [normed_add_comm_group F] [normed_space 𝕜 F] [normed_add_comm_group G] + [normed_space 𝕜 G] noncomputable theory diff --git a/src/analysis/normed_space/completion.lean b/src/analysis/normed_space/completion.lean index a4fa6337279de..4c7d363a7bdc2 100644 --- a/src/analysis/normed_space/completion.lean +++ b/src/analysis/normed_space/completion.lean @@ -20,7 +20,7 @@ noncomputable theory namespace uniform_space namespace completion -variables (𝕜 E : Type*) [normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] +variables (𝕜 E : Type*) [normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] @[priority 100] instance normed_space.to_has_uniform_continuous_const_smul : @@ -51,8 +51,8 @@ to_complₗᵢ.to_continuous_linear_map @[simp] lemma coe_to_complL : ⇑(to_complL : E →L[𝕜] completion E) = coe := rfl -@[simp] lemma norm_to_complL {𝕜 E : Type*} [nontrivially_normed_field 𝕜] - [normed_group E] [normed_space 𝕜 E] [nontrivial E] : ∥(to_complL : E →L[𝕜] completion E)∥ = 1 := +@[simp] lemma norm_to_complL {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] + [normed_space 𝕜 E] [nontrivial E] : ∥(to_complL : E →L[𝕜] completion E)∥ = 1 := (to_complₗᵢ : E →ₗᵢ[𝕜] completion E).norm_to_continuous_linear_map end completion diff --git a/src/analysis/normed_space/conformal_linear_map.lean b/src/analysis/normed_space/conformal_linear_map.lean index bb6ca05abae67..9f66f89da2471 100644 --- a/src/analysis/normed_space/conformal_linear_map.lean +++ b/src/analysis/normed_space/conformal_linear_map.lean @@ -45,14 +45,14 @@ open function linear_isometry continuous_linear_map /-- A continuous linear map `f'` is said to be conformal if it's a nonzero multiple of a linear isometry. -/ def is_conformal_map {R : Type*} {X Y : Type*} [normed_field R] - [semi_normed_group X] [semi_normed_group Y] [normed_space R X] [normed_space R Y] + [seminormed_add_comm_group X] [seminormed_add_comm_group Y] [normed_space R X] [normed_space R Y] (f' : X →L[R] Y) := ∃ (c : R) (hc : c ≠ 0) (li : X →ₗᵢ[R] Y), f' = c • li.to_continuous_linear_map variables {R M N G M' : Type*} [normed_field R] - [semi_normed_group M] [semi_normed_group N] [semi_normed_group G] + [seminormed_add_comm_group M] [seminormed_add_comm_group N] [seminormed_add_comm_group G] [normed_space R M] [normed_space R N] [normed_space R G] - [normed_group M'] [normed_space R M'] + [normed_add_comm_group M'] [normed_space R M'] {f : M →L[R] N} {g : N →L[R] G} {c : R} lemma is_conformal_map_id : is_conformal_map (id R M) := diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 837958ee17c8c..330977e666af5 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -42,9 +42,9 @@ submultiplicative: for a composition of maps, we have only `∥f.comp g∥ ≤ namespace continuous_affine_map variables {𝕜 R V W W₂ P Q Q₂ : Type*} -variables [normed_group V] [metric_space P] [normed_add_torsor V P] -variables [normed_group W] [metric_space Q] [normed_add_torsor W Q] -variables [normed_group W₂] [metric_space Q₂] [normed_add_torsor W₂ Q₂] +variables [normed_add_comm_group V] [metric_space P] [normed_add_torsor V P] +variables [normed_add_comm_group W] [metric_space Q] [normed_add_torsor W Q] +variables [normed_add_comm_group W₂] [metric_space Q₂] [normed_add_torsor W₂ Q₂] variables [normed_field R] [normed_space R V] [normed_space R W] [normed_space R W₂] variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 V] [normed_space 𝕜 W] [normed_space 𝕜 W₂] @@ -158,8 +158,8 @@ calc ∥f∥ = (max ∥f 0∥ ∥f.cont_linear∥) : by rw norm_def ... = (max 0 ∥f.cont_linear∥) : by rw [h, norm_zero] ... = ∥f.cont_linear∥ : max_eq_right (norm_nonneg _) -noncomputable instance : normed_group (V →A[𝕜] W) := -normed_group.of_core _ +noncomputable instance : normed_add_comm_group (V →A[𝕜] W) := +normed_add_comm_group.of_core _ { norm_eq_zero_iff := λ f, begin rw norm_def, diff --git a/src/analysis/normed_space/dual.lean b/src/analysis/normed_space/dual.lean index a5afb7dbfd9b2..bdd00bb90571d 100644 --- a/src/analysis/normed_space/dual.lean +++ b/src/analysis/normed_space/dual.lean @@ -19,7 +19,7 @@ version `normed_space.inclusion_in_double_dual_li` of the map which is of type a isometric embedding, `E →ₗᵢ[𝕜] (dual 𝕜 (dual 𝕜 E))`. Since a lot of elementary properties don't require `eq_of_dist_eq_zero` we start setting up the -theory for `semi_normed_group` and we specialize to `normed_group` when needed. +theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_group` when needed. ## Main definitions @@ -41,18 +41,18 @@ namespace normed_space section general variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] -variables (E : Type*) [semi_normed_group E] [normed_space 𝕜 E] -variables (F : Type*) [normed_group F] [normed_space 𝕜 F] +variables (E : Type*) [seminormed_add_comm_group E] [normed_space 𝕜 E] +variables (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] /-- The topological dual of a seminormed space `E`. -/ -@[derive [inhabited, semi_normed_group, normed_space 𝕜]] def dual := E →L[𝕜] 𝕜 +@[derive [inhabited, seminormed_add_comm_group, normed_space 𝕜]] def dual := E →L[𝕜] 𝕜 instance : continuous_linear_map_class (dual 𝕜 E) 𝕜 E 𝕜 := continuous_linear_map.continuous_semilinear_map_class instance : has_coe_to_fun (dual 𝕜 E) (λ _, E → 𝕜) := continuous_linear_map.to_fun -instance : normed_group (dual 𝕜 F) := continuous_linear_map.to_normed_group +instance : normed_add_comm_group (dual 𝕜 F) := continuous_linear_map.to_normed_add_comm_group instance [finite_dimensional 𝕜 E] : finite_dimensional 𝕜 (dual 𝕜 E) := continuous_linear_map.finite_dimensional @@ -90,7 +90,7 @@ end general section bidual_isometry variables (𝕜 : Type v) [is_R_or_C 𝕜] - {E : Type u} [normed_group E] [normed_space 𝕜 E] + {E : Type u} [normed_add_comm_group E] [normed_space 𝕜 E] /-- If one controls the norm of every `f x`, then one controls the norm of `x`. Compare `continuous_linear_map.op_norm_le_bound`. -/ @@ -144,11 +144,11 @@ open metric set normed_space `polar 𝕜 s` is the subset of `dual 𝕜 E` consisting of those functionals which evaluate to something of norm at most one at all points `z ∈ s`. -/ def polar (𝕜 : Type*) [nontrivially_normed_field 𝕜] - {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] : set E → set (dual 𝕜 E) := + {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] : set E → set (dual 𝕜 E) := (dual_pairing 𝕜 E).flip.polar variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] -variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] +variables {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] lemma mem_polar_iff {x' : dual 𝕜 E} (s : set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ∥x' z∥ ≤ 1 := iff.rfl @@ -217,8 +217,8 @@ calc ∥x' x∥ ≤ ∥x'∥ * ∥x∥ : x'.le_op_norm x /-- The `polar` of closed ball in a normed space `E` is the closed ball of the dual with inverse radius. -/ -lemma polar_closed_ball - {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {r : ℝ} (hr : 0 < r) : +lemma polar_closed_ball {𝕜 E : Type*} [is_R_or_C 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] + {r : ℝ} (hr : 0 < r) : polar 𝕜 (closed_ball (0 : E) r) = closed_ball (0 : dual 𝕜 E) r⁻¹ := begin refine subset.antisymm _ (closed_ball_inv_subset_polar_closed_ball _), diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index 27a35bf2ec2e5..4af078e526972 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -180,7 +180,7 @@ lemma finite_dist_eq (x y : e.finite_subspace) : dist x y = (e (x - y)).to_real lemma finite_edist_eq (x y : e.finite_subspace) : edist x y = e (x - y) := rfl /-- Normed group instance on `e.finite_subspace`. -/ -instance : normed_group e.finite_subspace := +instance : normed_add_comm_group e.finite_subspace := { norm := λ x, (e x).to_real, dist_eq := λ x y, rfl, .. finite_subspace.metric_space e, .. submodule.add_comm_group _ } diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index fd9f11ecc7330..66bfa28a352c4 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -32,7 +32,7 @@ Alternate forms which operate on `[is_scalar_tower ℝ 𝕜 F]` instead are prov open is_R_or_C -variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [semi_normed_group F] [normed_space 𝕜 F] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [seminormed_add_comm_group F] [normed_space 𝕜 F] local notation `abs𝕜` := @is_R_or_C.abs 𝕜 _ /-- Extend `fr : F →ₗ[ℝ] ℝ` to `F →ₗ[𝕜] 𝕜` in a way that will also be continuous and have its norm diff --git a/src/analysis/normed_space/extr.lean b/src/analysis/normed_space/extr.lean index ea541566064e6..603a1020ad2b3 100644 --- a/src/analysis/normed_space/extr.lean +++ b/src/analysis/normed_space/extr.lean @@ -21,7 +21,7 @@ Then we specialize it to the case `y = f c` and to different special cases of `i local maximum, normed space -/ -variables {α X E : Type*} [semi_normed_group E] [normed_space ℝ E] [topological_space X] +variables {α X E : Type*} [seminormed_add_comm_group E] [normed_space ℝ E] [topological_space X] section diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index a3a8704e38761..5d6cd07bd40a2 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -54,8 +54,8 @@ open linear_map variables {R : Type*} [semiring R] -variables {F E₁ : Type*} [semi_normed_group F] - [normed_group E₁] [module R E₁] +variables {F E₁ : Type*} [seminormed_add_comm_group F] + [normed_add_comm_group E₁] [module R E₁] variables {R₁ : Type*} [field R₁] [module R₁ E₁] [module R₁ F] [finite_dimensional R₁ E₁] [finite_dimensional R₁ F] @@ -84,7 +84,7 @@ open affine_map variables {𝕜 : Type*} {V₁ V₂ : Type*} {P₁ P₂ : Type*} [normed_field 𝕜] - [normed_group V₁] [semi_normed_group V₂] + [normed_add_comm_group V₁] [seminormed_add_comm_group V₂] [normed_space 𝕜 V₁] [normed_space 𝕜 V₂] [metric_space P₁] [pseudo_metric_space P₂] [normed_add_torsor V₁ P₁] [normed_add_torsor V₂ P₂] @@ -111,8 +111,8 @@ end affine_isometry section complete_field variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] -{E : Type v} [normed_group E] [normed_space 𝕜 E] -{F : Type w} [normed_group F] [normed_space 𝕜 F] +{E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type w} [normed_add_comm_group F] [normed_space 𝕜 F] {F' : Type x} [add_comm_group F'] [module 𝕜 F'] [topological_space F'] [topological_add_group F'] [has_continuous_smul 𝕜 F'] [complete_space 𝕜] @@ -168,12 +168,12 @@ vector space `E'` can be extended to a Lipschitz map on the whole space `α`, wi constant `C * K` where `C` only depends on `E'`. We record a working value for this constant `C` as `lipschitz_extension_constant E'`. -/ @[irreducible] def lipschitz_extension_constant - (E' : Type*) [normed_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : ℝ≥0 := + (E' : Type*) [normed_add_comm_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : ℝ≥0 := let A := (basis.of_vector_space ℝ E').equiv_fun.to_continuous_linear_equiv in max (∥A.symm.to_continuous_linear_map∥₊ * ∥A.to_continuous_linear_map∥₊) 1 lemma lipschitz_extension_constant_pos - (E' : Type*) [normed_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : + (E' : Type*) [normed_add_comm_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : 0 < lipschitz_extension_constant E' := by { rw lipschitz_extension_constant, exact zero_lt_one.trans_le (le_max_right _ _) } @@ -182,7 +182,7 @@ vector space `E'` can be extended to a Lipschitz map on the whole space `α`, wi constant `lipschitz_extension_constant E' * K`. -/ theorem lipschitz_on_with.extend_finite_dimension {α : Type*} [pseudo_metric_space α] - {E' : Type*} [normed_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] + {E' : Type*} [normed_add_comm_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] {s : set α} {f : α → E'} {K : ℝ≥0} (hf : lipschitz_on_with K f s) : ∃ (g : α → E'), lipschitz_with (lipschitz_extension_constant E' * K) g ∧ eq_on f g s := begin @@ -593,7 +593,7 @@ end complete_field section proper_field variables (𝕜 : Type u) [nontrivially_normed_field 𝕜] -(E : Type v) [normed_group E] [normed_space 𝕜 E] [proper_space 𝕜] +(E : Type v) [normed_add_comm_group E] [normed_space 𝕜 E] [proper_space 𝕜] /-- Any finite-dimensional vector space over a proper field is proper. We do not register this as an instance to avoid an instance loop when trying to prove the @@ -610,15 +610,15 @@ end proper_field /- Over the real numbers, we can register the previous statement as an instance as it will not cause problems in instance resolution since the properness of `ℝ` is already known. -/ @[priority 900] -instance finite_dimensional.proper_real - (E : Type u) [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] : proper_space E := +instance finite_dimensional.proper_real (E : Type u) [normed_add_comm_group E] [normed_space ℝ E] + [finite_dimensional ℝ E] : proper_space E := finite_dimensional.proper ℝ E /-- If `E` is a finite dimensional normed real vector space, `x : E`, and `s` is a neighborhood of `x` that is not equal to the whole space, then there exists a point `y ∈ frontier s` at distance `metric.inf_dist x sᶜ` from `x`. See also `is_compact.exists_mem_frontier_inf_dist_compl_eq_dist`. -/ -lemma exists_mem_frontier_inf_dist_compl_eq_dist {E : Type*} [normed_group E] +lemma exists_mem_frontier_inf_dist_compl_eq_dist {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {x : E} {s : set E} (hx : x ∈ s) (hs : s ≠ univ) : ∃ y ∈ frontier s, metric.inf_dist x sᶜ = dist x y := begin @@ -632,7 +632,7 @@ end /-- If `K` is a compact set in a nontrivial real normed space and `x ∈ K`, then there exists a point `y` of the boundary of `K` at distance `metric.inf_dist x Kᶜ` from `x`. See also `exists_mem_frontier_inf_dist_compl_eq_dist`. -/ -lemma is_compact.exists_mem_frontier_inf_dist_compl_eq_dist {E : Type*} [normed_group E] +lemma is_compact.exists_mem_frontier_inf_dist_compl_eq_dist {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [nontrivial E] {x : E} {K : set E} (hK : is_compact K) (hx : x ∈ K) : ∃ y ∈ frontier K, metric.inf_dist x Kᶜ = dist x y := begin @@ -652,8 +652,8 @@ end /-- In a finite dimensional vector space over `ℝ`, the series `∑ x, ∥f x∥` is unconditionally summable if and only if the series `∑ x, f x` is unconditionally summable. One implication holds in any complete normed space, while the other holds only in finite dimensional spaces. -/ -lemma summable_norm_iff {α E : Type*} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] - {f : α → E} : summable (λ x, ∥f x∥) ↔ summable f := +lemma summable_norm_iff {α E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [finite_dimensional ℝ E] {f : α → E} : summable (λ x, ∥f x∥) ↔ summable f := begin refine ⟨summable_of_summable_norm, λ hf, _⟩, -- First we use a finite basis to reduce the problem to the case `E = fin N → ℝ` @@ -675,32 +675,32 @@ begin { exact finset.sum_nonneg (λ _ _, norm_nonneg _) } end -lemma summable_of_is_O' {ι E F : Type*} [normed_group E] [complete_space E] [normed_group F] - [normed_space ℝ F] [finite_dimensional ℝ F] {f : ι → E} {g : ι → F} +lemma summable_of_is_O' {ι E F : Type*} [normed_add_comm_group E] [complete_space E] + [normed_add_comm_group F] [normed_space ℝ F] [finite_dimensional ℝ F] {f : ι → E} {g : ι → F} (hg : summable g) (h : f =O[cofinite] g) : summable f := summable_of_is_O (summable_norm_iff.mpr hg) h.norm_right -lemma summable_of_is_O_nat' {E F : Type*} [normed_group E] [complete_space E] [normed_group F] - [normed_space ℝ F] [finite_dimensional ℝ F] {f : ℕ → E} {g : ℕ → F} +lemma summable_of_is_O_nat' {E F : Type*} [normed_add_comm_group E] [complete_space E] + [normed_add_comm_group F] [normed_space ℝ F] [finite_dimensional ℝ F] {f : ℕ → E} {g : ℕ → F} (hg : summable g) (h : f =O[at_top] g) : summable f := summable_of_is_O_nat (summable_norm_iff.mpr hg) h.norm_right -lemma summable_of_is_equivalent {ι E : Type*} [normed_group E] [normed_space ℝ E] +lemma summable_of_is_equivalent {ι E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {f : ι → E} {g : ι → E} (hg : summable g) (h : f ~[cofinite] g) : summable f := hg.trans_sub (summable_of_is_O' hg h.is_o.is_O) -lemma summable_of_is_equivalent_nat {E : Type*} [normed_group E] [normed_space ℝ E] +lemma summable_of_is_equivalent_nat {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {f : ℕ → E} {g : ℕ → E} (hg : summable g) (h : f ~[at_top] g) : summable f := hg.trans_sub (summable_of_is_O_nat' hg h.is_o.is_O) -lemma is_equivalent.summable_iff {ι E : Type*} [normed_group E] [normed_space ℝ E] +lemma is_equivalent.summable_iff {ι E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {f : ι → E} {g : ι → E} (h : f ~[cofinite] g) : summable f ↔ summable g := ⟨λ hf, summable_of_is_equivalent hf h.symm, λ hg, summable_of_is_equivalent hg h⟩ -lemma is_equivalent.summable_iff_nat {E : Type*} [normed_group E] [normed_space ℝ E] +lemma is_equivalent.summable_iff_nat {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {f : ℕ → E} {g : ℕ → E} (h : f ~[at_top] g) : summable f ↔ summable g := ⟨λ hf, summable_of_is_equivalent_nat hf h.symm, λ hg, summable_of_is_equivalent_nat hg h⟩ diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index accee48ac378f..103c5f6d68263 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -31,7 +31,7 @@ of `𝕜`). universes u v namespace real -variables {E : Type*} [semi_normed_group E] [normed_space ℝ E] +variables {E : Type*} [seminormed_add_comm_group E] [normed_space ℝ E] /-- Hahn-Banach theorem for continuous linear functions over `ℝ`. -/ theorem exists_extension_norm_eq (p : subspace ℝ E) (f : p →L[ℝ] ℝ) : @@ -58,7 +58,7 @@ end real section is_R_or_C open is_R_or_C -variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [semi_normed_group F] [normed_space 𝕜 F] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [seminormed_add_comm_group F] [normed_space 𝕜 F] /-- Hahn-Banach theorem for continuous linear functions over `𝕜` satisyfing `is_R_or_C 𝕜`. -/ theorem exists_extension_norm_eq (p : subspace 𝕜 F) (f : p →L[𝕜] 𝕜) : @@ -103,7 +103,7 @@ end is_R_or_C section dual_vector variables (𝕜 : Type v) [is_R_or_C 𝕜] -variables {E : Type u} [normed_group E] [normed_space 𝕜 E] +variables {E : Type u} [normed_add_comm_group E] [normed_space 𝕜 E] open continuous_linear_equiv submodule open_locale classical diff --git a/src/analysis/normed_space/hahn_banach/separation.lean b/src/analysis/normed_space/hahn_banach/separation.lean index c1b099a188570..0c4d33ab05fe1 100644 --- a/src/analysis/normed_space/hahn_banach/separation.lean +++ b/src/analysis/normed_space/hahn_banach/separation.lean @@ -38,7 +38,7 @@ variables {𝕜 E : Type*} /-- Given a set `s` which is a convex neighbourhood of `0` and a point `x₀` outside of it, there is a continuous linear functional `f` separating `x₀` and `s`, in the sense that it sends `x₀` to 1 and all of `s` to values strictly below `1`. -/ -lemma separate_convex_open_set [semi_normed_group E] [normed_space ℝ E] {s : set E} +lemma separate_convex_open_set [seminormed_add_comm_group E] [normed_space ℝ E] {s : set E} (hs₀ : (0 : E) ∈ s) (hs₁ : convex ℝ s) (hs₂ : is_open s) {x₀ : E} (hx₀ : x₀ ∉ s) : ∃ f : E →L[ℝ] ℝ, f x₀ = 1 ∧ ∀ x ∈ s, f x < 1 := begin @@ -73,7 +73,7 @@ begin apply_instance } end -variables [normed_group E] [normed_space ℝ E] {s t : set E} {x y : E} +variables [normed_add_comm_group E] [normed_space ℝ E] {s t : set E} {x y : E} /-- A version of the **Hahn-Banach theorem**: given disjoint convex sets `s`, `t` where `s` is open, there is a continuous linear functional which separates them. -/ diff --git a/src/analysis/normed_space/indicator_function.lean b/src/analysis/normed_space/indicator_function.lean index 71019a37d71c5..a27af2a3975c7 100644 --- a/src/analysis/normed_space/indicator_function.lean +++ b/src/analysis/normed_space/indicator_function.lean @@ -15,7 +15,7 @@ This file contains a few simple lemmas about `set.indicator` and `norm`. indicator, norm -/ -variables {α E : Type*} [semi_normed_group E] {s t : set α} (f : α → E) (a : α) +variables {α E : Type*} [seminormed_add_comm_group E] {s t : set α} (f : α → E) (a : α) open set diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index ee947e962d0be..bad1305f97dd8 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -29,10 +29,10 @@ This file exists mainly to avoid importing `is_R_or_C` in the main normed space open metric @[simp, is_R_or_C_simps] lemma is_R_or_C.norm_coe_norm {𝕜 : Type*} [is_R_or_C 𝕜] - {E : Type*} [normed_group E] {z : E} : ∥(∥z∥ : 𝕜)∥ = ∥z∥ := + {E : Type*} [normed_add_comm_group E] {z : E} : ∥(∥z∥ : 𝕜)∥ = ∥z∥ := by { unfold_coes, simp only [norm_algebra_map', ring_hom.to_fun_eq_coe, norm_norm], } -variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] /-- Lemma to normalize a vector in a normed space `E` over either `ℂ` or `ℝ` to unit length. -/ @[simp] lemma norm_smul_inv_norm {x : E} (hx : x ≠ 0) : ∥(∥x∥⁻¹ : 𝕜) • x∥ = 1 := diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index 58eb3fdf14a45..737aa3b69e7c1 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -42,7 +42,7 @@ 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_group α, lattice α := + extends normed_add_comm_group α, lattice α := (add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) (solid : ∀ a b : α, |a| ≤ |b| → ∥a∥ ≤ ∥b∥) @@ -63,7 +63,7 @@ instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} Let `α` be a normed group with a partial order. Then the order dual is also a normed group. -/ @[priority 100] -- see Note [lower instance priority] -instance {α : Type*} : Π [normed_group α], normed_group αᵒᵈ := id +instance {α : Type*} : Π [normed_add_comm_group α], normed_add_comm_group αᵒᵈ := id variables {α : Type*} [normed_lattice_add_comm_group α] open lattice_ordered_comm_group diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index fdac94b0e337d..35497b695a19a 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -19,7 +19,7 @@ the star-linear versions. We also prove some trivial lemmas and provide convenience constructors. Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the -theory for `semi_normed_group` and we specialize to `normed_group` when needed. +theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_group` when needed. -/ open function set @@ -37,13 +37,13 @@ variables {R R₂ R₃ R₄ E E₂ E₃ E₄ F : Type*} [semiring R] [semiring R [ring_hom_comp_triple σ₂₃ σ₃₄ σ₂₄] [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] [ring_hom_comp_triple σ₃₂ σ₂₁ σ₃₁] [ring_hom_comp_triple σ₄₂ σ₂₁ σ₄₁] [ring_hom_comp_triple σ₄₃ σ₃₂ σ₄₂] [ring_hom_comp_triple σ₄₃ σ₃₁ σ₄₁] - [semi_normed_group E] [semi_normed_group E₂] [semi_normed_group E₃] [semi_normed_group E₄] - [module R E] [module R₂ E₂] [module R₃ E₃] [module R₄ E₄] - [normed_group F] [module R F] + [seminormed_add_comm_group E] [seminormed_add_comm_group E₂] [seminormed_add_comm_group E₃] + [seminormed_add_comm_group E₄] [module R E] [module R₂ E₂] [module R₃ E₃] [module R₄ E₄] + [normed_add_comm_group F] [module R F] /-- A `σ₁₂`-semilinear isometric embedding of a normed `R`-module into an `R₂`-module. -/ -structure linear_isometry (σ₁₂ : R →+* R₂) (E E₂ : Type*) [semi_normed_group E] - [semi_normed_group E₂] [module R E] [module R₂ E₂] extends E →ₛₗ[σ₁₂] E₂ := +structure linear_isometry (σ₁₂ : R →+* R₂) (E E₂ : Type*) [seminormed_add_comm_group E] + [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends E →ₛₗ[σ₁₂] E₂ := (norm_map' : ∀ x, ∥to_linear_map x∥ = ∥x∥) notation E ` →ₛₗᵢ[`:25 σ₁₂:25 `] `:0 E₂:0 := linear_isometry σ₁₂ E E₂ @@ -81,8 +81,8 @@ fun_like.coe_injective /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ -def simps.apply (σ₁₂ : R →+* R₂) (E E₂ : Type*) [semi_normed_group E] - [semi_normed_group E₂] [module R E] [module R₂ E₂] (h : E →ₛₗᵢ[σ₁₂] E₂) : E → E₂ := h +def simps.apply (σ₁₂ : R →+* R₂) (E E₂ : Type*) [seminormed_add_comm_group E] + [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] (h : E →ₛₗᵢ[σ₁₂] E₂) : E → E₂ := h initialize_simps_projections linear_isometry (to_linear_map_to_fun → apply) @@ -268,8 +268,8 @@ end submodule /-- A semilinear isometric equivalence between two normed vector spaces. -/ structure linear_isometry_equiv (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] - [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [semi_normed_group E] [semi_normed_group E₂] - [module R E] [module R₂ E₂] extends E ≃ₛₗ[σ₁₂] E₂ := + [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [seminormed_add_comm_group E] + [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends E ≃ₛₗ[σ₁₂] E₂ := (norm_map' : ∀ x, ∥to_linear_equiv x∥ = ∥x∥) notation E ` ≃ₛₗᵢ[`:25 σ₁₂:25 `] `:0 E₂:0 := linear_isometry_equiv σ₁₂ E E₂ @@ -422,12 +422,13 @@ def symm : E₂ ≃ₛₗᵢ[σ₂₁] E := /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ def simps.apply (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] - [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [semi_normed_group E] [semi_normed_group E₂] - [module R E] [module R₂ E₂] (h : E ≃ₛₗᵢ[σ₁₂] E₂) : E → E₂ := h + [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [seminormed_add_comm_group E] + [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] (h : E ≃ₛₗᵢ[σ₁₂] E₂) : E → E₂ := h /-- See Note [custom simps projection] -/ def simps.symm_apply (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] - [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [semi_normed_group E] [semi_normed_group E₂] + [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [seminormed_add_comm_group E] + [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] (h : E ≃ₛₗᵢ[σ₁₂] E₂) : E₂ → E := h.symm initialize_simps_projections linear_isometry_equiv diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 0c9e4bb7875ad..5a5ca99356ccc 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -28,7 +28,7 @@ The space `lp E p` is the subtype of elements of `Π i : α, E i` which satisfy if `p = 0`, `summable (λ a, ∥f a∥^p)` if `0 < p < ∞`, and `bdd_above (norm '' (set.range f))` if `p = ∞`. * `lp E p` : elements of `Π i : α, E i` such that `mem_ℓp f p`. Defined as an `add_subgroup` of - a type synonym `pre_lp` for `Π i : α, E i`, and equipped with a `normed_group` structure. + a type synonym `pre_lp` for `Π i : α, E i`, and equipped with a `normed_add_comm_group` structure. Under appropriate conditions, this is also equipped with the instances `lp.normed_space`, `lp.complete_space`. For `p=∞`, there is also `lp.infty_normed_ring`, `lp.infty_normed_algebra`. @@ -61,7 +61,7 @@ say that `∥-f∥ = ∥f∥`, instead of the non-working `f.norm_neg`. noncomputable theory open_locale nnreal ennreal big_operators -variables {α : Type*} {E : α → Type*} {p q : ℝ≥0∞} [Π i, normed_group (E i)] +variables {α : Type*} {E : α → Type*} {p q : ℝ≥0∞} [Π i, normed_add_comm_group (E i)] /-! ### `mem_ℓp` predicate @@ -293,12 +293,12 @@ We choose to deal with this issue by making a type synonym for `Π i, E i` rathe subgroup itself, because this allows all the spaces `lp E p` (for varying `p`) to be subgroups of the same ambient group, which permits lemma statements like `lp.monotone` (below). -/ @[derive add_comm_group, nolint unused_arguments] -def pre_lp (E : α → Type*) [Π i, normed_group (E i)] : Type* := Π i, E i +def pre_lp (E : α → Type*) [Π i, normed_add_comm_group (E i)] : Type* := Π i, E i instance pre_lp.unique [is_empty α] : unique (pre_lp E) := pi.unique_of_is_empty E /-- lp space -/ -def lp (E : α → Type*) [Π i, normed_group (E i)] +def lp (E : α → Type*) [Π i, normed_add_comm_group (E i)] (p : ℝ≥0∞) : add_subgroup (pre_lp E) := { carrier := {f | mem_ℓp f p}, zero_mem' := zero_mem_ℓp, @@ -456,8 +456,8 @@ begin simpa using lp.has_sum_norm hp f } end -instance [hp : fact (1 ≤ p)] : normed_group (lp E p) := -normed_group.of_core _ +instance [hp : fact (1 ≤ p)] : normed_add_comm_group (lp E p) := +normed_add_comm_group.of_core _ { norm_eq_zero_iff := norm_eq_zero_iff, triangle := λ f g, begin unfreezingI { rcases p.dichotomy with rfl | hp' }, @@ -661,7 +661,7 @@ instance : non_unital_normed_ring (lp B ∞) := ... ≤ ∥f∥ * ∥g∥ : mul_le_mul (lp.norm_apply_le_norm ennreal.top_ne_zero f i) (lp.norm_apply_le_norm ennreal.top_ne_zero g i) (norm_nonneg _) (norm_nonneg _)), - .. lp.normed_group } + .. lp.normed_add_comm_group } -- we also want a `non_unital_normed_comm_ring` instance, but this has to wait for #13719 @@ -905,7 +905,8 @@ begin have hp : p ≠ 0 := (ennreal.zero_lt_one.trans_le _i.elim).ne', rw uniform_continuous_pi, intros i, - rw normed_group.uniformity_basis_dist.uniform_continuous_iff normed_group.uniformity_basis_dist, + rw normed_add_comm_group.uniformity_basis_dist.uniform_continuous_iff + normed_add_comm_group.uniformity_basis_dist, intros ε hε, refine ⟨ε, hε, _⟩, rintros f g (hfg : ∥f - g∥ < ε), @@ -991,7 +992,7 @@ begin rw metric.nhds_basis_closed_ball.tendsto_right_iff, intros ε hε, have hε' : {p : (lp E p) × (lp E p) | ∥p.1 - p.2∥ < ε} ∈ 𝓤 (lp E p), - { exact normed_group.uniformity_basis_dist.mem_of_mem hε }, + { exact normed_add_comm_group.uniformity_basis_dist.mem_of_mem hε }, refine (hF.eventually_eventually hε').mono _, rintros n (hn : ∀ᶠ l in at_top, ∥(λ f, F n - f) (F l)∥ < ε), refine norm_le_of_tendsto (hn.mono (λ k hk, hk.le)) _, diff --git a/src/analysis/normed_space/matrix_exponential.lean b/src/analysis/normed_space/matrix_exponential.lean index 9426abbcf0608..1a1aea1e5afa5 100644 --- a/src/analysis/normed_space/matrix_exponential.lean +++ b/src/analysis/normed_space/matrix_exponential.lean @@ -23,10 +23,11 @@ without issue, so are not repeated here. The topological results specific to mat * `matrix.exp_block_diagonal'` Lemmas like `exp_add_of_commute` require a canonical norm on the type; while there are multiple -sensible choices for the norm of a `matrix` (`matrix.normed_group`, `matrix.frobenius_normed_group`, -`matrix.linfty_op_normed_group`), none of them are canonical. In an application where a particular -norm is chosen using `local attribute [instance]`, then the usual lemmas about `exp` are fine. When -choosing a norm is undesirable, the results in this file can be used. +sensible choices for the norm of a `matrix` (`matrix.normed_add_comm_group`, +`matrix.frobenius_normed_add_comm_group`, `matrix.linfty_op_normed_add_comm_group`), none of them +are canonical. In an application where a particular norm is chosen using +`local attribute [instance]`, then the usual lemmas about `exp` are fine. When choosing a norm is +undesirable, the results in this file can be used. In this file, we copy across the lemmas about `exp`, but hide the requirement for a norm inside the proof. diff --git a/src/analysis/normed_space/mazur_ulam.lean b/src/analysis/normed_space/mazur_ulam.lean index 862b9eb614518..2906940622db4 100644 --- a/src/analysis/normed_space/mazur_ulam.lean +++ b/src/analysis/normed_space/mazur_ulam.lean @@ -27,9 +27,9 @@ The formalization is based on [Jussi Väisälä, *A Proof of the Mazur-Ulam Theo isometry, affine map, linear map -/ -variables - {E PE : Type*} [normed_group E] [normed_space ℝ E] [metric_space PE] [normed_add_torsor E PE] - {F PF : Type*} [normed_group F] [normed_space ℝ F] [metric_space PF] [normed_add_torsor F PF] +variables {E PE F PF : Type*} [normed_add_comm_group E] [normed_space ℝ E] [metric_space PE] + [normed_add_torsor E PE] [normed_add_comm_group F] [normed_space ℝ F] [metric_space PF] + [normed_add_torsor F PF] open set affine_map affine_isometry_equiv diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 09bbbddfcf5c2..725c85f90e5ee 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -58,7 +58,7 @@ open_locale classical big_operators nnreal open finset metric local attribute [instance, priority 1001] -add_comm_group.to_add_comm_monoid normed_group.to_add_comm_group normed_space.to_module' +add_comm_group.to_add_comm_monoid normed_add_comm_group.to_add_comm_group normed_space.to_module' -- hack to speed up simp when dealing with complicated types local attribute [-instance] unique.subsingleton pi.subsingleton @@ -81,11 +81,11 @@ 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 𝕜] - [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] - [Π i, normed_group (E₁ i)] [Π i, normed_space 𝕜 (E₁ i)] - [Π i, normed_group (E' i)] [Π i, normed_space 𝕜 (E' i)] - [Π i, normed_group (Ei i)] [Π i, normed_space 𝕜 (Ei i)] - [normed_group G] [normed_space 𝕜 G] [normed_group G'] [normed_space 𝕜 G'] + [Π 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)] + [Π i, normed_add_comm_group (Ei i)] [Π i, normed_space 𝕜 (Ei i)] + [normed_add_comm_group G] [normed_space 𝕜 G] [normed_add_comm_group G'] [normed_space 𝕜 G'] /-! ### Continuity properties of multilinear maps @@ -123,7 +123,7 @@ begin obtain rfl : m = 0, from funext (is_empty.elim ‹_›), simp [univ_eq_empty, zero_le_one] }, obtain ⟨ε : ℝ, ε0 : 0 < ε, hε : ∀ m : Π i, E i, ∥m - 0∥ < ε → ∥f m - f 0∥ < 1⟩ := - normed_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one, + normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one, simp only [sub_zero, f.map_zero] at hε, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, have : 0 < (∥c∥ / ε) ^ fintype.card ι, from pow_pos (div_pos (zero_lt_one.trans hc) ε0) _, @@ -375,13 +375,14 @@ lemma op_norm_neg : ∥-f∥ = ∥f∥ := by { rw norm_def, apply congr_arg, ext /-- Continuous multilinear maps themselves form a normed space with respect to the operator norm. -/ -instance normed_group : normed_group (continuous_multilinear_map 𝕜 E G) := -normed_group.of_core _ ⟨op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ +instance normed_add_comm_group : normed_add_comm_group (continuous_multilinear_map 𝕜 E G) := +normed_add_comm_group.of_core _ ⟨op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ -/-- An alias of `continuous_multilinear_map.normed_group` with non-dependent types to help typeclass -search. -/ -instance normed_group' : normed_group (continuous_multilinear_map 𝕜 (λ i : ι, G) G') := -continuous_multilinear_map.normed_group +/-- An alias of `continuous_multilinear_map.normed_add_comm_group` with non-dependent types to help +typeclass search. -/ +instance normed_add_comm_group' : + normed_add_comm_group (continuous_multilinear_map 𝕜 (λ i : ι, G) G') := +continuous_multilinear_map.normed_add_comm_group instance normed_space : normed_space 𝕜' (continuous_multilinear_map 𝕜 E G) := ⟨λ c f, f.op_norm_smul_le c⟩ @@ -399,7 +400,7 @@ theorem le_op_norm_mul_pow_card_of_le {b : ℝ} (hm : ∀ i, ∥m i∥ ≤ b) : ∥f m∥ ≤ ∥f∥ * b ^ fintype.card ι := by simpa only [prod_const] using f.le_op_norm_mul_prod_of_le m hm -theorem le_op_norm_mul_pow_of_le {Ei : fin n → Type*} [Π i, normed_group (Ei i)] +theorem le_op_norm_mul_pow_of_le {Ei : fin n → Type*} [Π i, normed_add_comm_group (Ei i)] [Π i, normed_space 𝕜 (Ei i)] (f : continuous_multilinear_map 𝕜 Ei G) (m : Π i, Ei i) {b : ℝ} (hm : ∥m∥ ≤ b) : ∥f m∥ ≤ ∥f∥ * b ^ n := @@ -425,7 +426,7 @@ le_antisymm (f.op_norm_le_bound (norm_nonneg _) $ λ m, (le_max_left _ _).trans ((f.prod g).le_op_norm _)) (g.op_norm_le_bound (norm_nonneg _) $ λ m, (le_max_right _ _).trans ((f.prod g).le_op_norm _)) -lemma norm_pi {ι' : Type v'} [fintype ι'] {E' : ι' → Type wE'} [Π i', normed_group (E' i')] +lemma norm_pi {ι' : Type v'} [fintype ι'] {E' : ι' → Type wE'} [Π i', normed_add_comm_group (E' i')] [Π i', normed_space 𝕜 (E' i')] (f : Π i', continuous_multilinear_map 𝕜 E (E' i')) : ∥pi f∥ = ∥f∥ := begin @@ -459,7 +460,7 @@ def prodL : norm_map' := λ f, op_norm_prod f.1 f.2 } /-- `continuous_multilinear_map.pi` as a `linear_isometry_equiv`. -/ -def piₗᵢ {ι' : Type v'} [fintype ι'] {E' : ι' → Type wE'} [Π i', normed_group (E' i')] +def piₗᵢ {ι' : Type v'} [fintype ι'] {E' : ι' → Type wE'} [Π i', normed_add_comm_group (E' i')] [Π i', normed_space 𝕜 (E' i')] : @linear_isometry_equiv 𝕜 𝕜 _ _ (ring_hom.id 𝕜) _ _ _ (Π i', continuous_multilinear_map 𝕜 E (E' i')) (continuous_multilinear_map 𝕜 E (Π i, E' i)) _ _ diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index a1a2339a4ec29..06feeeb22c30c 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -15,7 +15,7 @@ Define the operator norm on the space of continuous (semi)linear maps between no prove its basic properties. In particular, show that this space is itself a normed space. Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the -theory for `semi_normed_group` and we specialize to `normed_group` at the end. +theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_group` at the end. Note that most of statements that apply to semilinear maps only hold when the ring homomorphism is isometric, as expressed by the typeclass `[ring_hom_isometric σ]`. @@ -30,8 +30,8 @@ variables {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} section semi_normed -variables [semi_normed_group E] [semi_normed_group Eₗ] [semi_normed_group F] [semi_normed_group Fₗ] -variables [semi_normed_group G] [semi_normed_group Gₗ] +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 @@ -110,7 +110,7 @@ lemma norm_image_of_norm_zero [semilinear_map_class 𝓕 σ₁₂ E F] (f : 𝓕 (hf : continuous f) {x : E} (hx : ∥x∥ = 0) : ∥f x∥ = 0 := begin refine le_antisymm (le_of_forall_pos_le_add (λ ε hε, _)) (norm_nonneg (f x)), - rcases normed_group.tendsto_nhds_nhds.1 (hf.tendsto 0) ε hε with ⟨δ, δ_pos, hδ⟩, + rcases normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) ε hε with ⟨δ, δ_pos, hδ⟩, replace hδ := hδ x, rw [sub_zero, hx] at hδ, replace hδ := le_of_lt (hδ δ_pos), @@ -140,7 +140,7 @@ controlled norm. The norm control for the original element follows by rescaling. lemma semilinear_map_class.bound_of_continuous [semilinear_map_class 𝓕 σ₁₂ E F] (f : 𝓕) (hf : continuous f) : ∃ C, 0 < C ∧ (∀ x : E, ∥f x∥ ≤ C * ∥x∥) := begin - rcases normed_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one with ⟨ε, ε_pos, hε⟩, + rcases normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one with ⟨ε, ε_pos, hε⟩, simp only [sub_zero, map_zero] at hε, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, have : 0 < ∥c∥ / ε, from div_pos (zero_lt_one.trans hc) ε_pos, @@ -382,8 +382,8 @@ lemma op_norm_smul_le {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' F /-- Continuous linear maps themselves form a seminormed space with respect to the operator norm. -/ -instance to_semi_normed_group : semi_normed_group (E →SL[σ₁₂] F) := -semi_normed_group.of_core _ ⟨op_norm_zero, λ x y, op_norm_add_le x y, op_norm_neg⟩ +instance to_seminormed_add_comm_group : seminormed_add_comm_group (E →SL[σ₁₂] F) := +seminormed_add_comm_group.of_core _ ⟨op_norm_zero, λ x y, op_norm_add_le x y, op_norm_neg⟩ lemma nnnorm_def (f : E →SL[σ₁₂] F) : ∥f∥₊ = Inf {c | ∀ x, ∥f x∥₊ ≤ c * ∥x∥₊} := begin @@ -436,7 +436,7 @@ omit σ₁₃ /-- Continuous linear maps form a seminormed ring with respect to the operator norm. -/ instance to_semi_normed_ring : semi_normed_ring (E →L[𝕜] E) := { norm_mul := λ f g, op_norm_comp_le f g, - .. continuous_linear_map.to_semi_normed_group, .. continuous_linear_map.ring } + .. continuous_linear_map.to_seminormed_add_comm_group, .. continuous_linear_map.ring } /-- For a normed space `E`, continuous linear endomorphisms form a normed algebra with respect to the operator norm. -/ @@ -753,10 +753,10 @@ def precompL (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : (Eₗ →L[𝕜] E) →L[ section prod universes u₁ u₂ u₃ u₄ -variables (M₁ : Type u₁) [semi_normed_group M₁] [normed_space 𝕜 M₁] - (M₂ : Type u₂) [semi_normed_group M₂] [normed_space 𝕜 M₂] - (M₃ : Type u₃) [semi_normed_group M₃] [normed_space 𝕜 M₃] - (M₄ : Type u₄) [semi_normed_group M₄] [normed_space 𝕜 M₄] +variables (M₁ : Type u₁) [seminormed_add_comm_group M₁] [normed_space 𝕜 M₁] + (M₂ : Type u₂) [seminormed_add_comm_group M₂] [normed_space 𝕜 M₂] + (M₃ : Type u₃) [seminormed_add_comm_group M₃] [normed_space 𝕜 M₃] + (M₄ : Type u₄) [seminormed_add_comm_group M₄] [normed_space 𝕜 M₄] variables {Eₗ} (𝕜) /-- `continuous_linear_map.prod_map` as a continuous linear map. -/ @@ -1117,7 +1117,7 @@ def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ₁₂] omit σ₂₁ namespace continuous_linear_map -variables {E' F' : Type*} [semi_normed_group E'] [semi_normed_group F'] +variables {E' F' : Type*} [seminormed_add_comm_group E'] [seminormed_add_comm_group F'] variables {𝕜₁' : Type*} {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₁'] [nontrivially_normed_field 𝕜₂'] [normed_space 𝕜₁' E'] [normed_space 𝕜₂' F'] @@ -1159,7 +1159,8 @@ end semi_normed section normed -variables [normed_group E] [normed_group F] [normed_group G] [normed_group Fₗ] +variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_add_comm_group G] + [normed_add_comm_group Fₗ] open metric continuous_linear_map @@ -1227,12 +1228,12 @@ instance norm_one_class [nontrivial E] : norm_one_class (E →L[𝕜] E) := ⟨n /-- Continuous linear maps themselves form a normed space with respect to the operator norm. -/ -instance to_normed_group [ring_hom_isometric σ₁₂] : normed_group (E →SL[σ₁₂] F) := -normed_group.of_core _ ⟨λ f, op_norm_zero_iff f, op_norm_add_le, op_norm_neg⟩ +instance to_normed_add_comm_group [ring_hom_isometric σ₁₂] : normed_add_comm_group (E →SL[σ₁₂] F) := +normed_add_comm_group.of_core _ ⟨λ f, op_norm_zero_iff f, op_norm_add_le, op_norm_neg⟩ /-- Continuous linear maps form a normed ring with respect to the operator norm. -/ instance to_normed_ring : normed_ring (E →L[𝕜] E) := -{ .. continuous_linear_map.to_normed_group, .. continuous_linear_map.to_semi_normed_ring } +{ .. continuous_linear_map.to_normed_add_comm_group, .. continuous_linear_map.to_semi_normed_ring } variable {f} @@ -1295,7 +1296,7 @@ section completeness open_locale topological_space open filter -variables {E' : Type*} [semi_normed_group E'] [normed_space 𝕜 E'] [ring_hom_isometric σ₁₂] +variables {E' : Type*} [seminormed_add_comm_group E'] [normed_space 𝕜 E'] [ring_hom_isometric σ₁₂] /-- Construct a bundled continuous (semi)linear map from a map `f : E → F` and a proof of the fact that it belongs to the closure of the image of a bounded set `s : set (E →SL[σ₁₂] F)` under coercion @@ -1558,7 +1559,7 @@ variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [normed_space 𝕜 Fₗ] (c : 𝕜) {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} -variables {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₂'] {F' : Type*} [normed_group F'] +variables {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₂'] {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜₂' F'] {σ₂' : 𝕜₂' →+* 𝕜₂} {σ₂'' : 𝕜₂ →+* 𝕜₂'} {σ₂₃' : 𝕜₂' →+* 𝕜₃} [ring_hom_inv_pair σ₂' σ₂''] [ring_hom_inv_pair σ₂'' σ₂'] @@ -1765,7 +1766,7 @@ end linear_equiv.coord_self 𝕜 E x h variables {𝕜} {𝕜₄ : Type*} [nontrivially_normed_field 𝕜₄] -variables {H : Type*} [normed_group H] [normed_space 𝕜₄ H] [normed_space 𝕜₃ G] +variables {H : Type*} [normed_add_comm_group H] [normed_space 𝕜₄ H] [normed_space 𝕜₃ G] variables {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} variables {σ₃₄ : 𝕜₃ →+* 𝕜₄} {σ₄₃ : 𝕜₄ →+* 𝕜₃} variables {σ₂₄ : 𝕜₂ →+* 𝕜₄} {σ₁₄ : 𝕜 →+* 𝕜₄} @@ -1798,7 +1799,7 @@ omit σ₂₁ σ₃₄ σ₁₃ σ₂₄ /-- A pair of continuous linear equivalences generates an continuous linear equivalence between the spaces of continuous linear maps. -/ -def arrow_congr {F H : Type*} [normed_group F] [normed_group H] +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) := @@ -1815,6 +1816,6 @@ A bounded bilinear form `B` in a real normed space is *coercive* if there is some positive constant C such that `C * ∥u∥ * ∥u∥ ≤ B u u`. -/ def is_coercive - [normed_group E] [normed_space ℝ E] + [normed_add_comm_group E] [normed_space ℝ E] (B : E →L[ℝ] E →L[ℝ] ℝ) : Prop := ∃ C, (0 < C) ∧ ∀ u, C * ∥u∥ * ∥u∥ ≤ B u u diff --git a/src/analysis/normed_space/ordered.lean b/src/analysis/normed_space/ordered.lean index 92612591ab467..218891f82576e 100644 --- a/src/analysis/normed_space/ordered.lean +++ b/src/analysis/normed_space/ordered.lean @@ -15,14 +15,14 @@ These are mostly useful to avoid diamonds during type class inference. open filter set open_locale topological_space -/-- A `normed_linear_ordered_group` is an additive group that is both a `normed_group` and +/-- A `normed_linear_ordered_group` is an additive group that is both a `normed_add_comm_group` and a `linear_ordered_add_comm_group`. This class is necessary to avoid diamonds. -/ class normed_linear_ordered_group (α : Type*) extends linear_ordered_add_comm_group α, has_norm α, metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) -@[priority 100] instance normed_linear_ordered_group.to_normed_group (α : Type*) - [normed_linear_ordered_group α] : normed_group α := +@[priority 100] instance normed_linear_ordered_group.to_normed_add_comm_group (α : Type*) + [normed_linear_ordered_group α] : normed_add_comm_group α := ⟨normed_linear_ordered_group.dist_eq⟩ /-- A `normed_linear_ordered_field` is a field that is both a `normed_field` and a diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 93a1b655f9629..fa3691400a417 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -280,47 +280,49 @@ antilipschitz_with_equiv_aux p β /-- seminormed group instance on the product of finitely many normed groups, using the `L^p` norm. -/ -instance semi_normed_group [Π i, semi_normed_group (β i)] : semi_normed_group (pi_Lp p β) := +instance seminormed_add_comm_group [Π i, seminormed_add_comm_group (β i)] : + seminormed_add_comm_group (pi_Lp p β) := { norm := λf, (∑ i, ∥f i∥ ^ p) ^ (1/p), dist_eq := λ x y, by simp [pi_Lp.dist_eq, dist_eq_norm, sub_eq_add_neg], .. pi.add_comm_group } /-- normed group instance on the product of finitely many normed groups, using the `L^p` norm. -/ -instance normed_group [Π i, normed_group (α i)] : normed_group (pi_Lp p α) := -{ ..pi_Lp.semi_normed_group p α } +instance normed_add_comm_group [Π i, normed_add_comm_group (α i)] : + normed_add_comm_group (pi_Lp p α) := +{ ..pi_Lp.seminormed_add_comm_group p α } omit fact_one_le_p lemma norm_eq {p : ℝ} [fact (1 ≤ p)] {β : ι → Type*} - [Π i, semi_normed_group (β i)] (f : pi_Lp p β) : + [Π i, seminormed_add_comm_group (β i)] (f : pi_Lp p β) : ∥f∥ = (∑ i, ∥f i∥ ^ p) ^ (1/p) := rfl lemma nnnorm_eq {p : ℝ} [fact (1 ≤ p)] {β : ι → Type*} - [Π i, semi_normed_group (β i)] (f : pi_Lp p β) : + [Π i, seminormed_add_comm_group (β i)] (f : pi_Lp p β) : ∥f∥₊ = (∑ i, ∥f i∥₊ ^ p) ^ (1/p) := by { ext, simp [nnreal.coe_sum, norm_eq] } lemma norm_eq_of_nat {p : ℝ} [fact (1 ≤ p)] {β : ι → Type*} - [Π i, semi_normed_group (β i)] (n : ℕ) (h : p = n) (f : pi_Lp p β) : + [Π i, seminormed_add_comm_group (β i)] (n : ℕ) (h : p = n) (f : pi_Lp p β) : ∥f∥ = (∑ i, ∥f i∥ ^ n) ^ (1/(n : ℝ)) := by simp [norm_eq, h, real.sqrt_eq_rpow, ←real.rpow_nat_cast] -lemma norm_eq_of_L2 {β : ι → Type*} [Π i, semi_normed_group (β i)] (x : pi_Lp 2 β) : +lemma norm_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) : ∥x∥ = sqrt (∑ (i : ι), ∥x i∥ ^ 2) := by { rw [norm_eq_of_nat 2]; simp [sqrt_eq_rpow] } -lemma nnnorm_eq_of_L2 {β : ι → Type*} [Π i, semi_normed_group (β i)] (x : pi_Lp 2 β) : +lemma nnnorm_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) : ∥x∥₊ = nnreal.sqrt (∑ (i : ι), ∥x i∥₊ ^ 2) := subtype.ext $ by { push_cast, exact norm_eq_of_L2 x } -lemma dist_eq_of_L2 {β : ι → Type*} [Π i, semi_normed_group (β i)] (x y : pi_Lp 2 β) : +lemma dist_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) : dist x y = (∑ i, dist (x i) (y i) ^ 2).sqrt := by simp_rw [dist_eq_norm, norm_eq_of_L2, pi.sub_apply] -lemma nndist_eq_of_L2 {β : ι → Type*} [Π i, semi_normed_group (β i)] (x y : pi_Lp 2 β) : +lemma nndist_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) : nndist x y = (∑ i, nndist (x i) (y i) ^ 2).sqrt := subtype.ext $ by { push_cast, exact dist_eq_of_L2 _ _ } -lemma edist_eq_of_L2 {β : ι → Type*} [Π i, semi_normed_group (β i)] (x y : pi_Lp 2 β) : +lemma edist_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) : edist x y = (∑ i, edist (x i) (y i) ^ 2) ^ (1 / 2 : ℝ) := by simp_rw [pi_Lp.edist_eq, ennreal.rpow_two] @@ -329,7 +331,7 @@ include fact_one_le_p variables [normed_field 𝕜] /-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/ -instance normed_space [Π i, semi_normed_group (β i)] [Π i, normed_space 𝕜 (β i)] : +instance normed_space [Π i, seminormed_add_comm_group (β i)] [Π i, normed_space 𝕜 (β i)] : normed_space 𝕜 (pi_Lp p β) := { norm_smul_le := begin @@ -342,14 +344,14 @@ instance normed_space [Π i, semi_normed_group (β i)] [Π i, normed_space 𝕜 end, .. pi.module ι β 𝕜 } -instance finite_dimensional [Π i, semi_normed_group (β i)] [Π i, normed_space 𝕜 (β i)] +instance finite_dimensional [Π i, seminormed_add_comm_group (β i)] [Π i, normed_space 𝕜 (β i)] [I : ∀ i, finite_dimensional 𝕜 (β i)] : finite_dimensional 𝕜 (pi_Lp p β) := 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, semi_normed_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 @@ -361,7 +363,7 @@ variables (x y : pi_Lp p β) (x' y' : Π i, β i) (i : ι) variables {ι' : Type*} variables [fintype ι'] -variables (p 𝕜) (E : Type*) [normed_group E] [normed_space 𝕜 E] +variables (p 𝕜) (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] /-- An equivalence of finite domains induces a linearly isometric equivalence of finitely supported functions-/ @@ -418,7 +420,7 @@ end @[simp] lemma equiv_symm_smul : (pi_Lp.equiv p β).symm (c • x') = c • (pi_Lp.equiv p β).symm x' := rfl -lemma nnnorm_equiv_symm_const {β} [semi_normed_group β] (b : β) : +lemma nnnorm_equiv_symm_const {β} [seminormed_add_comm_group β] (b : β) : ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥₊ = fintype.card ι ^ (1 / p) * ∥b∥₊ := begin have : p ≠ 0 := (zero_lt_one.trans_le (fact.out $ 1 ≤ p)).ne', @@ -427,15 +429,15 @@ begin nnreal.rpow_one], end -lemma norm_equiv_symm_const {β} [semi_normed_group β] (b : β) : +lemma norm_equiv_symm_const {β} [seminormed_add_comm_group β] (b : β) : ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥ = fintype.card ι ^ (1 / p) * ∥b∥ := (congr_arg coe $ nnnorm_equiv_symm_const b).trans $ by simp -lemma nnnorm_equiv_symm_one {β} [semi_normed_group β] [has_one β] : +lemma nnnorm_equiv_symm_one {β} [seminormed_add_comm_group β] [has_one β] : ∥(pi_Lp.equiv p (λ _ : ι, β)).symm 1∥₊ = fintype.card ι ^ (1 / p) * ∥(1 : β)∥₊ := (nnnorm_equiv_symm_const (1 : β)).trans rfl -lemma norm_equiv_symm_one {β} [semi_normed_group β] [has_one β] : +lemma norm_equiv_symm_one {β} [seminormed_add_comm_group β] [has_one β] : ∥(pi_Lp.equiv p (λ _ : ι, β)).symm 1∥ = fintype.card ι ^ (1 / p) * ∥(1 : β)∥ := (norm_equiv_symm_const (1 : β)).trans rfl diff --git a/src/analysis/normed_space/pointwise.lean b/src/analysis/normed_space/pointwise.lean index 5868265fa1dc5..f965936ea9ec0 100644 --- a/src/analysis/normed_space/pointwise.lean +++ b/src/analysis/normed_space/pointwise.lean @@ -19,8 +19,8 @@ open_locale pointwise topological_space variables {𝕜 E : Type*} [normed_field 𝕜] -section semi_normed_group -variables [semi_normed_group E] [normed_space 𝕜 E] +section seminormed_add_comm_group +variables [seminormed_add_comm_group E] [normed_space 𝕜 E] theorem smul_ball {c : 𝕜} (hc : c ≠ 0) (x : E) (r : ℝ) : c • ball x r = ball (c • x) (∥c∥ * r) := @@ -306,10 +306,10 @@ lemma closed_ball_sub_closed_ball [proper_space E] (hε : 0 ≤ ε) (hδ : 0 ≤ closed_ball a ε - closed_ball b δ = closed_ball (a - b) (ε + δ) := by simp_rw [sub_eq_add_neg, neg_closed_ball, closed_ball_add_closed_ball hε hδ] -end semi_normed_group +end seminormed_add_comm_group -section normed_group -variables [normed_group E] [normed_space 𝕜 E] +section normed_add_comm_group +variables [normed_add_comm_group E] [normed_space 𝕜 E] theorem smul_closed_ball (c : 𝕜) (x : E) {r : ℝ} (hr : 0 ≤ r) : c • closed_ball x r = closed_ball (c • x) (∥c∥ * r) := @@ -360,4 +360,4 @@ lemma affinity_unit_closed_ball {r : ℝ} (hr : 0 ≤ r) (x : E) : x +ᵥ r • closed_ball 0 1 = closed_ball x r := by rw [smul_closed_unit_ball, real.norm_of_nonneg hr, vadd_closed_ball_zero] -end normed_group +end normed_add_comm_group diff --git a/src/analysis/normed_space/ray.lean b/src/analysis/normed_space/ray.lean index 992d0ead9f667..338d77836a30d 100644 --- a/src/analysis/normed_space/ray.lean +++ b/src/analysis/normed_space/ray.lean @@ -16,8 +16,8 @@ norms and `∥y∥ • x = ∥x∥ • y`. open real -variables {E : Type*} [semi_normed_group E] [normed_space ℝ E] - {F : Type*} [normed_group F] [normed_space ℝ F] +variables {E : Type*} [seminormed_add_comm_group E] [normed_space ℝ E] + {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] namespace same_ray diff --git a/src/analysis/normed_space/riesz_lemma.lean b/src/analysis/normed_space/riesz_lemma.lean index a38eb7fa8e3d5..ac4c101c7b26d 100644 --- a/src/analysis/normed_space/riesz_lemma.lean +++ b/src/analysis/normed_space/riesz_lemma.lean @@ -24,8 +24,8 @@ open set metric open_locale topological_space variables {𝕜 : Type*} [normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [semi_normed_group F] [normed_space ℝ F] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [seminormed_add_comm_group F] [normed_space ℝ F] /-- Riesz's lemma, which usually states that it is possible to find a vector with norm 1 whose distance to a closed proper subspace is diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 225882bb214ea..0ee496043638a 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -35,7 +35,7 @@ open_locale topological_space local postfix `⋆`:std.prec.max_plus := star /-- A normed star group is a normed group with a compatible `star` which is isometric. -/ -class normed_star_group (E : Type*) [semi_normed_group E] [star_add_monoid E] : Prop := +class normed_star_group (E : Type*) [seminormed_add_comm_group E] [star_add_monoid E] : Prop := (norm_star : ∀ x : E, ∥x⋆∥ = ∥x∥) export normed_star_group (norm_star) @@ -44,12 +44,12 @@ attribute [simp] norm_star variables {𝕜 E α : Type*} section normed_star_group -variables [semi_normed_group E] [star_add_monoid E] [normed_star_group E] +variables [seminormed_add_comm_group E] [star_add_monoid E] [normed_star_group E] @[simp] lemma nnnorm_star (x : E) : ∥star x∥₊ = ∥x∥₊ := subtype.ext $ norm_star _ /-- The `star` map in a normed star group is a normed group homomorphism. -/ -def star_normed_group_hom : normed_group_hom E E := +def star_normed_add_group_hom : normed_add_group_hom E E := { bound' := ⟨1, λ v, le_trans (norm_star _).le (one_mul _).symm.le⟩, .. star_add_equiv } @@ -179,7 +179,7 @@ nnnorm_pow_two_pow_of_self_adjoint x.property _ section starₗᵢ variables [comm_semiring 𝕜] [star_ring 𝕜] -variables [semi_normed_group E] [star_add_monoid E] [normed_star_group E] +variables [seminormed_add_comm_group E] [star_add_monoid E] [normed_star_group E] variables [module 𝕜 E] [star_module 𝕜 E] variables (𝕜) diff --git a/src/analysis/normed_space/star/matrix.lean b/src/analysis/normed_space/star/matrix.lean index 356d1b43b9431..62fc8b30f3606 100644 --- a/src/analysis/normed_space/star/matrix.lean +++ b/src/analysis/normed_space/star/matrix.lean @@ -53,7 +53,7 @@ begin exact norm_sum, end -local attribute [instance] matrix.normed_group +local attribute [instance] matrix.normed_add_comm_group /-- The entrywise sup norm of a unitary matrix is at most 1. -/ lemma entrywise_sup_norm_bound_of_unitary {U : matrix n n 𝕜} (hU : U ∈ matrix.unitary_group n 𝕜) : diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index 4f45144673891..03f09ac772803 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -99,7 +99,7 @@ by the dual-norm (i.e. the operator-norm). -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] +variables {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] namespace normed_space diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index fd31357a301a1..6963ae10a748a 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -896,7 +896,7 @@ end seminorm /-! ### The norm as a seminorm -/ section norm_seminorm -variables (𝕜) (E) [normed_field 𝕜] [semi_normed_group E] [normed_space 𝕜 E] {r : ℝ} +variables (𝕜) (E) [normed_field 𝕜] [seminormed_add_comm_group E] [normed_space 𝕜 E] {r : ℝ} /-- The norm of a seminormed group as an add_monoid seminorm. -/ def norm_add_group_seminorm : add_group_seminorm E := diff --git a/src/analysis/special_functions/arsinh.lean b/src/analysis/special_functions/arsinh.lean index 2b8ef08eb5d7b..49d41ae88a67b 100644 --- a/src/analysis/special_functions/arsinh.lean +++ b/src/analysis/special_functions/arsinh.lean @@ -167,7 +167,7 @@ end continuous section fderiv -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {s : set E} {a : E} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {s : set E} {a : E} {f' : E →L[ℝ] ℝ} {n : with_top ℕ} lemma has_strict_fderiv_at.arsinh (hf : has_strict_fderiv_at f f' a) : diff --git a/src/analysis/special_functions/complex/log_deriv.lean b/src/analysis/special_functions/complex/log_deriv.lean index 75d63fafccf35..176a8c6e0ea00 100644 --- a/src/analysis/special_functions/complex/log_deriv.lean +++ b/src/analysis/special_functions/complex/log_deriv.lean @@ -68,7 +68,7 @@ section log_deriv open complex filter open_locale topological_space -variables {α : Type*} [topological_space α] {E : Type*} [normed_group E] [normed_space ℂ E] +variables {α : Type*} [topological_space α] {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] lemma has_strict_fderiv_at.clog {f : E → ℂ} {f' : E →L[ℂ] ℂ} {x : E} (h₁ : has_strict_fderiv_at f f' x) (h₂ : 0 < (f x).re ∨ (f x).im ≠ 0) : diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 136d5a6db6d90..4c88a9c96fef7 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -102,7 +102,7 @@ end section variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 ℂ] - {E : Type*} [normed_group E] [normed_space 𝕜 E] {f : E → ℂ} {f' : E →L[𝕜] ℂ} + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {f : E → ℂ} {f' : E →L[𝕜] ℂ} {x : E} {s : set E} lemma has_strict_fderiv_at.cexp (hf : has_strict_fderiv_at f f' x) : @@ -213,7 +213,7 @@ section /-! Register lemmas for the derivatives of the composition of `real.exp` with a differentiable function, for standalone use and use with `simp`. -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} {x : E} {s : set E} lemma cont_diff.exp {n} (hf : cont_diff ℝ n f) : diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index 2a95b04277468..dc732754a436d 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -115,8 +115,8 @@ end deriv section fderiv -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {x : E} {f' : E →L[ℝ] ℝ} - {s : set E} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {x : E} + {f' : E →L[ℝ] ℝ} {s : set E} lemma has_fderiv_within_at.log (hf : has_fderiv_within_at f f' s x) (hx : f x ≠ 0) : has_fderiv_within_at (λ x, log (f x)) ((f x)⁻¹ • f') s x := diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index e23d717bfece3..07377437342e9 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -39,8 +39,8 @@ integrable function open_locale measure_theory topological_space interval nnreal ennreal open measure_theory topological_space set filter asymptotics interval_integral -variables {E F : Type*} [normed_group E] [normed_space ℝ E] [second_countable_topology E] -[complete_space E] [normed_group F] +variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [second_countable_topology E] +[complete_space E] [normed_add_comm_group F] /-- If `f` is eventually differentiable along a nontrivial filter `l : filter ℝ` that is generated by convex sets, the norm of `f` tends to infinity along `l`, and `f' = O(g)` along `l`, where `f'` diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 17ac6a49383f2..7bf5c29748baa 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -125,7 +125,7 @@ begin end theorem integral_comp_polar_coord_symm - {E : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] (f : ℝ × ℝ → E) : + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (f : ℝ × ℝ → E) : ∫ p in polar_coord.target, p.1 • f (polar_coord.symm p) = ∫ p, f p := begin set B : (ℝ × ℝ) → ((ℝ × ℝ) →L[ℝ] (ℝ × ℝ)) := λ p, diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow_deriv.lean index 018aff857af47..d43a491aa916d 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow_deriv.lean @@ -68,7 +68,7 @@ section fderiv open complex -variables {E : Type*} [normed_group E] [normed_space ℂ E] {f g : E → ℂ} {f' g' : E →L[ℂ] ℂ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {f g : E → ℂ} {f' g' : E →L[ℂ] ℂ} {x : E} {s : set E} {c : ℂ} lemma has_strict_fderiv_at.cpow (hf : has_strict_fderiv_at f f' x) @@ -339,7 +339,7 @@ open real section fderiv -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f g : E → ℝ} {f' g' : E →L[ℝ] ℝ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f g : E → ℝ} {f' g' : E →L[ℝ] ℝ} {x : E} {s : set E} {c p : ℝ} {n : with_top ℕ} lemma has_fderiv_within_at.rpow (hf : has_fderiv_within_at f f' s x) diff --git a/src/analysis/special_functions/sqrt.lean b/src/analysis/special_functions/sqrt.lean index f9aabd7ad5ad8..6a28132dbe8d3 100644 --- a/src/analysis/special_functions/sqrt.lean +++ b/src/analysis/special_functions/sqrt.lean @@ -98,7 +98,7 @@ end deriv section fderiv -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {n : with_top ℕ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {n : with_top ℕ} {s : set E} {x : E} {f' : E →L[ℝ] ℝ} lemma has_fderiv_at.sqrt (hf : has_fderiv_at f f' x) (hx : f x ≠ 0) : diff --git a/src/analysis/special_functions/trigonometric/arctan_deriv.lean b/src/analysis/special_functions/trigonometric/arctan_deriv.lean index 48990ebb7bde6..80962df871ee9 100644 --- a/src/analysis/special_functions/trigonometric/arctan_deriv.lean +++ b/src/analysis/special_functions/trigonometric/arctan_deriv.lean @@ -131,8 +131,8 @@ end deriv section fderiv -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} {x : E} - {s : set E} {n : with_top ℕ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} + {x : E} {s : set E} {n : with_top ℕ} lemma has_strict_fderiv_at.arctan (hf : has_strict_fderiv_at f f' x) : has_strict_fderiv_at (λ x, arctan (f x)) ((1 / (1 + (f x)^2)) • f') x := diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index 4b03da617ab50..22f70695f6738 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -241,7 +241,7 @@ end section /-! ### Simp lemmas for derivatives of `λ x, complex.cos (f x)` etc., `f : E → ℂ` -/ -variables {E : Type*} [normed_group E] [normed_space ℂ E] {f : E → ℂ} {f' : E →L[ℂ] ℂ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {f : E → ℂ} {f' : E →L[ℂ] ℂ} {x : E} {s : set E} /-! #### `complex.cos` -/ @@ -713,7 +713,7 @@ section /-! ### Simp lemmas for derivatives of `λ x, real.cos (f x)` etc., `f : E → ℝ` -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} {x : E} {s : set E} /-! #### `real.cos` -/ diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index 61de52b86a593..30eebbd5ef473 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -37,7 +37,7 @@ lemma summable_of_absolute_convergence_real {f : ℕ → ℝ} : /-! ### Powers -/ -lemma tendsto_norm_zero' {𝕜 : Type*} [normed_group 𝕜] : +lemma tendsto_norm_zero' {𝕜 : Type*} [normed_add_comm_group 𝕜] : tendsto (norm : 𝕜 → ℝ) (𝓝[≠] 0) (𝓝[>] 0) := tendsto_norm_zero.inf $ tendsto_principal_principal.2 $ λ x hx, norm_pos_iff.2 hx @@ -59,7 +59,7 @@ end /-- The (scalar) product of a sequence that tends to zero with a bounded one also tends to zero. -/ lemma tendsto_zero_smul_of_tendsto_zero_of_bounded {ι 𝕜 𝔸 : Type*} [normed_field 𝕜] - [normed_group 𝔸] [normed_space 𝕜 𝔸] {l : filter ι} {ε : ι → 𝕜} {f : ι → 𝔸} + [normed_add_comm_group 𝔸] [normed_space 𝕜 𝔸] {l : filter ι} {ε : ι → 𝕜} {f : ι → 𝔸} (hε : tendsto ε l (𝓝 0)) (hf : filter.is_bounded_under (≤) l (norm ∘ f)) : tendsto (ε • f) l (𝓝 0) := begin @@ -342,9 +342,9 @@ end mul_geometric section summable_le_geometric -variables [semi_normed_group α] {r C : ℝ} {f : ℕ → α} +variables [seminormed_add_comm_group α] {r C : ℝ} {f : ℕ → α} -lemma semi_normed_group.cauchy_seq_of_le_geometric {C : ℝ} {r : ℝ} (hr : r < 1) +lemma seminormed_add_comm_group.cauchy_seq_of_le_geometric {C : ℝ} {r : ℝ} (hr : r < 1) {u : ℕ → α} (h : ∀ n, ∥u n - u (n + 1)∥ ≤ C*r^n) : cauchy_seq u := cauchy_seq_of_le_geometric r C hr (by simpa [dist_eq_norm] using h) @@ -386,11 +386,11 @@ lemma cauchy_series_of_le_geometric {C : ℝ} {u : ℕ → α} {r : ℝ} (hr : r < 1) (h : ∀ n, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range n, u k) := cauchy_seq_of_le_geometric r C hr (by simp [h]) -lemma normed_group.cauchy_series_of_le_geometric' {C : ℝ} {u : ℕ → α} {r : ℝ} (hr : r < 1) +lemma normed_add_comm_group.cauchy_series_of_le_geometric' {C : ℝ} {u : ℕ → α} {r : ℝ} (hr : r < 1) (h : ∀ n, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := (cauchy_series_of_le_geometric hr h).comp_tendsto $ tendsto_add_at_top_nat 1 -lemma normed_group.cauchy_series_of_le_geometric'' {C : ℝ} {u : ℕ → α} {N : ℕ} {r : ℝ} +lemma normed_add_comm_group.cauchy_series_of_le_geometric'' {C : ℝ} {u : ℕ → α} {N : ℕ} {r : ℝ} (hr₀ : 0 < r) (hr₁ : r < 1) (h : ∀ n ≥ N, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := begin @@ -400,7 +400,8 @@ begin have : ∀ n ≥ N, u n = v n, { intros n hn, simp [v, hn, if_neg (not_lt.mpr hn)] }, - refine cauchy_seq_sum_of_eventually_eq this (normed_group.cauchy_series_of_le_geometric' hr₁ _), + refine cauchy_seq_sum_of_eventually_eq this (normed_add_comm_group.cauchy_series_of_le_geometric' + hr₁ _), { exact C }, intro n, dsimp [v], @@ -473,8 +474,8 @@ end normed_ring_geometric /-! ### Summability tests based on comparison with geometric series -/ -lemma summable_of_ratio_norm_eventually_le {α : Type*} [semi_normed_group α] [complete_space α] - {f : ℕ → α} {r : ℝ} (hr₁ : r < 1) +lemma summable_of_ratio_norm_eventually_le {α : Type*} [seminormed_add_comm_group α] + [complete_space α] {f : ℕ → α} {r : ℝ} (hr₁ : r < 1) (h : ∀ᶠ n in at_top, ∥f (n+1)∥ ≤ r * ∥f n∥) : summable f := begin by_cases hr₀ : 0 ≤ r, @@ -495,7 +496,7 @@ begin exact not_lt.mpr (norm_nonneg _) (lt_of_le_of_lt hn $ mul_neg_of_neg_of_pos hr₀ h), }, end -lemma summable_of_ratio_test_tendsto_lt_one {α : Type*} [normed_group α] [complete_space α] +lemma summable_of_ratio_test_tendsto_lt_one {α : Type*} [normed_add_comm_group α] [complete_space α] {f : ℕ → α} {l : ℝ} (hl₁ : l < 1) (hf : ∀ᶠ n in at_top, f n ≠ 0) (h : tendsto (λ n, ∥f (n+1)∥/∥f n∥) at_top (𝓝 l)) : summable f := begin @@ -505,7 +506,7 @@ begin rwa ← div_le_iff (norm_pos_iff.mpr h₁), end -lemma not_summable_of_ratio_norm_eventually_ge {α : Type*} [semi_normed_group α] +lemma not_summable_of_ratio_norm_eventually_ge {α : Type*} [seminormed_add_comm_group α] {f : ℕ → α} {r : ℝ} (hr : 1 < r) (hf : ∃ᶠ n in at_top, ∥f n∥ ≠ 0) (h : ∀ᶠ n in at_top, r * ∥f n∥ ≤ ∥f (n+1)∥) : ¬ summable f := begin @@ -528,7 +529,7 @@ begin ac_refl } end -lemma not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [semi_normed_group α] +lemma not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [seminormed_add_comm_group α] {f : ℕ → α} {l : ℝ} (hl : 1 < l) (h : tendsto (λ n, ∥f (n+1)∥/∥f n∥) at_top (𝓝 l)) : ¬ summable f := begin @@ -545,7 +546,7 @@ end section /-! ### Dirichlet and alternating series tests -/ -variables {E : Type*} [normed_group E] [normed_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] variables {b : ℝ} {f : ℕ → ℝ} {z : ℕ → E} /-- **Dirichlet's Test** for monotone sequences. -/ diff --git a/src/analysis/von_neumann_algebra/basic.lean b/src/analysis/von_neumann_algebra/basic.lean index 657402fed8c75..720b1dd2b0401 100644 --- a/src/analysis/von_neumann_algebra/basic.lean +++ b/src/analysis/von_neumann_algebra/basic.lean @@ -43,7 +43,7 @@ and we may be unhappy with the resulting opaqueness of the definition. -/ class wstar_algebra (M : Type u) [normed_ring M] [star_ring M] [cstar_ring M] [module ℂ M] [normed_algebra ℂ M] [star_module ℂ M] := -(exists_predual : ∃ (X : Type u) [normed_group X] [normed_space ℂ X] [complete_space X], +(exists_predual : ∃ (X : Type u) [normed_add_comm_group X] [normed_space ℂ X] [complete_space X], nonempty (normed_space.dual ℂ X ≃ₗᵢ⋆[ℂ] M)) -- TODO: Without this, `von_neumann_algebra` times out. Why? diff --git a/src/combinatorics/additive/salem_spencer.lean b/src/combinatorics/additive/salem_spencer.lean index 71b65b62a5537..fd646013d86b3 100644 --- a/src/combinatorics/additive/salem_spencer.lean +++ b/src/combinatorics/additive/salem_spencer.lean @@ -263,8 +263,8 @@ begin (add_halves _) hc.2, end -lemma add_salem_spencer_sphere [normed_group E] [normed_space ℝ E] [strict_convex_space ℝ E] (x : E) - (r : ℝ) : add_salem_spencer (sphere x r) := +lemma add_salem_spencer_sphere [normed_add_comm_group E] [normed_space ℝ E] + [strict_convex_space ℝ E] (x : E) (r : ℝ) : add_salem_spencer (sphere x r) := begin obtain rfl | hr := eq_or_ne r 0, { rw sphere_zero, diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index 914b4a9ec76d2..1395d096a85bc 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -699,7 +699,7 @@ library_note "is_R_or_C instance" simp [re_add_im a, algebra.smul_def, algebra_map_eq_of_real] end⟩⟩ -variables (K) (E : Type*) [normed_group E] [normed_space K E] +variables (K) (E : Type*) [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. diff --git a/src/geometry/manifold/algebra/left_invariant_derivation.lean b/src/geometry/manifold/algebra/left_invariant_derivation.lean index 869534128aedf..0d3851627160c 100644 --- a/src/geometry/manifold/algebra/left_invariant_derivation.lean +++ b/src/geometry/manifold/algebra/left_invariant_derivation.lean @@ -24,7 +24,7 @@ noncomputable theory open_locale lie_group manifold derivation variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) (G : Type*) [topological_space G] [charted_space H G] [monoid G] [has_smooth_mul I G] (g h : G) diff --git a/src/geometry/manifold/algebra/lie_group.lean b/src/geometry/manifold/algebra/lie_group.lean index 81d4dde0809c6..7866fc867b9a0 100644 --- a/src/geometry/manifold/algebra/lie_group.lean +++ b/src/geometry/manifold/algebra/lie_group.lean @@ -46,7 +46,7 @@ the addition and negation operations are smooth. -/ @[ancestor has_smooth_add] class lie_add_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] - {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [add_group G] [topological_space G] [charted_space H G] extends has_smooth_add I G : Prop := (smooth_neg : smooth I I (λ a:G, -a)) @@ -57,7 +57,7 @@ the multiplication and inverse operations are smooth. -/ @[ancestor has_smooth_mul, to_additive] class lie_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] - {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [group G] [topological_space G] [charted_space H G] extends has_smooth_mul I G : Prop := (smooth_inv : smooth I I (λ a:G, a⁻¹)) @@ -66,13 +66,13 @@ section lie_group variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] -{E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} -{F : Type*} [normed_group F] [normed_space 𝕜 F] {J : model_with_corners 𝕜 F F} +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {J : model_with_corners 𝕜 F F} {G : Type*} [topological_space G] [charted_space H G] [group G] [lie_group I G] -{E' : Type*} [normed_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'} {M : Type*} [topological_space M] [charted_space H' M] -{E'' : Type*} [normed_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''} {M' : Type*} [topological_space M'] [charted_space H'' M'] @@ -122,9 +122,9 @@ section prod_lie_group /- Instance of product group -/ @[to_additive] instance {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] - {E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [topological_space G] [charted_space H G] [group G] [lie_group I G] - {E' : Type*} [normed_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'} {G' : Type*} [topological_space G'] [charted_space H' G'] [group G'] [lie_group I' G'] : @@ -137,7 +137,7 @@ end prod_lie_group /-! ### Normed spaces are Lie groups -/ instance normed_space_lie_add_group {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] : + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] : lie_add_group (𝓘(𝕜, E)) E := { smooth_add := smooth_iff.2 ⟨continuous_add, λ x y, cont_diff_add.cont_diff_on⟩, smooth_neg := smooth_iff.2 ⟨continuous_neg, λ x y, cont_diff_neg.cont_diff_on⟩, diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index 3c0f0e31719ac..ce893551579e5 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -44,7 +44,7 @@ instances `add_monoid α` and `has_smooth_add α`. -/ @[ancestor smooth_manifold_with_corners] class has_smooth_add {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] - {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [has_add G] [topological_space G] [charted_space H G] extends smooth_manifold_with_corners I G : Prop := (smooth_add : smooth (I.prod I) I (λ p : G×G, p.1 + p.2)) @@ -56,7 +56,7 @@ and `has_smooth_mul I G`. -/ @[ancestor smooth_manifold_with_corners, to_additive] class has_smooth_mul {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] - {E : Type*} [normed_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (I : model_with_corners 𝕜 E H) (G : Type*) [has_mul G] [topological_space G] [charted_space H G] extends smooth_manifold_with_corners I G : Prop := (smooth_mul : smooth (I.prod I) I (λ p : G×G, p.1 * p.2)) @@ -65,9 +65,9 @@ section has_smooth_mul variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] -{E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [has_mul G] [topological_space G] [charted_space H G] [has_smooth_mul I G] -{E' : Type*} [normed_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'} {M : Type*} [topological_space M] [charted_space H' M] @@ -187,11 +187,11 @@ end /- Instance of product -/ @[to_additive] instance has_smooth_mul.prod {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) (G : Type*) [topological_space G] [charted_space H G] [has_mul G] [has_smooth_mul I G] - {E' : Type*} [normed_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') (G' : Type*) [topological_space G'] [charted_space H' G'] [has_mul G'] [has_smooth_mul I' G'] : @@ -206,10 +206,10 @@ section monoid variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] -{E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [monoid G] [topological_space G] [charted_space H G] [has_smooth_mul I G] {H' : Type*} [topological_space H'] -{E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {I' : model_with_corners 𝕜 E' H'} +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {I' : model_with_corners 𝕜 E' H'} {G' : Type*} [monoid G'] [topological_space G'] [charted_space H' G'] [has_smooth_mul I' G'] lemma smooth_pow : ∀ n : ℕ, smooth I I (λ a : G, a ^ n) @@ -250,9 +250,9 @@ open_locale big_operators variables {ι 𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] -{E : Type*} [normed_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {I : model_with_corners 𝕜 E H} {G : Type*} [comm_monoid G] [topological_space G] [charted_space H G] [has_smooth_mul I G] -{E' : Type*} [normed_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'} {M : Type*} [topological_space M] [charted_space H' M] {s : set M} {x : M} {t : finset ι} {f : ι → M → G} {n : with_top ℕ} {p : ι → Prop} diff --git a/src/geometry/manifold/algebra/smooth_functions.lean b/src/geometry/manifold/algebra/smooth_functions.lean index 8e5669a3fe367..fe18fe42b2c32 100644 --- a/src/geometry/manifold/algebra/smooth_functions.lean +++ b/src/geometry/manifold/algebra/smooth_functions.lean @@ -17,12 +17,12 @@ noncomputable theory open_locale manifold variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{E' : Type*} [normed_group E'] [normed_space 𝕜 E'] +{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'} {N : Type*} [topological_space N] [charted_space H N] -{E'' : Type*} [normed_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''} {N' : Type*} [topological_space N'] [charted_space H'' N'] @@ -179,26 +179,26 @@ In this section we show that smooth functions valued in a vector space `M` over field `𝕜` inherit a vector space structure. -/ -instance has_smul {V : Type*} [normed_group V] [normed_space 𝕜 V] : +instance has_smul {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] : has_smul 𝕜 C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := ⟨λ r f, ⟨r • f, smooth_const.smul f.smooth⟩⟩ @[simp] -lemma coe_smul {V : Type*} [normed_group V] [normed_space 𝕜 V] +lemma coe_smul {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] (r : 𝕜) (f : C^∞⟮I, N; 𝓘(𝕜, V), V⟯) : ⇑(r • f) = r • f := rfl -@[simp] lemma smul_comp {V : Type*} [normed_group V] [normed_space 𝕜 V] +@[simp] lemma smul_comp {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] (r : 𝕜) (g : C^∞⟮I'', N'; 𝓘(𝕜, V), V⟯) (h : C^∞⟮I, N; I'', N'⟯) : (r • g).comp h = r • (g.comp h) := rfl -instance module {V : Type*} [normed_group V] [normed_space 𝕜 V] : +instance module {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] : module 𝕜 C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := function.injective.module 𝕜 coe_fn_add_monoid_hom cont_mdiff_map.coe_inj coe_smul /-- Coercion to a function as a `linear_map`. -/ @[simps] -def coe_fn_linear_map {V : Type*} [normed_group V] [normed_space 𝕜 V] : +def coe_fn_linear_map {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] : C^∞⟮I, N; 𝓘(𝕜, V), V⟯ →ₗ[𝕜] (N → V) := { to_fun := coe_fn, map_smul' := coe_smul, @@ -254,15 +254,15 @@ section module_over_continuous_functions If `V` is a module over `𝕜`, then we show that the space of smooth functions from `N` to `V` is naturally a vector space over the ring of smooth functions from `N` to `𝕜`. -/ -instance has_smul' {V : Type*} [normed_group V] [normed_space 𝕜 V] : +instance has_smul' {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] : has_smul C^∞⟮I, N; 𝕜⟯ C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := ⟨λ f g, ⟨λ x, (f x) • (g x), (smooth.smul f.2 g.2)⟩⟩ -@[simp] lemma smul_comp' {V : Type*} [normed_group V] [normed_space 𝕜 V] +@[simp] lemma smul_comp' {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] (f : C^∞⟮I'', N'; 𝕜⟯) (g : C^∞⟮I'', N'; 𝓘(𝕜, V), V⟯) (h : C^∞⟮I, N; I'', N'⟯) : (f • g).comp h = (f.comp h) • (g.comp h) := rfl -instance module' {V : Type*} [normed_group V] [normed_space 𝕜 V] : +instance module' {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] : module C^∞⟮I, N; 𝓘(𝕜), 𝕜⟯ C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := { smul := (•), smul_add := λ c f g, by ext x; exact smul_add (c x) (f x) (g x), diff --git a/src/geometry/manifold/algebra/structures.lean b/src/geometry/manifold/algebra/structures.lean index 2ab748bf19d44..b8fad067971af 100644 --- a/src/geometry/manifold/algebra/structures.lean +++ b/src/geometry/manifold/algebra/structures.lean @@ -17,7 +17,7 @@ open_locale manifold section smooth_ring variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H : Type*} [topological_space H] -{E : Type*} [normed_group E] [normed_space 𝕜 E] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] set_option default_priority 100 -- see Note [default priority] @@ -55,7 +55,7 @@ instance field_smooth_ring {𝕜 : Type*} [nontrivially_normed_field 𝕜] : ..normed_space_lie_add_group } variables {𝕜 R E H : Type*} [topological_space R] [topological_space H] - [nontrivially_normed_field 𝕜] [normed_group E] [normed_space 𝕜 E] + [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] [charted_space H R] (I : model_with_corners 𝕜 E H) /-- A smooth (semi)ring is a topological (semi)ring. This is not an instance for technical reasons, diff --git a/src/geometry/manifold/bump_function.lean b/src/geometry/manifold/bump_function.lean index c9d5755434bde..a9e8b916682db 100644 --- a/src/geometry/manifold/bump_function.lean +++ b/src/geometry/manifold/bump_function.lean @@ -28,7 +28,7 @@ manifold, smooth bump function universes uE uF uH uM variables -{E : Type uE} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +{E : Type uE} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {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] @@ -309,7 +309,7 @@ protected lemma continuous : continuous f := f.smooth.continuous /-- If `f : smooth_bump_function I c` is a smooth bump function and `g : M → G` is a function smooth on the source of the chart at `c`, then `f • g` is smooth on the whole manifold. -/ -lemma smooth_smul {G} [normed_group G] [normed_space ℝ G] +lemma smooth_smul {G} [normed_add_comm_group G] [normed_space ℝ G] {g : M → G} (hg : smooth_on I 𝓘(ℝ, G) g (chart_at H c).source) : smooth I 𝓘(ℝ, G) (λ x, f x • g x) := begin diff --git a/src/geometry/manifold/conformal_groupoid.lean b/src/geometry/manifold/conformal_groupoid.lean index b1f8c1fc6573a..7ea9bc0fe5c17 100644 --- a/src/geometry/manifold/conformal_groupoid.lean +++ b/src/geometry/manifold/conformal_groupoid.lean @@ -20,7 +20,7 @@ In this file we define the groupoid of conformal maps on normed spaces. conformal, groupoid -/ -variables {X : Type*} [normed_group X] [normed_space ℝ X] +variables {X : Type*} [normed_add_comm_group X] [normed_space ℝ X] /-- The pregroupoid of conformal maps. -/ def conformal_pregroupoid : pregroupoid X := diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 92239f4f4db43..56be1c95a03ff 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -54,19 +54,19 @@ open_locale topological_space manifold variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -- declare a smooth manifold `M` over the pair `(E, H)`. -{E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] [Is : smooth_manifold_with_corners I M] -- declare a smooth manifold `M'` over the pair `(E', H')`. -{E' : Type*} [normed_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') {M' : Type*} [topological_space M'] [charted_space H' M'] [I's : smooth_manifold_with_corners I' M'] -- declare a smooth manifold `N` over the pair `(F, G)`. -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {G : Type*} [topological_space G] {J : model_with_corners 𝕜 F G} {N : Type*} [topological_space N] [charted_space G N] [Js : smooth_manifold_with_corners J N] -- declare a smooth manifold `N'` over the pair `(F', G')`. -{F' : Type*} [normed_group F'] [normed_space 𝕜 F'] +{F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] {G' : Type*} [topological_space G'] {J' : model_with_corners 𝕜 F' G'} {N' : Type*} [topological_space N'] [charted_space G' N'] [J's : smooth_manifold_with_corners J' N'] -- declare functions, sets, points and smoothness indices @@ -778,7 +778,7 @@ lemma cont_mdiff_of_locally_cont_mdiff_on section composition -variables {E'' : Type*} [normed_group E''] [normed_space 𝕜 E''] +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''] @@ -1871,7 +1871,7 @@ We have no `model_with_corners.pi` yet, so we prove lemmas about functions `f : use `𝓘(𝕜, Π i, F i)` as the model space. -/ -variables {ι : Type*} [fintype ι] {Fi : ι → Type*} [Π i, normed_group (Fi i)] +variables {ι : Type*} [fintype ι] {Fi : ι → Type*} [Π i, normed_add_comm_group (Fi i)] [Π i, normed_space 𝕜 (Fi i)] {φ : M → Π i, Fi i} lemma cont_mdiff_within_at_pi_space : @@ -1925,7 +1925,7 @@ L.cont_diff.cont_mdiff /-! ### Smoothness of standard operations -/ -variables {V : Type*} [normed_group V] [normed_space 𝕜 V] +variables {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] /-- On any vector space, multiplication by a scalar is a smooth operation. -/ lemma smooth_smul : smooth (𝓘(𝕜).prod 𝓘(𝕜, V)) 𝓘(𝕜, V) (λp : 𝕜 × V, p.1 • p.2) := diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 925680dc95af7..eb0ec6c483c5a 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -15,14 +15,14 @@ bundled maps. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{E' : Type*} [normed_group E'] [normed_space 𝕜 E'] +{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_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''} {M'' : Type*} [topological_space M''] [charted_space H'' M''] diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index c0ac8c6f4b208..68308fc1b2656 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -21,7 +21,7 @@ of the Lie algebra for a Lie group. -/ variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] (n : with_top ℕ) @@ -95,7 +95,7 @@ lemma eval_at_apply (x : M) : eval_at x X f = (X f) x := rfl end derivation -variables {I} {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] +variables {I} {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'] @@ -131,7 +131,7 @@ localized "notation `𝒅ₕ` := hfdifferential" in manifold @[simp] lemma apply_hfdifferential {f : C^∞⟮I, M; I', M'⟯} {x : M} {y : M'} (h : f x = y) (v : point_derivation I x) (g : C^∞⟮I', M'; 𝕜⟯) : 𝒅ₕh v g = 𝒅f x v g := rfl -variables {E'' : Type*} [normed_group E''] [normed_space 𝕜 E''] +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''] diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 1ed56603bd220..72a80945f79fb 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -47,9 +47,9 @@ open_locale manifold topological_space open function set variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{E' : Type*} [normed_group E'] [normed_space 𝕜 E'] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {H : Type*} [topological_space H] {H' : Type*} [topological_space H'] {G : Type*} [topological_space G] diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 96b30a99a9f9b..0c89673dd021e 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -374,7 +374,7 @@ begin (ℝ ∙ ((-v):E))ᗮ.subtypeL.cont_diff).comp U.symm.cont_diff).cont_diff_on } end -variables {F : Type*} [normed_group F] [normed_space ℝ F] +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] variables {H : Type*} [topological_space H] {I : model_with_corners ℝ F H} variables {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] diff --git a/src/geometry/manifold/metrizable.lean b/src/geometry/manifold/metrizable.lean index 34a44698146d2..432c3b5b1dae0 100644 --- a/src/geometry/manifold/metrizable.lean +++ b/src/geometry/manifold/metrizable.lean @@ -19,7 +19,7 @@ open topological_space /-- A σ-compact Hausdorff topological manifold over a finite dimensional real vector space is metrizable. -/ lemma manifold_with_corners.metrizable_space - {E : Type*} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {H : Type*} [topological_space H] (I : model_with_corners ℝ E H) (M : Type*) [topological_space M] [charted_space H M] [sigma_compact_space M] [t2_space M] : metrizable_space M := diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 02735f0b81395..c43b8c20ea839 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -111,10 +111,10 @@ We use the names `mdifferentiable` and `mfderiv`, where the prefix letter `m` me -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] -{E' : Type*} [normed_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') {M' : Type*} [topological_space M'] [charted_space H' M'] @@ -240,13 +240,13 @@ section derivatives_properties /-! ### Unique differentiability sets in manifolds -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] -- -{E' : Type*} [normed_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'} {M' : Type*} [topological_space M'] [charted_space H' M'] -{E'' : Type*} [normed_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''} {M'' : Type*} [topological_space M''] [charted_space H'' M''] {f f₀ f₁ : M → M'} @@ -854,8 +854,8 @@ this and related statements. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{E' : Type*} [normed_group E'] [normed_space 𝕜 E'] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {f : E → E'} {s : set E} {x : E} lemma unique_mdiff_within_at_iff_unique_diff_within_at : @@ -960,10 +960,10 @@ section specific_functions /-! ### Differentiability of specific functions -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -{E' : Type*} [normed_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') {M' : Type*} [topological_space M'] [charted_space H' M'] [smooth_manifold_with_corners I' M'] @@ -1264,13 +1264,13 @@ end specific_functions namespace local_homeomorph.mdifferentiable variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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} {M : Type*} [topological_space M] [charted_space H M] -{E' : Type*} [normed_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'} {M' : Type*} [topological_space M'] [charted_space H' M'] -{E'' : Type*} [normed_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''} {M'' : Type*} [topological_space M''] [charted_space H'' M''] {e : local_homeomorph M M'} (he : e.mdifferentiable I I') @@ -1376,7 +1376,7 @@ end local_homeomorph.mdifferentiable section ext_chart_at variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] {s : set M} {x y : M} @@ -1403,10 +1403,10 @@ end ext_chart_at section unique_mdiff variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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} {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -{E' : Type*} [normed_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'} {M' : Type*} [topological_space M'] [charted_space H' M'] {s : set M} @@ -1516,7 +1516,7 @@ begin exact this.unique_diff_on_target_inter _ end -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] (Z : basic_smooth_vector_bundle_core I M F) /-- In a smooth fiber bundle constructed from core, the preimage under the projection of a set with diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index bf4b138e4dadb..fc5fba78dd5e7 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -64,8 +64,8 @@ open_locale topological_space manifold classical filter big_operators noncomputable theory variables {ι : Type uι} -{E : Type uE} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] -{F : Type uF} [normed_group F] [normed_space ℝ F] +{E : Type uE} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +{F : Type uF} [normed_add_comm_group F] [normed_space ℝ F] {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] @@ -214,7 +214,7 @@ end smooth_partition_of_unity namespace bump_covering -- Repeat variables to drop [finite_dimensional ℝ E] and [smooth_manifold_with_corners I M] -lemma smooth_to_partition_of_unity {E : Type uE} [normed_group E] [normed_space ℝ E] +lemma smooth_to_partition_of_unity {E : Type uE} [normed_add_comm_group E] [normed_space ℝ E] {H : Type uH} [topological_space H] {I : model_with_corners ℝ E H} {M : Type uM} [topological_space M] [charted_space H M] {s : set M} (f : bump_covering ι M s) (hf : ∀ i, smooth I 𝓘(ℝ) (f i)) (i : ι) : diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 4106ac16e6aa0..abc1771348f2c 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -63,7 +63,7 @@ In the same way, it would not apply to product manifolds, modelled on The right invocation does not focus on one specific construction, but on all constructions sharing the right properties, like - `variables {E : Type*} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] + `variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {I : model_with_corners ℝ E E} [I.boundaryless] {M : Type*} [topological_space M] [charted_space E M] [smooth_manifold_with_corners I M]` @@ -127,7 +127,7 @@ define a smooth manifold with model space `H`, and model vector space `E`. -/ @[nolint has_inhabited_instance] structure model_with_corners (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] extends local_equiv H E := (source_eq : source = univ) (unique_diff' : unique_diff_on 𝕜 to_local_equiv.target) @@ -138,7 +138,7 @@ attribute [simp, mfld_simps] model_with_corners.source_eq /-- A vector space is a model with corners. -/ def model_with_corners_self (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] : model_with_corners 𝕜 E E := + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] : model_with_corners 𝕜 E E := { to_local_equiv := local_equiv.refl E, source_eq := rfl, unique_diff' := unique_diff_on_univ, @@ -151,7 +151,7 @@ localized "notation `𝓘(` 𝕜 `)` := model_with_corners_self 𝕜 𝕜" in ma section variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) namespace model_with_corners @@ -164,12 +164,12 @@ protected def symm : local_equiv E H := I.to_local_equiv.symm /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ def simps.apply (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] (I : model_with_corners 𝕜 E H) : H → E := I /-- See Note [custom simps projection] -/ def simps.symm_apply (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] (I : model_with_corners 𝕜 E H) : E → H := I.symm initialize_simps_projections model_with_corners @@ -306,10 +306,10 @@ structure on the tangent bundle to a manifold modelled on `(E, H)`: it will be m vs `H × H'`. -/ @[simps (lemmas_only)] def model_with_corners.prod {𝕜 : Type u} [nontrivially_normed_field 𝕜] - {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] - (I : model_with_corners 𝕜 E H) - {E' : Type v'} [normed_group E'] [normed_space 𝕜 E'] {H' : Type w'} [topological_space H'] - (I' : model_with_corners 𝕜 E' H') : model_with_corners 𝕜 (E × E') (model_prod H H') := + {E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] + (I : model_with_corners 𝕜 E H) {E' : Type v'} [normed_add_comm_group E'] [normed_space 𝕜 E'] + {H' : Type w'} [topological_space H'] (I' : model_with_corners 𝕜 E' H') : + model_with_corners 𝕜 (E × E') (model_prod H H') := { to_fun := λ x, (I x.1, I' x.2), inv_fun := λ x, (I.symm x.1, I'.symm x.2), source := {x | x.1 ∈ I.source ∧ x.2 ∈ I'.source}, @@ -324,7 +324,7 @@ corners `pi I` on `(Π i, E i, model_pi H)`. See note [Manifold type tags] for e `model_pi H`. -/ def model_with_corners.pi {𝕜 : Type u} [nontrivially_normed_field 𝕜] {ι : Type v} [fintype ι] - {E : ι → Type w} [Π i, normed_group (E i)] [Π i, normed_space 𝕜 (E i)] + {E : ι → Type w} [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] {H : ι → Type u'} [Π i, topological_space (H i)] (I : Π i, model_with_corners 𝕜 (E i) (H i)) : model_with_corners 𝕜 (Π i, E i) (model_pi H) := { to_local_equiv := local_equiv.pi (λ i, (I i).to_local_equiv), @@ -337,13 +337,14 @@ def model_with_corners.pi as the model to tangent bundles. -/ @[reducible] def model_with_corners.tangent {𝕜 : Type u} [nontrivially_normed_field 𝕜] - {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] + {E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] (I : model_with_corners 𝕜 E H) : model_with_corners 𝕜 (E × E) (model_prod H E) := I.prod (𝓘(𝕜, E)) -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] -{F : Type*} [normed_group F] [normed_space 𝕜 F] {F' : Type*} [normed_group F'] [normed_space 𝕜 F'] +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'] {F : Type*} + [normed_add_comm_group F] [normed_space 𝕜 F] {F' : Type*} [normed_add_comm_group F'] + [normed_space 𝕜 F'] {H : Type*} [topological_space H] {H' : Type*} [topological_space H'] {G : Type*} [topological_space G] {G' : Type*} [topological_space G'] {I : model_with_corners 𝕜 E H} {J : model_with_corners 𝕜 F G} @@ -366,20 +367,21 @@ section boundaryless /-- Property ensuring that the model with corners `I` defines manifolds without boundary. -/ class model_with_corners.boundaryless {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) : Prop := (range_eq_univ : range I = univ) /-- The trivial model with corners has no boundary -/ instance model_with_corners_self_boundaryless (𝕜 : Type*) [nontrivially_normed_field 𝕜] - (E : Type*) [normed_group E] [normed_space 𝕜 E] : (model_with_corners_self 𝕜 E).boundaryless := + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] : + (model_with_corners_self 𝕜 E).boundaryless := ⟨by simp⟩ /-- If two model with corners are boundaryless, their product also is -/ instance model_with_corners.range_eq_univ_prod {𝕜 : Type u} [nontrivially_normed_field 𝕜] - {E : Type v} [normed_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] - (I : model_with_corners 𝕜 E H) [I.boundaryless] - {E' : Type v'} [normed_group E'] [normed_space 𝕜 E'] {H' : Type w'} [topological_space H'] + {E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type w} [topological_space H] + (I : model_with_corners 𝕜 E H) [I.boundaryless] {E' : Type v'} [normed_add_comm_group E'] + [normed_space 𝕜 E'] {H' : Type w'} [topological_space H'] (I' : model_with_corners 𝕜 E' H') [I'.boundaryless] : (I.prod I').boundaryless := begin @@ -395,7 +397,7 @@ section cont_diff_groupoid /-! ### Smooth functions on models with corners -/ variables {m n : with_top ℕ} {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) {M : Type*} [topological_space M] @@ -504,7 +506,7 @@ begin (of_set_mem_cont_diff_groupoid n I e.open_target) this end -variables {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] {H' : Type*} [topological_space H'] +variables {E' H' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] [topological_space H'] /-- The product of two smooth local homeomorphisms is smooth. -/ lemma cont_diff_groupoid_prod @@ -550,13 +552,13 @@ section smooth_manifold_with_corners field `𝕜` and with infinite smoothness to simplify typeclass search and statements later on. -/ @[ancestor has_groupoid] class smooth_manifold_with_corners {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] extends has_groupoid M (cont_diff_groupoid ∞ I) : Prop lemma smooth_manifold_with_corners.mk' {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] [gr : has_groupoid M (cont_diff_groupoid ∞ I)] : @@ -564,7 +566,7 @@ lemma smooth_manifold_with_corners.mk' {𝕜 : Type*} [nontrivially_normed_field lemma smooth_manifold_with_corners_of_cont_diff_on {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] (h : ∀ (e e' : local_homeomorph M H), e ∈ atlas H M → e' ∈ atlas H M → @@ -579,7 +581,7 @@ lemma smooth_manifold_with_corners_of_cont_diff_on /-- For any model with corners, the model space is a smooth manifold -/ instance model_space_smooth {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} : smooth_manifold_with_corners I H := { .. has_groupoid_model_space _ _ } @@ -591,7 +593,7 @@ charted space with a structure groupoid, avoiding the need to specify the groupo `cont_diff_groupoid ∞ I` explicitly. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] @@ -618,8 +620,8 @@ structure_groupoid.compatible_of_mem_maximal_atlas he he' /-- The product of two smooth manifolds with corners is naturally a smooth manifold with corners. -/ instance prod {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] - {E' : Type*} [normed_group E'] [normed_space 𝕜 E'] + {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] @@ -638,7 +640,7 @@ end smooth_manifold_with_corners lemma local_homeomorph.singleton_smooth_manifold_with_corners {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) {M : Type*} [topological_space M] (e : local_homeomorph M H) (h : e.source = set.univ) : @@ -648,7 +650,7 @@ e.singleton_has_groupoid h (cont_diff_groupoid ∞ I) lemma open_embedding.singleton_smooth_manifold_with_corners {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) {M : Type*} [topological_space M] [nonempty M] {f : M → H} (h : open_embedding f) : @@ -660,7 +662,7 @@ namespace topological_space.opens open topological_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] (s : opens M) @@ -673,7 +675,7 @@ section extended_charts open_locale topological_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] - {E : Type*} [normed_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) {M : Type*} [topological_space M] [charted_space H M] (x : M) {s t : set M} @@ -892,6 +894,6 @@ end extended_charts /-- In the case of the manifold structure on a vector space, the extended charts are just the identity.-/ lemma ext_chart_model_space_eq_id (𝕜 : Type*) [nontrivially_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] (x : E) : + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (x : E) : ext_chart_at (model_with_corners_self 𝕜 E) x = local_equiv.refl E := by simp only with mfld_simps diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean index c2fca4e33bfe7..a0771fdfc9d5d 100644 --- a/src/geometry/manifold/tangent_bundle.lean +++ b/src/geometry/manifold/tangent_bundle.lean @@ -84,10 +84,10 @@ of `M`. This structure registers the changes in the fibers when one changes coor 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_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) (M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -(F : Type*) [normed_group F] [normed_space 𝕜 F] := +(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, @@ -99,10 +99,10 @@ structure basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_fi /-- 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_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) (M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -(F : Type*) [normed_group F] [normed_space 𝕜 F] : basic_smooth_vector_bundle_core I M F := +(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, @@ -111,10 +111,10 @@ def trivial_basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_ namespace basic_smooth_vector_bundle_core variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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} {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -{F : Type*} [normed_group F] [normed_space 𝕜 F] +{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) := @@ -290,7 +290,7 @@ end basic_smooth_vector_bundle_core section tangent_bundle variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_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) (M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] diff --git a/src/geometry/manifold/whitney_embedding.lean b/src/geometry/manifold/whitney_embedding.lean index fe0a799e38d9f..b5be8c77521f8 100644 --- a/src/geometry/manifold/whitney_embedding.lean +++ b/src/geometry/manifold/whitney_embedding.lean @@ -27,7 +27,7 @@ partition of unity, smooth bump function, whitney theorem universes uι uE uH uM variables {ι : Type uι} -{E : Type uE} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +{E : Type uE} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {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] diff --git a/src/information_theory/hamming.lean b/src/information_theory/hamming.lean index 6d0c9a16d68fc..4e8d64251a95f 100644 --- a/src/information_theory/hamming.lean +++ b/src/information_theory/hamming.lean @@ -23,7 +23,7 @@ code. * `hamming β`: a type synonym for `Π i, β i` with `dist` and `norm` provided by the above. * `hamming.to_hamming`, `hamming.of_hamming`: functions for casting between `hamming β` and `Π i, β i`. -* `hamming.normed_group`: the Hamming norm forms a normed group on `hamming β`. +* `hamming.normed_add_comm_group`: the Hamming norm forms a normed group on `hamming β`. -/ section hamming_dist_norm @@ -187,7 +187,7 @@ end hamming_dist_norm /-! ### The `hamming` type synonym -/ /-- Type synonym for a Pi type which inherits the usual algebraic instances, but is equipped with -the Hamming metric and norm, instead of `pi.normed_group` which uses the sup norm. -/ +the Hamming metric and norm, instead of `pi.normed_add_comm_group` which uses the sup norm. -/ def hamming {ι : Type*} (β : ι → Type*) : Type* := Π i, β i namespace hamming @@ -302,13 +302,14 @@ instance [Π i, has_zero (β i)] : has_norm (hamming β) := ⟨λ x, hamming_nor @[simp, push_cast] lemma norm_eq_hamming_norm [Π i, has_zero (β i)] (x : hamming β) : ∥x∥ = hamming_norm (of_hamming x) := rfl -instance [Π i, add_comm_group (β i)] : semi_normed_group (hamming β) := +instance [Π i, add_comm_group (β i)] : seminormed_add_comm_group (hamming β) := { dist_eq := by { push_cast, exact_mod_cast hamming_dist_eq_hamming_norm }, ..pi.add_comm_group } @[simp, push_cast] lemma nnnorm_eq_hamming_norm [Π i, add_comm_group (β i)] (x : hamming β) : ∥x∥₊ = hamming_norm (of_hamming x) := rfl -instance [Π i, add_comm_group (β i)] : normed_group (hamming β) := { ..hamming.semi_normed_group } +instance [Π i, add_comm_group (β i)] : normed_add_comm_group (hamming β) := +{ ..hamming.seminormed_add_comm_group } end diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 50cde7b97745a..4f172c2639b7d 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1847,9 +1847,9 @@ lemma ae_measurable.coe_ereal_ennreal {f : α → ℝ≥0∞} {μ : measure α} ae_measurable (λ x, (f x : ereal)) μ := measurable_coe_ennreal_ereal.comp_ae_measurable hf -section normed_group +section normed_add_comm_group -variables [normed_group α] [opens_measurable_space α] [measurable_space β] +variables [normed_add_comm_group α] [opens_measurable_space α] [measurable_space β] @[measurability] lemma measurable_norm : measurable (norm : α → ℝ) := @@ -1891,7 +1891,7 @@ lemma ae_measurable.ennnorm {f : β → α} {μ : measure β} (hf : ae_measurabl ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := measurable_ennnorm.comp_ae_measurable hf -end normed_group +end normed_add_comm_group section limits @@ -2052,9 +2052,9 @@ end limits namespace continuous_linear_map variables {𝕜 : Type*} [normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E] -variables [opens_measurable_space E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] [measurable_space F] [borel_space F] +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 := @@ -2069,8 +2069,8 @@ end continuous_linear_map namespace continuous_linear_map variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] - {F : Type*} [normed_group F] [normed_space 𝕜 F] +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 _ @@ -2097,8 +2097,8 @@ end continuous_linear_map section continuous_linear_map_nontrivially_normed_field variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E] [borel_space E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] +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) : @@ -2114,8 +2114,8 @@ end continuous_linear_map_nontrivially_normed_field section normed_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] [measurable_space 𝕜] -variables [borel_space 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] [measurable_space E] [borel_space E] +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 := diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index b022a56081e07..8fc239fc3e4eb 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -90,7 +90,7 @@ end variables [measurable_space α] [measurable_space α'] [measurable_space β] [measurable_space β'] variables [measurable_space γ] variables {μ : measure α} {ν : measure β} {τ : measure γ} -variables [normed_group E] +variables [normed_add_comm_group E] /-! ### Measurability @@ -930,7 +930,7 @@ begin rw [← integral_map measurable_swap.ae_measurable hf, prod_swap] end -variables {E' : Type*} [normed_group E'] [complete_space E'] [normed_space ℝ E'] +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. -/ diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index 51ec7115bc341..02288e290a362 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -49,7 +49,7 @@ noncomputable theory namespace besicovitch -variables {E : Type*} [normed_group E] +variables {E : Type*} [normed_add_comm_group E] namespace satellite_config variables [normed_space ℝ E] {N : ℕ} {τ : ℝ} (a : satellite_config E N τ) @@ -123,7 +123,7 @@ end satellite_config /-- The maximum cardinality of a `1`-separated set in the ball of radius `2`. This is also the optimal number of families in the Besicovitch covering theorem. -/ -def multiplicity (E : Type*) [normed_group E] := +def multiplicity (E : Type*) [normed_add_comm_group E] := Sup {N | ∃ s : finset E, s.card = N ∧ (∀ c ∈ s, ∥c∥ ≤ 2) ∧ (∀ c ∈ s, ∀ d ∈ s, c ≠ d → 1 ≤ ∥c - d∥)} section diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index 13bdcfebb7146..137691bae6898 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -70,7 +70,7 @@ local notation `⟪`x`, `y`⟫` := y x variables (𝕜) -lemma ae_eq_zero_of_forall_dual_of_is_separable [normed_group E] [normed_space 𝕜 E] +lemma ae_eq_zero_of_forall_dual_of_is_separable [normed_add_comm_group E] [normed_space 𝕜 E] {t : set E} (ht : topological_space.is_separable t) {f : α → E} (hf : ∀ c : dual 𝕜 E, (λ x, ⟪f x, c⟫) =ᵐ[μ] 0) (h't : ∀ᵐ x ∂μ, f x ∈ t) : f =ᵐ[μ] 0 := @@ -105,8 +105,8 @@ begin exact A (f x) h'x hx, end -lemma ae_eq_zero_of_forall_dual [normed_group E] [normed_space 𝕜 E] [second_countable_topology E] - {f : α → E} (hf : ∀ c : dual 𝕜 E, (λ x, ⟪f x, c⟫) =ᵐ[μ] 0) : +lemma ae_eq_zero_of_forall_dual [normed_add_comm_group E] [normed_space 𝕜 E] + [second_countable_topology E] {f : α → E} (hf : ∀ c : dual 𝕜 E, (λ x, ⟪f x, c⟫) =ᵐ[μ] 0) : f =ᵐ[μ] 0 := ae_eq_zero_of_forall_dual_of_is_separable 𝕜 (is_separable_of_separable_space (set.univ : set E)) hf (eventually_of_forall (λ x, set.mem_univ _)) @@ -118,7 +118,7 @@ end ae_eq_of_forall variables {α E : Type*} {m m0 : measurable_space α} {μ : measure α} {s t : set α} - [normed_group E] [normed_space ℝ E] + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] {p : ℝ≥0∞} diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index a31e6a938552f..245ad2e843b1e 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -227,15 +227,15 @@ variables {α β γ E E' F F' G G' H 𝕜 : Type*} {p : ℝ≥0∞} [inner_product_space 𝕜 E'] [complete_space E'] [normed_space ℝ E'] -- F for a Lp submodule - [normed_group F] [normed_space 𝕜 F] + [normed_add_comm_group F] [normed_space 𝕜 F] -- F' for integrals on a Lp submodule - [normed_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] + [normed_add_comm_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] -- G for a Lp add_subgroup - [normed_group G] + [normed_add_comm_group G] -- G' for integrals on a Lp add_subgroup - [normed_group G'] [normed_space ℝ G'] [complete_space G'] + [normed_add_comm_group G'] [normed_space ℝ G'] [complete_space G'] -- H for a normed group (hypotheses of mem_ℒp) - [normed_group H] + [normed_add_comm_group H] section Lp_meas diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index 80df4cea7918f..f89292c96e266 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -46,7 +46,7 @@ open_locale ennreal nnreal topological_space bounded_continuous_function open measure_theory topological_space continuous_map variables {α : Type*} [measurable_space α] [topological_space α] [normal_space α] [borel_space α] -variables (E : Type*) [normed_group E] +variables (E : Type*) [normed_add_comm_group E] [second_countable_topology_either α E] variables {p : ℝ≥0∞} [_i : fact (1 ≤ p)] (hp : p ≠ ∞) (μ : measure α) @@ -154,7 +154,7 @@ begin 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_group _ gc_cont (∥c∥) _, rfl⟩, + 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] }, diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index 0fcdd9622af26..060a8a5bb4a54 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -49,7 +49,7 @@ def tendsto_in_measure [has_dist E] {m : measurable_space α} (μ : measure α) (f : ι → α → E) (l : filter ι) (g : α → E) : Prop := ∀ ε (hε : 0 < ε), tendsto (λ i, μ {x | ε ≤ dist (f i x) (g x)}) l (𝓝 0) -lemma tendsto_in_measure_iff_norm [semi_normed_group E] {l : filter ι} +lemma tendsto_in_measure_iff_norm [seminormed_add_comm_group E] {l : filter ι} {f : ι → α → E} {g : α → E} : tendsto_in_measure μ f l g ↔ ∀ ε (hε : 0 < ε), tendsto (λ i, μ {x | ε ≤ ∥f i x - g x∥}) l (𝓝 0) := @@ -259,7 +259,7 @@ end exists_seq_tendsto_ae section ae_measurable_of -variables [measurable_space E] [normed_group E] [borel_space E] +variables [measurable_space E] [normed_add_comm_group E] [borel_space E] lemma tendsto_in_measure.ae_measurable {u : filter ι} [ne_bot u] [is_countably_generated u] @@ -275,7 +275,7 @@ end ae_measurable_of section tendsto_in_measure_of -variables [normed_group E] {p : ℝ≥0∞} +variables [normed_add_comm_group E] {p : ℝ≥0∞} variables {f : ι → α → E} {g : α → E} /-- This lemma is superceded by `measure_theory.tendsto_in_measure_of_tendsto_snorm` where we @@ -326,8 +326,8 @@ end /-- See also `measure_theory.tendsto_in_measure_of_tendsto_snorm` which work for general Lp-convergence for all `p ≠ 0`. -/ -lemma tendsto_in_measure_of_tendsto_snorm_top {E} [normed_group E] {f : ι → α → E} {g : α → E} - {l : filter ι} (hfg : tendsto (λ n, snorm (f n - g) ∞ μ) l (𝓝 0)) : +lemma tendsto_in_measure_of_tendsto_snorm_top {E} [normed_add_comm_group E] {f : ι → α → E} + {g : α → E} {l : filter ι} (hfg : tendsto (λ n, snorm (f n - g) ∞ μ) l (𝓝 0)) : tendsto_in_measure μ f l g := begin intros δ hδ, diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 5c801a92066b2..c36c5707b59b6 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -89,8 +89,8 @@ open measure_theory measure_theory.measure metric filter set finite_dimensional topological_space open_locale nnreal ennreal topological_space pointwise -variables {E F : Type*} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] -[normed_group F] [normed_space ℝ F] {s : set E} {f : E → E} {f' : E → E →L[ℝ] E} +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} /-! ### Decomposition lemmas diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index bd0ba3038840b..a9fce93c233dc 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -20,15 +20,15 @@ classes of integrable functions, already defined as a special case of `L^p` spac ## Notation -* `α →₁[μ] β` is the type of `L¹` space, where `α` is a `measure_space` and `β` is a `normed_group` - with a `second_countable_topology`. `f : α →ₘ β` is a "function" in `L¹`. In comments, `[f]` is - also used to denote an `L¹` function. +* `α →₁[μ] β` is the type of `L¹` space, where `α` is a `measure_space` and `β` is a + `normed_add_comm_group` with a `second_countable_topology`. `f : α →ₘ β` is a "function" in `L¹`. + In comments, `[f]` is also used to denote an `L¹` function. `₁` can be typed as `\1`. ## Main definitions -* Let `f : α → β` be a function, where `α` is a `measure_space` and `β` a `normed_group`. +* Let `f : α → β` be a function, where `α` is a `measure_space` and `β` a `normed_add_comm_group`. Then `has_finite_integral f` means `(∫⁻ a, ∥f a∥₊) < ∞`. * If `β` is moreover a `measurable_space` then `f` is called `integrable` if @@ -52,8 +52,8 @@ open_locale classical topological_space big_operators ennreal measure_theory nnr open set filter topological_space ennreal emetric measure_theory variables {α β γ δ : Type*} {m : measurable_space α} {μ ν : measure α} [measurable_space δ] -variables [normed_group β] -variables [normed_group γ] +variables [normed_add_comm_group β] +variables [normed_add_comm_group γ] namespace measure_theory @@ -707,7 +707,7 @@ begin end section -variables {E : Type*} [normed_group E] [normed_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] lemma integrable_with_density_iff_integrable_coe_smul {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : @@ -778,7 +778,7 @@ begin end section -variables {E : Type*} [normed_group E] [normed_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] lemma mem_ℒ1_smul_of_L1_with_density {f : α → ℝ≥0} (f_meas : measurable f) (u : Lp E 1 (μ.with_density (λ x, f x))) : @@ -907,7 +907,7 @@ end normed_space section normed_space_over_complete_field variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] +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 μ := @@ -955,7 +955,8 @@ end inner_product section trim -variables {H : Type*} [normed_group H] {m0 : measurable_space α} {μ' : measure α} {f : α → H} +variables {H : Type*} [normed_add_comm_group H] {m0 : measurable_space α} {μ' : measure α} + {f : α → H} lemma integrable.trim (hm : m ≤ m0) (hf_int : integrable f μ') (hf : strongly_measurable[m] f) : integrable f (μ'.trim hm) := @@ -980,7 +981,7 @@ end trim section sigma_finite -variables {E : Type*} {m0 : measurable_space α} [normed_group E] +variables {E : Type*} {m0 : measurable_space α} [normed_add_comm_group E] lemma integrable_of_forall_fin_meas_le' {μ : measure α} (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : ae_strongly_measurable f μ) @@ -1186,9 +1187,9 @@ end measure_theory open measure_theory -variables {E : Type*} [normed_group E] +variables {E : Type*} [normed_add_comm_group E] {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] - {H : Type*} [normed_group H] [normed_space 𝕜 H] + {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] lemma measure_theory.integrable.apply_continuous_linear_map {φ : α → H →L[𝕜] E} (φ_int : integrable φ μ) (v : H) : integrable (λ a, φ a v) μ := diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index c9e52cd5a9fc9..b3a7ca2d8f3b5 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -29,7 +29,7 @@ namespace measure_theory section -variables {α F : Type*} {m : measurable_space α} {μ : measure α} [normed_group F] +variables {α F : Type*} {m : measurable_space α} {μ : measure α} [normed_add_comm_group F] lemma mem_ℒp.integrable_sq {f : α → ℝ} (h : mem_ℒp f 2 μ) : integrable (λ x, (f x)^2) μ := @@ -58,7 +58,7 @@ end namespace L2 variables {α E F 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] {μ : measure α} - [inner_product_space 𝕜 E] [normed_group F] + [inner_product_space 𝕜 E] [normed_add_comm_group F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index e8d07e8aeebf3..68bd275ea2647 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -25,7 +25,7 @@ open_locale topological_space interval variables {X Y E : Type*} [measurable_space X] [topological_space X] variables [measurable_space Y] [topological_space Y] -variables [normed_group E] {f : X → E} {μ : measure X} +variables [normed_add_comm_group E] {f : X → E} {μ : measure X} namespace measure_theory diff --git a/src/measure_theory/function/lp_order.lean b/src/measure_theory/function/lp_order.lean index 733d44f3a135b..1c2719ccf8fe5 100644 --- a/src/measure_theory/function/lp_order.lean +++ b/src/measure_theory/function/lp_order.lean @@ -96,7 +96,7 @@ instance [fact (1 ≤ p)] : normed_lattice_add_comm_group (Lp E p μ) := rw [hxf, hxg] at hx, exact solid hx, end, - ..Lp.lattice, ..Lp.normed_group, } + ..Lp.lattice, ..Lp.normed_add_comm_group, } end order diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index f567e377dd856..adf23f5ebe61e 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -78,7 +78,7 @@ open topological_space measure_theory filter open_locale nnreal ennreal big_operators topological_space measure_theory variables {α E F G : Type*} {m m0 : measurable_space α} {p : ℝ≥0∞} {q : ℝ} {μ ν : measure α} - [normed_group E] [normed_group F] [normed_group G] + [normed_add_comm_group E] [normed_add_comm_group F] [normed_add_comm_group G] namespace measure_theory @@ -1329,17 +1329,17 @@ The space of equivalence classes of measurable functions for which `snorm f p μ -/ @[simp] lemma snorm_ae_eq_fun {α E : Type*} [measurable_space α] {μ : measure α} - [normed_group E] {p : ℝ≥0∞} {f : α → E} (hf : ae_strongly_measurable f μ) : + [normed_add_comm_group E] {p : ℝ≥0∞} {f : α → E} (hf : ae_strongly_measurable f μ) : snorm (ae_eq_fun.mk f hf) p μ = snorm f p μ := snorm_congr_ae (ae_eq_fun.coe_fn_mk _ _) lemma mem_ℒp.snorm_mk_lt_top {α E : Type*} [measurable_space α] {μ : measure α} - [normed_group E] {p : ℝ≥0∞} {f : α → E} (hfp : mem_ℒp f p μ) : + [normed_add_comm_group E] {p : ℝ≥0∞} {f : α → E} (hfp : mem_ℒp f p μ) : snorm (ae_eq_fun.mk f hfp.1) p μ < ∞ := by simp [hfp.2] /-- Lp space -/ -def Lp {α} (E : Type*) {m : measurable_space α} [normed_group E] +def Lp {α} (E : Type*) {m : measurable_space α} [normed_add_comm_group E] (p : ℝ≥0∞) (μ : measure α . volume_tac) : add_subgroup (α →ₘ[μ] E) := { carrier := {f | snorm f p μ < ∞}, zero_mem' := by simp [snorm_congr_ae ae_eq_fun.coe_fn_zero, snorm_zero], @@ -1556,12 +1556,12 @@ begin rw max_eq_left hC end -instance [hp : fact (1 ≤ p)] : normed_group (Lp E p μ) := +instance [hp : fact (1 ≤ p)] : normed_add_comm_group (Lp E p μ) := { edist := edist, edist_dist := λ f g, by rw [edist_def, dist_def, ←snorm_congr_ae (coe_fn_sub _ _), ennreal.of_real_to_real (snorm_ne_top (f - g))], - .. normed_group.of_core (Lp E p μ) + .. normed_add_comm_group.of_core (Lp E p μ) { norm_eq_zero_iff := λ f, norm_eq_zero_iff (ennreal.zero_lt_one.trans_le hp.1), triangle := begin assume f g, @@ -1669,7 +1669,7 @@ end variables (hs) -lemma snorm_indicator_le {E : Type*} [normed_group E] (f : α → E) : +lemma snorm_indicator_le {E : Type*} [normed_add_comm_group E] (f : α → E) : snorm (s.indicator f) p μ ≤ snorm f p μ := begin refine snorm_mono_ae (eventually_of_forall (λ x, _)), @@ -1898,8 +1898,8 @@ section composition variables {g : E → F} {c : ℝ≥0} lemma lipschitz_with.comp_mem_ℒp {α E F} {K} [measurable_space α] {μ : measure α} - [normed_group E] [normed_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 μ := + [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, _), @@ -1909,7 +1909,7 @@ begin end lemma measure_theory.mem_ℒp.of_comp_antilipschitz_with {α E F} {K'} - [measurable_space α] {μ : measure α} [normed_group E] [normed_group F] + [measurable_space α] {μ : measure α} [normed_add_comm_group E] [normed_add_comm_group F] {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 @@ -1925,7 +1925,7 @@ end namespace lipschitz_with lemma mem_ℒp_comp_iff_of_antilipschitz {α E F} {K K'} [measurable_space α] {μ : measure α} - [normed_group E] [normed_group F] + [normed_add_comm_group E] [normed_add_comm_group F] {f : α → E} {g : E → F} (hg : lipschitz_with K g) (hg' : antilipschitz_with K' g) (g0 : g 0 = 0) : mem_ℒp (g ∘ f) p μ ↔ mem_ℒp f p μ := ⟨λ h, h.of_comp_antilipschitz_with hg.uniform_continuous hg' g0, λ h, hg.comp_mem_ℒp g0 h⟩ @@ -2207,8 +2207,8 @@ begin exact (continuous_nnnorm.tendsto (f_lim a)).comp ha, end -lemma snorm'_lim_le_liminf_snorm' {E} [normed_group E] {f : ℕ → α → E} {p : ℝ} (hp_pos : 0 < p) - (hf : ∀ n, ae_strongly_measurable (f n) μ) {f_lim : α → E} +lemma snorm'_lim_le_liminf_snorm' {E} [normed_add_comm_group E] {f : ℕ → α → E} {p : ℝ} + (hp_pos : 0 < p) (hf : ∀ n, ae_strongly_measurable (f n) μ) {f_lim : α → E} (h_lim : ∀ᵐ (x : α) ∂μ, tendsto (λ n, f n x) at_top (𝓝 (f_lim x))) : snorm' f_lim p μ ≤ at_top.liminf (λ n, snorm' (f n) p μ) := begin @@ -2247,7 +2247,7 @@ begin exact ennreal.ess_sup_liminf_le (λ n, (λ x, (∥f n x∥₊ : ℝ≥0∞))), end -lemma snorm_lim_le_liminf_snorm {E} [normed_group E] +lemma snorm_lim_le_liminf_snorm {E} [normed_add_comm_group E] {f : ℕ → α → E} (hf : ∀ n, ae_strongly_measurable (f n) μ) (f_lim : α → E) (h_lim : ∀ᵐ (x : α) ∂μ, tendsto (λ n, f n x) at_top (𝓝 (f_lim x))) : snorm f_lim p μ ≤ at_top.liminf (λ n, snorm (f n) p μ) := @@ -2642,7 +2642,7 @@ variables (p μ) /-- The normed group homomorphism of considering a bounded continuous function on a finite-measure space as an element of `Lp`. -/ -def to_Lp_hom [fact (1 ≤ p)] : normed_group_hom (α →ᵇ E) (Lp E p μ) := +def to_Lp_hom [fact (1 ≤ p)] : normed_add_group_hom (α →ᵇ E) (Lp E p μ) := { bound' := ⟨_, Lp_norm_le⟩, .. add_monoid_hom.cod_restrict ((continuous_map.to_ae_eq_fun_add_hom μ).comp (to_continuous_map_add_hom α E)) diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index afefae1a8803e..93bf695847081 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -19,9 +19,10 @@ by a sequence of simple functions. ## Main results -* `tendsto_approx_on_univ_Lp` (Lᵖ convergence): If `E` is a `normed_group` and `f` is measurable - and `mem_ℒp` (for `p < ∞`), then the simple functions `simple_func.approx_on f hf s 0 h₀ n` may - be considered as elements of `Lp E p μ`, and they tend in Lᵖ to `f`. +* `tendsto_approx_on_univ_Lp` (Lᵖ convergence): If `E` is a `normed_add_comm_group` and `f` is + measurable and `mem_ℒp` (for `p < ∞`), then the simple functions + `simple_func.approx_on f hf s 0 h₀ n` may be considered as elements of `Lp E p μ`, and they tend + in Lᵖ to `f`. * `Lp.simple_func.dense_embedding`: the embedding `coe_to_Lp` of the `Lp` simple functions into `Lp` is dense. * `Lp.simple_func.induction`, `Lp.induction`, `mem_ℒp.induction`, `integrable.induction`: to prove @@ -52,8 +53,8 @@ namespace simple_func /-! ### Lp approximation by simple functions -/ section Lp -variables [measurable_space β] -variables [measurable_space E] [normed_group E] [normed_group F] {q : ℝ} {p : ℝ≥0∞} +variables [measurable_space β] [measurable_space E] [normed_add_comm_group E] + [normed_add_comm_group F] {q : ℝ} {p : ℝ≥0∞} lemma nnnorm_approx_on_le [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} {y₀ : E} (h₀ : y₀ ∈ s) [separable_space s] (x : β) (n : ℕ) : @@ -191,7 +192,7 @@ end Lp section integrable variables [measurable_space β] -variables [measurable_space E] [normed_group E] +variables [measurable_space E] [normed_add_comm_group E] lemma tendsto_approx_on_L1_nnnorm [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} {y₀ : E} (h₀ : y₀ ∈ s) [separable_space s] @@ -235,7 +236,7 @@ end integrable section simple_func_properties variables [measurable_space α] -variables [normed_group E] [normed_group F] +variables [normed_add_comm_group E] [normed_add_comm_group F] variables {μ : measure α} {p : ℝ≥0∞} /-! @@ -383,7 +384,8 @@ namespace Lp open ae_eq_fun -variables [measurable_space α] [normed_group E] [normed_group F] (p : ℝ≥0∞) (μ : measure α) +variables [measurable_space α] [normed_add_comm_group E] [normed_add_comm_group F] (p : ℝ≥0∞) + (μ : measure α) variables (E) @@ -829,7 +831,7 @@ end simple_func end Lp -variables [measurable_space α] [normed_group E] {f : α → E} {p : ℝ≥0∞} {μ : measure α} +variables [measurable_space α] [normed_add_comm_group E] {f : α → E} {p : ℝ≥0∞} {μ : measure α} /-- To prove something for an arbitrary `Lp` function in a second countable Borel normed group, it suffices to show that diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index df026a9299c01..05e68ddec5413 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -694,17 +694,17 @@ protected lemma dist {m : measurable_space α} {β : Type*} [pseudo_metric_space strongly_measurable (λ x, dist (f x) (g x)) := continuous_dist.comp_strongly_measurable (hf.prod_mk hg) -protected lemma norm {m : measurable_space α} {β : Type*} [normed_group β] {f : α → β} +protected lemma norm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : strongly_measurable (λ x, ∥f x∥) := continuous_norm.comp_strongly_measurable hf -protected lemma nnnorm {m : measurable_space α} {β : Type*} [normed_group β] {f : α → β} +protected lemma nnnorm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : strongly_measurable (λ x, ∥f x∥₊) := continuous_nnnorm.comp_strongly_measurable hf -protected lemma ennnorm {m : measurable_space α} {β : Type*} [normed_group β] {f : α → β} +protected lemma ennnorm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := (ennreal.continuous_coe.comp_strongly_measurable hf.nnnorm).measurable @@ -1266,19 +1266,22 @@ protected lemma dist {β : Type*} [pseudo_metric_space β] {f g : α → β} ae_strongly_measurable (λ x, dist (f x) (g x)) μ := continuous_dist.comp_ae_strongly_measurable (hf.prod_mk hg) -protected lemma norm {β : Type*} [normed_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : +protected lemma norm {β : Type*} [normed_add_comm_group β] {f : α → β} + (hf : ae_strongly_measurable f μ) : ae_strongly_measurable (λ x, ∥f x∥) μ := continuous_norm.comp_ae_strongly_measurable hf -protected lemma nnnorm {β : Type*} [normed_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : +protected lemma nnnorm {β : Type*} [normed_add_comm_group β] {f : α → β} + (hf : ae_strongly_measurable f μ) : ae_strongly_measurable (λ x, ∥f x∥₊) μ := continuous_nnnorm.comp_ae_strongly_measurable hf -protected lemma ennnorm {β : Type*} [normed_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : +protected lemma ennnorm {β : Type*} [normed_add_comm_group β] {f : α → β} + (hf : ae_strongly_measurable f μ) : ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := (ennreal.continuous_coe.comp_ae_strongly_measurable hf.nnnorm).ae_measurable -protected lemma edist {β : Type*} [normed_group β] {f g : α → β} +protected lemma edist {β : Type*} [normed_add_comm_group β] {f g : α → β} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : ae_measurable (λ a, edist (f a) (g a)) μ := (continuous_edist.comp_ae_strongly_measurable (hf.prod_mk hg)).ae_measurable @@ -1512,7 +1515,7 @@ lemma smul_measure {R : Type*} [monoid R] [distrib_mul_action R ℝ≥0∞] section normed_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] lemma _root_.ae_strongly_measurable_smul_const_iff {f : α → 𝕜} {c : E} (hc : c ≠ 0) : ae_strongly_measurable (λ x, f x • c) μ ↔ ae_strongly_measurable f μ := @@ -1545,9 +1548,9 @@ end mul_action section continuous_linear_map_nontrivially_normed_field variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] -variables {F : Type*} [normed_group F] [normed_space 𝕜 F] -variables {G : Type*} [normed_group G] [normed_space 𝕜 G] +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] lemma _root_.strongly_measurable.apply_continuous_linear_map {m : measurable_space α} {φ : α → F →L[𝕜] E} (hφ : strongly_measurable φ) (v : F) : @@ -1567,8 +1570,8 @@ L.continuous₂.comp_ae_strongly_measurable $ hf.prod_mk hg end continuous_linear_map_nontrivially_normed_field -lemma _root_.ae_strongly_measurable_with_density_iff {E : Type*} [normed_group E] [normed_space ℝ E] - {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : +lemma _root_.ae_strongly_measurable_with_density_iff {E : Type*} [normed_add_comm_group E] + [normed_space ℝ E] {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : ae_strongly_measurable g (μ.with_density (λ x, (f x : ℝ≥0∞))) ↔ ae_strongly_measurable (λ x, (f x : ℝ) • g x) μ := begin @@ -1709,7 +1712,7 @@ end ae_fin_strongly_measurable section second_countable_topology variables {G : Type*} {p : ℝ≥0∞} {m m0 : measurable_space α} {μ : measure α} - [normed_group G] [measurable_space G] [borel_space G] [second_countable_topology G] + [normed_add_comm_group G] [measurable_space G] [borel_space G] [second_countable_topology G] {f : α → G} /-- In a space with second countable topology and a sigma-finite measure, `fin_strongly_measurable` diff --git a/src/measure_theory/function/strongly_measurable_lp.lean b/src/measure_theory/function/strongly_measurable_lp.lean index 571c39a8569c0..44f3b46bf86ca 100644 --- a/src/measure_theory/function/strongly_measurable_lp.lean +++ b/src/measure_theory/function/strongly_measurable_lp.lean @@ -32,7 +32,7 @@ namespace measure_theory local infixr ` →ₛ `:25 := simple_func variables {α G : Type*} {p : ℝ≥0∞} {m m0 : measurable_space α} {μ : measure α} - [normed_group G] + [normed_add_comm_group G] {f : α → G} lemma mem_ℒp.fin_strongly_measurable_of_strongly_measurable diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index bc774fd28acf6..ff7dac8012669 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -48,7 +48,7 @@ namespace measure_theory open set filter topological_space -variables {α β ι : Type*} {m : measurable_space α} {μ : measure α} [normed_group β] +variables {α β ι : Type*} {m : measurable_space α} {μ : measure α} [normed_add_comm_group β] /-- Uniform integrability in the measure theory sense. diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index 5c6d9641c328d..e67cab39d1f8f 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -54,7 +54,7 @@ structure is_fundamental_domain (G : Type*) {α : Type*} [has_one G] [has_smul G namespace is_fundamental_domain variables {G α E : Type*} [group G] [mul_action G α] [measurable_space α] - [normed_group E] {s t : set α} {μ : measure α} + [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 `α`. -/ diff --git a/src/measure_theory/group/integration.lean b/src/measure_theory/group/integration.lean index 1143a97806c65..fee69426abdf3 100644 --- a/src/measure_theory/group/integration.lean +++ b/src/measure_theory/group/integration.lean @@ -20,7 +20,7 @@ open measure topological_space open_locale ennreal variables {𝕜 M α G E F : Type*} [measurable_space G] -variables [normed_group E] [normed_space ℝ E] [complete_space E] [normed_group F] +variables [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] [normed_add_comm_group F] variables {μ : measure G} {f : G → E} {g : G} section measurable_inv diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index 1946376e3d5ab..16e6d6bfd37c1 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -575,8 +575,9 @@ 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_group E] [normed_space ℝ E] [nontrivial E] [finite_dimensional ℝ E] - [measurable_space E] [borel_space E] (μ : measure E) [is_add_haar_measure μ] : +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 diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index 7bb6581827581..1e61e65b5d9cd 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -31,8 +31,8 @@ open measure_theory measure_theory.measure metric set filter topological_space f open_locale topological_space big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} - [normed_group E] [normed_space ℝ E] [complete_space E] - [normed_group F] [normed_space ℝ F] [complete_space F] + [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} /-! diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 8cde6c4a3b66d..76badaf315264 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -153,7 +153,7 @@ section weighted_smul open continuous_linear_map -variables [normed_group F] [normed_space ℝ F] {m : measurable_space α} {μ : measure α} +variables [normed_add_comm_group F] [normed_space ℝ F] {m : measurable_space α} {μ : measure α} /-- Given a set `s`, return the continuous linear map `λ x, (μ s).to_real • x`. The extension of that set function through `set_to_L1` gives the Bochner integral of L1 functions. -/ @@ -278,8 +278,8 @@ and prove basic property of this integral. -/ open finset -variables [normed_group E] [normed_group F] [normed_space ℝ F] {p : ℝ≥0∞} - {G F' : Type*} [normed_group G] [normed_group F'] [normed_space ℝ F'] +variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_space ℝ F] {p : ℝ≥0∞} + {G F' : Type*} [normed_add_comm_group G] [normed_add_comm_group F'] [normed_space ℝ F'] {m : measurable_space α} {μ : measure α} /-- Bochner integral of simple functions whose codomain is a real `normed_space`. @@ -424,7 +424,8 @@ namespace L1 open ae_eq_fun Lp.simple_func Lp -variables [normed_group E] [normed_group F] {m : measurable_space α} {μ : measure α} +variables [normed_add_comm_group E] [normed_add_comm_group F] {m : measurable_space α} + {μ : measure α} variables {α E μ} @@ -466,7 +467,7 @@ Define the Bochner integral on `α →₁ₛ[μ] E` by extension from the simple and prove basic properties of this integral. -/ variables [normed_field 𝕜] [normed_space 𝕜 E] [normed_space ℝ E] [smul_comm_class ℝ 𝕜 E] - {F' : Type*} [normed_group F'] [normed_space ℝ F'] + {F' : Type*} [normed_add_comm_group F'] [normed_space ℝ F'] local attribute [instance] simple_func.normed_space @@ -498,7 +499,7 @@ begin exact (to_simple_func f).norm_integral_le_integral_norm (simple_func.integrable f) end -variables {E' : Type*} [normed_group E'] [normed_space ℝ E'] [normed_space 𝕜 E'] +variables {E' : Type*} [normed_add_comm_group E'] [normed_space ℝ E'] [normed_space 𝕜 E'] variables (α E μ 𝕜) @@ -686,9 +687,9 @@ functions, and 0 otherwise; prove its basic properties. -/ -variables [normed_group E] [normed_space ℝ E] [complete_space E] +variables [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] - [normed_group F] [normed_space ℝ F] [complete_space F] + [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] section open_locale classical @@ -969,7 +970,7 @@ begin rw [this, hfi], refl } end -lemma integral_norm_eq_lintegral_nnnorm {G} [normed_group G] +lemma integral_norm_eq_lintegral_nnnorm {G} [normed_add_comm_group G] {f : α → G} (hf : ae_strongly_measurable f μ) : ∫ x, ∥f x∥ ∂μ = ennreal.to_real ∫⁻ x, ∥f x∥₊ ∂μ := begin @@ -978,7 +979,7 @@ begin { refine ae_of_all _ _, simp_rw [pi.zero_apply, norm_nonneg, imp_true_iff] }, end -lemma of_real_integral_norm_eq_lintegral_nnnorm {G} [normed_group G] {f : α → G} +lemma of_real_integral_norm_eq_lintegral_nnnorm {G} [normed_add_comm_group G] {f : α → G} (hf : integrable f μ) : ennreal.of_real ∫ x, ∥f x∥ ∂μ = ∫⁻ x, ∥f x∥₊ ∂μ := by rw [integral_norm_eq_lintegral_nnnorm hf.ae_strongly_measurable, @@ -1074,8 +1075,8 @@ lemma integral_pos_iff_support_of_nonneg {f : α → ℝ} (hf : 0 ≤ f) (hfi : (0 < ∫ x, f x ∂μ) ↔ 0 < μ (function.support f) := integral_pos_iff_support_of_nonneg_ae (eventually_of_forall hf) hfi -section normed_group -variables {H : Type*} [normed_group H] +section normed_add_comm_group +variables {H : Type*} [normed_add_comm_group H] lemma L1.norm_eq_integral_norm (f : α →₁[μ] H) : ∥f∥ = ∫ a, ∥f a∥ ∂μ := begin @@ -1111,7 +1112,7 @@ begin exact (lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp1 hp2 hf.2).ne end -end normed_group +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 ∂μ := @@ -1431,7 +1432,7 @@ attribute [irreducible] integral L1.integral section integral_trim -variables {H β γ : Type*} [normed_group H] +variables {H β γ : Type*} [normed_add_comm_group H] {m m0 : measurable_space β} {μ : measure β} /-- Simple function seen as simple function of a larger `measurable_space`. -/ diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index b55bee6f504c1..111cfaef670a9 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -65,7 +65,7 @@ some lemmas use, e.g., `(z - c)⁻¹ • f z` instead of `f z / (z - c)`. integral, circle, Cauchy integral -/ -variables {E : Type*} [normed_group E] +variables {E : Type*} [normed_add_comm_group E] noncomputable theory diff --git a/src/measure_theory/integral/circle_integral_transform.lean b/src/measure_theory/integral/circle_integral_transform.lean index 7dbb237f5f6db..81ce0c737c03b 100644 --- a/src/measure_theory/integral/circle_integral_transform.lean +++ b/src/measure_theory/integral/circle_integral_transform.lean @@ -25,7 +25,7 @@ open_locale interval real noncomputable theory -variables {E : Type} [normed_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/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index da7069c3afe83..d931f9c9e463f 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -53,7 +53,7 @@ universes u namespace measure_theory -variables {E : Type u} [normed_group E] [normed_space ℝ E] [complete_space E] +variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] section variables {n : ℕ} @@ -296,7 +296,7 @@ end /-- An auxiliary lemma that is used to specialize the general divergence theorem to spaces that do not have the form `fin n → ℝ`. -/ lemma integral_divergence_of_has_fderiv_within_at_off_countable_of_equiv - {F : Type*} [normed_group F] [normed_space ℝ F] [partial_order F] [measure_space F] + {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [partial_order F] [measure_space F] [borel_space F] (eL : F ≃L[ℝ] ℝⁿ⁺¹) (he_ord : ∀ x y, eL x ≤ eL y ↔ x ≤ y) (he_vol : measure_preserving eL volume volume) (f : fin (n + 1) → F → E) (f' : fin (n + 1) → F → F →L[ℝ] E) (s : set F) (hs : s.countable) diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index c36a42b963b08..f7c9c136d28ed 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -65,15 +65,15 @@ end namespace measure_theory -section normed_group +section normed_add_comm_group -lemma has_finite_integral_restrict_of_bounded [normed_group E] {f : α → E} {s : set α} +lemma has_finite_integral_restrict_of_bounded [normed_add_comm_group E] {f : α → E} {s : set α} {μ : measure α} {C} (hs : μ s < ∞) (hf : ∀ᵐ x ∂(μ.restrict s), ∥f x∥ ≤ C) : has_finite_integral f (μ.restrict s) := by haveI : is_finite_measure (μ.restrict s) := ⟨by rwa [measure.restrict_apply_univ]⟩; exact has_finite_integral_of_bounded hf -variables [normed_group E] {f g : α → E} {s t : set α} {μ ν : measure α} +variables [normed_add_comm_group E] {f g : α → E} {s t : set α} {μ ν : measure α} /-- A function is `integrable_on` a set `s` if it is almost everywhere strongly measurable on `s` and if the integral of its pointwise norm over `s` is less than infinity. -/ @@ -216,7 +216,7 @@ lemma integrable.indicator (h : integrable f μ) (hs : measurable_set s) : integrable (indicator s f) μ := h.integrable_on.indicator hs -lemma integrable_indicator_const_Lp {E} [normed_group E] +lemma integrable_indicator_const_Lp {E} [normed_add_comm_group E] {p : ℝ≥0∞} {s : set α} (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : E) : integrable (indicator_const_Lp p hs hμs c) μ := begin @@ -234,7 +234,7 @@ begin rwa [← indicator_eq_self.2 h1s, integrable_indicator_iff h2s] end -lemma integrable_on_Lp_of_measure_ne_top {E} [normed_group E] +lemma integrable_on_Lp_of_measure_ne_top {E} [normed_add_comm_group E] {p : ℝ≥0∞} {s : set α} (f : Lp E p μ) (hp : 1 ≤ p) (hμs : μ s ≠ ∞) : integrable_on f s μ := begin @@ -327,13 +327,13 @@ begin { rw ← indicator_add_eq_right h, exact hfg.indicator hg.measurable_set_support } end -end normed_group +end normed_add_comm_group end measure_theory open measure_theory -variables [normed_group E] +variables [normed_add_comm_group E] /-- A function which is continuous on a set `s` is almost everywhere measurable with respect to `μ.restrict s`. -/ diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index ed7b5adcea130..bb5663185543c 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -385,7 +385,7 @@ end lintegral section integrable variables {α ι E : Type*} [measurable_space α] {μ : measure α} {l : filter ι} - [normed_group E] + [normed_add_comm_group E] lemma ae_cover.integrable_of_lintegral_nnnorm_bounded [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ) (hfm : ae_strongly_measurable f μ) @@ -469,7 +469,7 @@ end integrable section integral variables {α ι E : Type*} [measurable_space α] {μ : measure α} {l : filter ι} - [normed_group E] [normed_space ℝ E] [complete_space E] + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] lemma ae_cover.integral_tendsto_of_countably_generated [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (hfi : integrable f μ) : @@ -505,7 +505,7 @@ section integrable_of_interval_integral variables {ι E : Type*} {μ : measure ℝ} {l : filter ι} [filter.ne_bot l] [is_countably_generated l] - [normed_group E] + [normed_add_comm_group E] {a b : ι → ℝ} {f : ℝ → E} lemma integrable_of_interval_integral_norm_bounded @@ -605,7 +605,7 @@ section integral_of_interval_integral variables {ι E : Type*} {μ : measure ℝ} {l : filter ι} [is_countably_generated l] - [normed_group E] [normed_space ℝ E] [complete_space E] + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] {a b : ι → ℝ} {f : ℝ → E} lemma interval_integral_tendsto_integral diff --git a/src/measure_theory/integral/interval_average.lean b/src/measure_theory/integral/interval_average.lean index 450d4fe78d1a0..635fb91aefdba 100644 --- a/src/measure_theory/integral/interval_average.lean +++ b/src/measure_theory/integral/interval_average.lean @@ -27,7 +27,7 @@ We also prove that `⨍ x in a..b, f x = ⨍ x in b..a, f x`, see `interval_aver open measure_theory set topological_space open_locale interval -variables {E : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] notation `⨍` binders ` in ` a `..` b `, ` r:(scoped:60 f, average (measure.restrict volume (Ι a b)) f) := r diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 657ea7c50ef5e..ee45f7dcc4c88 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -169,7 +169,7 @@ open measure_theory set classical filter function open_locale classical topological_space filter ennreal big_operators interval nnreal -variables {ι 𝕜 E F : Type*} [normed_group E] +variables {ι 𝕜 E F : Type*} [normed_add_comm_group E] /-! ### Integrability at an interval @@ -202,7 +202,7 @@ lemma interval_integrable_iff_integrable_Ioc_of_le (hab : a ≤ b) : by rw [interval_integrable_iff, interval_oc_of_le hab] lemma integrable_on_Icc_iff_integrable_on_Ioc' - {E : Type*} [normed_group E] {f : ℝ → E} (ha : μ {a} ≠ ∞) : + {E : Type*} [normed_add_comm_group E] {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, @@ -213,12 +213,12 @@ begin end lemma integrable_on_Icc_iff_integrable_on_Ioc - {E : Type*}[normed_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : + {E : Type*}[normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := integrable_on_Icc_iff_integrable_on_Ioc' (by simp) lemma integrable_on_Ioc_iff_integrable_on_Ioo' - {E : Type*} [normed_group E] + {E : Type*} [normed_add_comm_group E] {f : ℝ → E} {a b : ℝ} (hb : μ {b} ≠ ∞) : integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := begin @@ -230,12 +230,12 @@ begin end lemma integrable_on_Ioc_iff_integrable_on_Ioo - {E : Type*} [normed_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : + {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := integrable_on_Ioc_iff_integrable_on_Ioo' (by simp) lemma integrable_on_Icc_iff_integrable_on_Ioo - {E : Type*} [normed_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : + {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioo a b) μ := by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable_on_Ioo] @@ -243,13 +243,13 @@ lemma interval_integrable_iff' [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (interval a b) μ := by rw [interval_integrable_iff, interval, interval_oc, integrable_on_Icc_iff_integrable_on_Ioc] -lemma interval_integrable_iff_integrable_Icc_of_le {E : Type*} [normed_group E] +lemma interval_integrable_iff_integrable_Icc_of_le {E : Type*} [normed_add_comm_group E] {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) {μ : measure ℝ} [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (Icc a b) μ := by rw [interval_integrable_iff_integrable_Ioc_of_le hab, integrable_on_Icc_iff_integrable_on_Ioc] lemma integrable_on_Ici_iff_integrable_on_Ioi' - {E : Type*} [normed_group E] {f : ℝ → E} (ha : μ {a} ≠ ∞) : + {E : Type*} [normed_add_comm_group E] {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, @@ -258,7 +258,7 @@ begin end lemma integrable_on_Ici_iff_integrable_on_Ioi - {E : Type*} [normed_group E] [has_no_atoms μ] {f : ℝ → E} : + {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} : integrable_on f (Ici a) μ ↔ integrable_on f (Ioi a) μ := integrable_on_Ici_iff_integrable_on_Ioi' (by simp) @@ -343,7 +343,7 @@ lemma mono_set_ae (hf : interval_integrable f μ a b) (h : Ι c d ≤ᵐ[μ] Ι interval_integrable f μ c d := interval_integrable_iff.mpr $ hf.def.mono_set_ae h -lemma mono_fun [normed_group F] {g : ℝ → F} +lemma mono_fun [normed_add_comm_group F] {g : ℝ → F} (hf : interval_integrable f μ a b) (hgm : ae_strongly_measurable g (μ.restrict (Ι a b))) (hle : (λ x, ∥g x∥) ≤ᵐ[μ.restrict (Ι a b)] (λ x, ∥f x∥)) : interval_integrable g μ a b := interval_integrable_iff.2 $ hf.def.integrable.mono hgm hle @@ -661,7 +661,7 @@ lemma integral_smul_measure (c : ℝ≥0∞) : ∫ x in a..b, f x ∂(c • μ) = c.to_real • ∫ x in a..b, f x ∂μ := by simp only [interval_integral, measure.restrict_smul, integral_smul_measure, smul_sub] -variables [normed_group F] [complete_space F] [normed_space ℝ F] +variables [normed_add_comm_group F] [complete_space F] [normed_space ℝ F] lemma _root_.continuous_linear_map.interval_integral_comp_comm (L : E →L[ℝ] F) (hf : interval_integrable f μ a b) : diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 5779306ddc4d6..c44276bbd9a16 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -2434,7 +2434,7 @@ begin { exact hf (measurable_set_singleton 0).compl }, end -lemma ae_measurable_with_density_iff {E : Type*} [normed_group E] [normed_space ℝ E] +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) μ := @@ -2731,7 +2731,7 @@ by rw [lintegral_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk), section sigma_finite -variables {E : Type*} [normed_group E] [measurable_space E] +variables {E : Type*} [normed_add_comm_group E] [measurable_space E] [opens_measurable_space E] lemma univ_le_of_forall_fin_meas_le {μ : measure α} (hm : m ≤ m0) [sigma_finite (μ.trim hm)] diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 48f6c027dede1..9e9acb3aaca47 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -27,7 +27,7 @@ begin simpa only [add_comm x] using exists_unique_add_zsmul_mem_Ioc hT x t end -variables {E : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] namespace function diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index a5f595292c4fb..44c48877bb234 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -54,9 +54,9 @@ variables {α β E F : Type*} [measurable_space α] namespace measure_theory -section normed_group +section normed_add_comm_group -variables [normed_group E] {f g : α → E} {s t : set α} {μ ν : measure α} +variables [normed_add_comm_group E] {f g : α → E} {s t : set α} {μ ν : measure α} {l l' : filter α} variables [complete_space E] [normed_space ℝ E] @@ -417,7 +417,7 @@ lemma integral_Ioc_eq_integral_Ioo [partial_order α] {f : α → E} {a b : α} ∫ t in Ioc a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := integral_Ioc_eq_integral_Ioo' $ measure_singleton b -end normed_group +end normed_add_comm_group section mono @@ -521,7 +521,7 @@ end nonneg section tendsto_mono -variables {μ : measure α} [normed_group E] [complete_space E] [normed_space ℝ E] +variables {μ : measure α} [normed_add_comm_group E] [complete_space E] [normed_space ℝ E] {s : ℕ → set α} {f : α → E} lemma _root_.antitone.tendsto_set_integral (hsm : ∀ i, measurable_set (s i)) @@ -552,8 +552,8 @@ end tendsto_mono We prove that for any set `s`, the function `λ f : α →₁[μ] E, ∫ x in s, f x ∂μ` is continuous. -/ section continuous_set_integral -variables [normed_group E] {𝕜 : Type*} [normed_field 𝕜] [normed_group F] [normed_space 𝕜 F] - {p : ℝ≥0∞} {μ : measure α} +variables [normed_add_comm_group E] {𝕜 : Type*} [normed_field 𝕜] [normed_add_comm_group F] + [normed_space 𝕜 F] {p : ℝ≥0∞} {μ : measure α} /-- For `f : Lp E p μ`, we can define an element of `Lp E p (μ.restrict s)` by `(Lp.mem_ℒp f).restrict s).to_Lp f`. This map is additive. -/ @@ -634,7 +634,7 @@ end measure_theory open measure_theory asymptotics metric -variables {ι : Type*} [normed_group E] +variables {ι : Type*} [normed_add_comm_group E] /-- Fundamental theorem of calculus for set integrals: if `μ` is a measure that is finite at a filter `l` and `f` is a measurable function that has a finite limit `b` at `l ⊓ μ.ae`, then `∫ x in @@ -744,7 +744,7 @@ as `continuous_linear_map.comp_Lp`. We take advantage of this construction here. open_locale complex_conjugate variables {μ : measure α} {𝕜 : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] - [normed_group F] [normed_space 𝕜 F] + [normed_add_comm_group F] [normed_space 𝕜 F] {p : ennreal} namespace continuous_linear_map @@ -786,7 +786,7 @@ begin all_goals { assumption } end -lemma integral_apply {H : Type*} [normed_group H] [normed_space 𝕜 H] +lemma integral_apply {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] {φ : α → H →L[𝕜] E} (φ_int : integrable φ μ) (v : H) : (∫ a, φ a ∂μ) v = ∫ a, φ a v ∂μ := ((continuous_linear_map.apply 𝕜 E v).integral_comp_comm φ_int).symm diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index 3987e9b4686c5..01ba2d27259f4 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -74,10 +74,10 @@ open set filter topological_space ennreal emetric namespace measure_theory variables {α E F F' G 𝕜 : Type*} {p : ℝ≥0∞} - [normed_group E] [normed_space ℝ E] - [normed_group F] [normed_space ℝ F] - [normed_group F'] [normed_space ℝ F'] - [normed_group G] + [normed_add_comm_group E] [normed_space ℝ E] + [normed_add_comm_group F] [normed_space ℝ F] + [normed_add_comm_group F'] [normed_space ℝ F'] + [normed_add_comm_group G] {m : measurable_space α} {μ : measure α} local infixr ` →ₛ `:25 := simple_func @@ -180,13 +180,13 @@ end fin_meas_additive /-- A `fin_meas_additive` set function whose norm on every set is less than the measure of the set (up to a multiplicative constant). -/ -def dominated_fin_meas_additive {β} [semi_normed_group β] {m : measurable_space α} +def dominated_fin_meas_additive {β} [seminormed_add_comm_group β] {m : measurable_space α} (μ : measure α) (T : set α → β) (C : ℝ) : Prop := fin_meas_additive μ T ∧ ∀ s, measurable_set s → μ s < ∞ → ∥T s∥ ≤ C * (μ s).to_real namespace dominated_fin_meas_additive -variables {β : Type*} [semi_normed_group β] {T T' : set α → β} {C C' : ℝ} +variables {β : Type*} [seminormed_add_comm_group β] {T T' : set α → β} {C C' : ℝ} lemma zero {m : measurable_space α} (μ : measure α) (hC : 0 ≤ C) : dominated_fin_meas_additive μ (0 : set α → β) C := @@ -196,7 +196,7 @@ begin exact mul_nonneg hC to_real_nonneg, end -lemma eq_zero_of_measure_zero {β : Type*} [normed_group β] {T : set α → β} {C : ℝ} +lemma eq_zero_of_measure_zero {β : Type*} [normed_add_comm_group β] {T : set α → β} {C : ℝ} (hT : dominated_fin_meas_additive μ T C) {s : set α} (hs : measurable_set s) (hs_zero : μ s = 0) : T s = 0 := @@ -206,7 +206,7 @@ begin rw [hs_zero, ennreal.zero_to_real, mul_zero], end -lemma eq_zero {β : Type*} [normed_group β] {T : set α → β} {C : ℝ} +lemma eq_zero {β : Type*} [normed_add_comm_group β] {T : set α → β} {C : ℝ} {m : measurable_space α} (hT : dominated_fin_meas_additive (0 : measure α) T C) {s : set α} (hs : measurable_set s) : T s = 0 := @@ -511,7 +511,7 @@ calc set_to_simple_func T (c • f) = ∑ x in f.range, T (f ⁻¹' {x}) (c • ... = c • set_to_simple_func T f : by simp only [set_to_simple_func, smul_sum, smul_smul, mul_comm] -lemma set_to_simple_func_smul {E} [normed_group E] [normed_field 𝕜] +lemma set_to_simple_func_smul {E} [normed_add_comm_group E] [normed_field 𝕜] [normed_space 𝕜 E] [normed_space ℝ E] [normed_space 𝕜 F] (T : set α → E →L[ℝ] F) (h_add : fin_meas_additive μ T) (h_smul : ∀ c : 𝕜, ∀ s x, T s (c • x) = c • T s x) (c : 𝕜) {f : α →ₛ E} (hf : integrable f μ) : @@ -818,7 +818,7 @@ begin exact smul_to_simple_func c f, end -lemma set_to_L1s_smul {E} [normed_group E] [normed_space ℝ E] +lemma set_to_L1s_smul {E} [normed_add_comm_group E] [normed_space ℝ E] [normed_space 𝕜 E] [normed_space 𝕜 F] (T : set α → E →L[ℝ] F) (h_zero : ∀ s, measurable_set s → μ s = 0 → T s = 0) (h_add : fin_meas_additive μ T) diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean index d4f8f02512a3e..01fdf2694ed15 100644 --- a/src/measure_theory/integral/torus_integral.lean +++ b/src/measure_theory/integral/torus_integral.lean @@ -53,7 +53,7 @@ integral, torus -/ variable {n : ℕ} -variables {E : Type*} [normed_group E] +variables {E : Type*} [normed_add_comm_group E] noncomputable theory diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index e6cf45bc7797c..a29416200a9bd 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -92,7 +92,7 @@ namespace measure /-- If a set is disjoint of its translates by infinitely many bounded vectors, then it has measure zero. This auxiliary lemma proves this assuming additionally that the set is bounded. -/ lemma add_haar_eq_zero_of_disjoint_translates_aux - {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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} (u : ℕ → E) (sb : bounded s) (hu : bounded (range u)) (hs : pairwise (disjoint on (λ n, {u n} + s))) (h's : measurable_set s) : @@ -114,7 +114,7 @@ end /-- If a set is disjoint of its translates by infinitely many bounded vectors, then it has measure zero. -/ lemma add_haar_eq_zero_of_disjoint_translates - {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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} (u : ℕ → E) (hu : bounded (range u)) (hs : pairwise (disjoint on (λ n, {u n} + s))) (h's : measurable_set s) : @@ -135,7 +135,7 @@ end /-- A strict vector subspace has measure zero. -/ lemma add_haar_submodule - {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 : submodule ℝ E) (hs : s ≠ ⊤) : μ s = 0 := begin @@ -165,7 +165,7 @@ end /-- A strict affine subspace has measure zero. -/ lemma add_haar_affine_subspace - {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 : affine_subspace ℝ E) (hs : s ≠ ⊤) : μ s = 0 := begin @@ -200,7 +200,7 @@ begin end lemma map_linear_map_add_haar_eq_smul_add_haar - {E : Type*} [normed_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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)⁻¹) • μ := @@ -235,7 +235,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -248,7 +248,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -257,7 +257,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -270,7 +270,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -279,7 +279,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -301,7 +301,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -310,7 +310,7 @@ 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_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] + {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 := @@ -320,7 +320,7 @@ 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_group E] [measurable_space E] [normed_space ℝ E] +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 μ] lemma map_add_haar_smul {r : ℝ} (hr : r ≠ 0) : @@ -373,7 +373,7 @@ general Haar measures on general commutative groups. -/ /-! ### Measure of balls -/ lemma add_haar_ball_center - {E : Type*} [normed_group E] [measurable_space E] + {E : Type*} [normed_add_comm_group E] [measurable_space E] [borel_space E] (μ : measure E) [is_add_haar_measure μ] (x : E) (r : ℝ) : μ (ball x r) = μ (ball (0 : E) r) := begin @@ -382,7 +382,7 @@ begin end lemma add_haar_closed_ball_center - {E : Type*} [normed_group E] [measurable_space E] + {E : Type*} [normed_add_comm_group E] [measurable_space E] [borel_space E] (μ : measure E) [is_add_haar_measure μ] (x : E) (r : ℝ) : μ (closed_ball x r) = μ (closed_ball (0 : E) r) := begin diff --git a/src/measure_theory/measure/with_density_vector_measure.lean b/src/measure_theory/measure/with_density_vector_measure.lean index b67c51df4c2b9..fb74f73525300 100644 --- a/src/measure_theory/measure/with_density_vector_measure.lean +++ b/src/measure_theory/measure/with_density_vector_measure.lean @@ -31,7 +31,7 @@ namespace measure_theory open topological_space variables {μ ν : measure α} -variables {E : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] /-- Given a measure `μ` and an integrable function `f`, `μ.with_densityᵥ f` is the vector measure which maps the set `s` to `∫ₛ f ∂μ`. -/ diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index ec96482548042..0f32ee0bba55b 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -108,7 +108,7 @@ end bottom_row section tendsto_lemmas open filter continuous_linear_map -local attribute [instance] matrix.normed_group matrix.normed_space +local attribute [instance] matrix.normed_add_comm_group matrix.normed_space local attribute [simp] coe_smul /-- The function `(c,d) → |cz+d|^2` is proper, that is, preimages of bounded-above sets are finite. diff --git a/src/probability/conditional_expectation.lean b/src/probability/conditional_expectation.lean index ac7a142cbb884..77acd08812dae 100644 --- a/src/probability/conditional_expectation.lean +++ b/src/probability/conditional_expectation.lean @@ -27,7 +27,7 @@ namespace measure_theory open probability_theory -variables {α E : Type*} [normed_group E] [normed_space ℝ E] [complete_space E] +variables {α E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] {m₁ m₂ m : measurable_space α} {μ : measure α} {f : α → E} /-- If `m₁, m₂` are independent σ-algebras and `f` is `m₁`-measurable, then `𝔼[f | m₂] = 𝔼[f]` diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index b876aaa80a4bc..b86122e18c847 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -171,7 +171,7 @@ begin ← lintegral_map' ae_measurable_id h.ae_measurable_snd, h.map_eq], end -lemma integral_eq [normed_group γ] [normed_space ℝ γ] [complete_space γ] [borel_space γ] +lemma integral_eq [normed_add_comm_group γ] [normed_space ℝ γ] [complete_space γ] [borel_space γ] (h : ident_distrib f g μ ν) : ∫ x, f x ∂μ = ∫ x, g x ∂ν := begin by_cases hf : ae_strongly_measurable f μ, @@ -191,7 +191,8 @@ begin rw integral_non_ae_strongly_measurable hf } end -lemma snorm_eq [normed_group γ] [opens_measurable_space γ] (h : ident_distrib f g μ ν) (p : ℝ≥0∞) : +lemma snorm_eq [normed_add_comm_group γ] [opens_measurable_space γ] (h : ident_distrib f g μ ν) + (p : ℝ≥0∞) : snorm f p μ = snorm g p ν := begin by_cases h0 : p = 0, @@ -208,7 +209,7 @@ begin (measurable.pow_const (measurable_coe_nnreal_ennreal.comp measurable_nnnorm) p.to_real), end -lemma mem_ℒp_snd [normed_group γ] [borel_space γ] +lemma mem_ℒp_snd [normed_add_comm_group γ] [borel_space γ] {p : ℝ≥0∞} (h : ident_distrib f g μ ν) (hf : mem_ℒp f p μ) : mem_ℒp g p ν := begin @@ -217,26 +218,26 @@ begin exact hf.2 end -lemma mem_ℒp_iff [normed_group γ] [borel_space γ] {p : ℝ≥0∞} (h : ident_distrib f g μ ν) : +lemma mem_ℒp_iff [normed_add_comm_group γ] [borel_space γ] {p : ℝ≥0∞} (h : ident_distrib f g μ ν) : mem_ℒp f p μ ↔ mem_ℒp g p ν := ⟨λ hf, h.mem_ℒp_snd hf, λ hg, h.symm.mem_ℒp_snd hg⟩ -lemma integrable_snd [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) +lemma integrable_snd [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) (hf : integrable f μ) : integrable g ν := begin rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact h.mem_ℒp_snd hf end -lemma integrable_iff [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : +lemma integrable_iff [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : integrable f μ ↔ integrable g ν := ⟨λ hf, h.integrable_snd hf, λ hg, h.symm.integrable_snd hg⟩ -protected lemma norm [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : +protected lemma norm [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : ident_distrib (λ x, ∥f x∥) (λ x, ∥g x∥) μ ν := h.comp measurable_norm -protected lemma nnnorm [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : +protected lemma nnnorm [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : ident_distrib (λ x, ∥f x∥₊) (λ x, ∥g x∥₊) μ ν := h.comp measurable_nnnorm diff --git a/src/probability/martingale.lean b/src/probability/martingale.lean index 178c420e9db06..119e8184fedbd 100644 --- a/src/probability/martingale.lean +++ b/src/probability/martingale.lean @@ -42,7 +42,7 @@ namespace measure_theory variables {α E ι : Type*} [preorder ι] {m0 : measurable_space α} {μ : measure α} - [normed_group E] [normed_space ℝ E] [complete_space E] + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] {f g : ι → α → E} {ℱ : filtration ι m0} /-- A family of functions `f : ι → α → E` is a martingale with respect to a filtration `ℱ` if `f` diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index ae804da5a3940..ce4f010b0544c 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -1361,9 +1361,9 @@ end end add_comm_monoid -section normed_group +section normed_add_comm_group -variables [normed_group β] {p : ℝ≥0∞} {μ : measure α} +variables [normed_add_comm_group β] {p : ℝ≥0∞} {μ : measure α} lemma mem_ℒp_stopped_process (hτ : is_stopping_time f τ) (hu : ∀ n, mem_ℒp (u n) p μ) (n : ℕ) : mem_ℒp (stopped_process u τ n) p μ := @@ -1399,7 +1399,7 @@ lemma integrable_stopped_value (hτ : is_stopping_time f τ) integrable (stopped_value u τ) μ := by { simp_rw ← mem_ℒp_one_iff_integrable at hu ⊢, exact mem_ℒp_stopped_value hτ hu hbdd, } -end normed_group +end normed_add_comm_group end nat diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 6f98349849b8b..f817af85adeeb 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -635,7 +635,8 @@ instance : has_lipschitz_add (α →ᵇ β) := apply max_le_max; exact dist_coe_le_dist x, end⟩ } -/-- Coercion of a `normed_group_hom` is an `add_monoid_hom`. Similar to `add_monoid_hom.coe_fn` -/ +/-- Coercion of a `normed_add_group_hom` is an `add_monoid_hom`. Similar to +`add_monoid_hom.coe_fn`. -/ @[simps] def coe_fn_add_hom : (α →ᵇ β) →+ (α → β) := { to_fun := coe_fn, map_zero' := coe_zero, map_add' := coe_add } @@ -670,12 +671,12 @@ by simp end comm_has_lipschitz_add -section normed_group +section normed_add_comm_group /- In this section, if β is a normed group, then we show that the space of bounded continuous functions from α to β inherits a normed group structure, by using pointwise operations and checking that they are compatible with the uniform distance. -/ -variables [topological_space α] [semi_normed_group β] +variables [topological_space α] [seminormed_add_comm_group β] variables (f g : α →ᵇ β) {x : α} {C : ℝ} instance : has_norm (α →ᵇ β) := ⟨λu, dist u 0⟩ @@ -756,33 +757,33 @@ le_antisymm (norm_const_le b) $ h.elim $ λ x, (const α b).norm_coe_le_norm x /-- Constructing a bounded continuous function from a uniformly bounded continuous function taking values in a normed group. -/ -def of_normed_group {α : Type u} {β : Type v} [topological_space α] [semi_normed_group β] - (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : α →ᵇ β := +def of_normed_add_comm_group {α : Type u} {β : Type v} [topological_space α] + [seminormed_add_comm_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : + α →ᵇ β := ⟨⟨λn, f n, Hf⟩, ⟨_, dist_le_two_norm' H⟩⟩ -@[simp] lemma coe_of_normed_group - {α : Type u} {β : Type v} [topological_space α] [semi_normed_group β] +@[simp] lemma coe_of_normed_add_comm_group + {α : Type u} {β : Type v} [topological_space α] [seminormed_add_comm_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : - (of_normed_group f Hf C H : α → β) = f := rfl + (of_normed_add_comm_group f Hf C H : α → β) = f := rfl -lemma norm_of_normed_group_le {f : α → β} (hfc : continuous f) {C : ℝ} (hC : 0 ≤ C) - (hfC : ∀ x, ∥f x∥ ≤ C) : ∥of_normed_group f hfc C hfC∥ ≤ C := +lemma norm_of_normed_add_comm_group_le {f : α → β} (hfc : continuous f) {C : ℝ} (hC : 0 ≤ C) + (hfC : ∀ x, ∥f x∥ ≤ C) : ∥of_normed_add_comm_group f hfc C hfC∥ ≤ C := (norm_le hC).2 hfC /-- Constructing a bounded continuous function from a uniformly bounded function on a discrete space, taking values in a normed group -/ -def of_normed_group_discrete {α : Type u} {β : Type v} - [topological_space α] [discrete_topology α] [semi_normed_group β] +def of_normed_add_comm_group_discrete {α : Type u} {β : Type v} + [topological_space α] [discrete_topology α] [seminormed_add_comm_group β] (f : α → β) (C : ℝ) (H : ∀x, norm (f x) ≤ C) : α →ᵇ β := -of_normed_group f continuous_of_discrete_topology C H +of_normed_add_comm_group f continuous_of_discrete_topology C H -@[simp] lemma coe_of_normed_group_discrete - {α : Type u} {β : Type v} [topological_space α] [discrete_topology α] [semi_normed_group β] - (f : α → β) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : - (of_normed_group_discrete f C H : α → β) = f := rfl +@[simp] lemma coe_of_normed_add_comm_group_discrete {α : Type u} {β : Type v} [topological_space α] + [discrete_topology α] [seminormed_add_comm_group β] (f : α → β) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : + (of_normed_add_comm_group_discrete f C H : α → β) = f := rfl -/-- Taking the pointwise norm of a bounded continuous function with values in a `semi_normed_group`, -yields a bounded continuous function with values in ℝ. -/ +/-- Taking the pointwise norm of a bounded continuous function with values in a +`seminormed_add_comm_group` yields a bounded continuous function with values in ℝ. -/ def norm_comp : α →ᵇ ℝ := f.comp norm lipschitz_with_one_norm @@ -799,12 +800,12 @@ by simp_rw [norm_def, dist_eq_supr, coe_zero, pi.zero_apply, dist_zero_right] /-- The pointwise opposite of a bounded continuous function is again bounded continuous. -/ instance : has_neg (α →ᵇ β) := -⟨λf, of_normed_group (-f) f.continuous.neg ∥f∥ $ λ x, +⟨λf, of_normed_add_comm_group (-f) f.continuous.neg ∥f∥ $ λ x, trans_rel_right _ (norm_neg _) (f.norm_coe_le_norm x)⟩ /-- The pointwise difference of two bounded continuous functions is again bounded continuous. -/ instance : has_sub (α →ᵇ β) := -⟨λf g, of_normed_group (f - g) (f.continuous.sub g.continuous) (∥f∥ + ∥g∥) $ λ x, +⟨λf g, of_normed_add_comm_group (f - g) (f.continuous.sub g.continuous) (∥f∥ + ∥g∥) $ λ x, by { simp only [sub_eq_add_neg], exact le_trans (norm_add_le _ _) (add_le_add (f.norm_coe_le_norm x) $ trans_rel_right _ (norm_neg _) (g.norm_coe_le_norm x)) }⟩ @@ -837,11 +838,11 @@ instance : add_comm_group (α →ᵇ β) := fun_like.coe_injective.add_comm_group _ coe_zero coe_add coe_neg coe_sub (λ _ _, coe_nsmul _ _) (λ _ _, coe_zsmul _ _) -instance : semi_normed_group (α →ᵇ β) := +instance : seminormed_add_comm_group (α →ᵇ β) := { dist_eq := λ f g, by simp only [norm_eq, dist_eq, dist_eq_norm, sub_apply] } -instance {α β} [topological_space α] [normed_group β] : normed_group (α →ᵇ β) := -{ ..bounded_continuous_function.semi_normed_group } +instance {α β} [topological_space α] [normed_add_comm_group β] : normed_add_comm_group (α →ᵇ β) := +{ ..bounded_continuous_function.seminormed_add_comm_group } lemma nnnorm_def : ∥f∥₊ = nndist f 0 := rfl @@ -873,7 +874,7 @@ lemma norm_comp_continuous_le [topological_space γ] (f : α →ᵇ β) (g : C( ((lipschitz_comp_continuous g).dist_le_mul f 0).trans $ by rw [nnreal.coe_one, one_mul, dist_zero_right] -end normed_group +end normed_add_comm_group section has_bounded_smul /-! @@ -979,16 +980,16 @@ continuous functions from `α` to `β` inherits a normed space structure, by usi pointwise operations and checking that they are compatible with the uniform distance. -/ variables {𝕜 : Type*} -variables [topological_space α] [semi_normed_group β] +variables [topological_space α] [seminormed_add_comm_group β] variables {f g : α →ᵇ β} {x : α} {C : ℝ} instance [normed_field 𝕜] [normed_space 𝕜 β] : normed_space 𝕜 (α →ᵇ β) := ⟨λ c f, begin - refine norm_of_normed_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, + refine norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, exact (λ x, trans_rel_right _ (norm_smul _ _) (mul_le_mul_of_nonneg_left (f.norm_coe_le_norm _) (norm_nonneg _))) end⟩ variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 β] -variables [semi_normed_group γ] [normed_space 𝕜 γ] +variables [seminormed_add_comm_group γ] [normed_space 𝕜 γ] variables (α) -- TODO does this work in the `has_bounded_smul` setting, too? @@ -1000,7 +1001,7 @@ Upgraded version of `continuous_linear_map.comp_left_continuous`, similar to protected def _root_.continuous_linear_map.comp_left_continuous_bounded (g : β →L[𝕜] γ) : (α →ᵇ β) →L[𝕜] (α →ᵇ γ) := linear_map.mk_continuous - { to_fun := λ f, of_normed_group + { to_fun := λ f, of_normed_add_comm_group (g ∘ f) (g.continuous.comp f.continuous) (∥g∥ * ∥f∥) @@ -1008,7 +1009,7 @@ linear_map.mk_continuous map_add' := λ f g, by ext; simp, map_smul' := λ c f, by ext; simp } ∥g∥ - (λ f, norm_of_normed_group_le _ (mul_nonneg (norm_nonneg g) (norm_nonneg f)) _) + (λ f, norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg g) (norm_nonneg f)) _) @[simp] lemma _root_.continuous_linear_map.comp_left_continuous_bounded_apply (g : β →L[𝕜] γ) (f : α →ᵇ β) (x : α) : @@ -1033,7 +1034,7 @@ section semi_normed variables [non_unital_semi_normed_ring R] instance : has_mul (α →ᵇ R) := -{ mul := λ f g, of_normed_group (f * g) (f.continuous.mul g.continuous) (∥f∥ * ∥g∥) $ λ x, +{ mul := λ f g, of_normed_add_comm_group (f * g) (f.continuous.mul g.continuous) (∥f∥ * ∥g∥) $ λ x, le_trans (norm_mul_le (f x) (g x)) $ mul_le_mul (f.norm_coe_le_norm x) (g.norm_coe_le_norm x) (norm_nonneg _) (norm_nonneg _) } @@ -1045,14 +1046,15 @@ fun_like.coe_injective.non_unital_ring _ coe_zero coe_add coe_mul coe_neg coe_su (λ _ _, coe_nsmul _ _) (λ _ _, coe_zsmul _ _) instance : non_unital_semi_normed_ring (α →ᵇ R) := -{ norm_mul := λ f g, norm_of_normed_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, - .. bounded_continuous_function.semi_normed_group } +{ norm_mul := λ f g, norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) + _, + .. bounded_continuous_function.seminormed_add_comm_group } end semi_normed instance [non_unital_normed_ring R] : non_unital_normed_ring (α →ᵇ R) := { .. bounded_continuous_function.non_unital_semi_normed_ring, - .. bounded_continuous_function.normed_group } + .. bounded_continuous_function.normed_add_comm_group } end non_unital @@ -1115,10 +1117,10 @@ instance [semi_normed_comm_ring R] : comm_ring (α →ᵇ R) := .. bounded_continuous_function.ring } instance [semi_normed_comm_ring R] : semi_normed_comm_ring (α →ᵇ R) := -{ .. bounded_continuous_function.comm_ring, .. bounded_continuous_function.semi_normed_group } +{ ..bounded_continuous_function.comm_ring, ..bounded_continuous_function.seminormed_add_comm_group } instance [normed_comm_ring R] : normed_comm_ring (α →ᵇ R) := -{ .. bounded_continuous_function.comm_ring, .. bounded_continuous_function.normed_group } +{ .. bounded_continuous_function.comm_ring, .. bounded_continuous_function.normed_add_comm_group } end normed_comm_ring @@ -1131,7 +1133,7 @@ continuous functions from `α` to `γ` inherits a normed algebra structure, by u pointwise operations and checking that they are compatible with the uniform distance. -/ variables {𝕜 : Type*} [normed_field 𝕜] -variables [topological_space α] [semi_normed_group β] [normed_space 𝕜 β] +variables [topological_space α] [seminormed_add_comm_group β] [normed_space 𝕜 β] variables [normed_ring γ] [normed_algebra 𝕜 γ] variables {f g : α →ᵇ γ} {x : α} {c : 𝕜} @@ -1165,7 +1167,7 @@ functions from `α` to `β` is naturally a module over the algebra of bounded co functions from `α` to `𝕜`. -/ instance has_smul' : has_smul (α →ᵇ 𝕜) (α →ᵇ β) := -⟨λ (f : α →ᵇ 𝕜) (g : α →ᵇ β), of_normed_group (λ x, (f x) • (g x)) +⟨λ (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∥ * ∥g∥ : mul_le_mul (f.norm_coe_le_norm _) (g.norm_coe_le_norm _) (norm_nonneg _) @@ -1180,7 +1182,7 @@ module.of_core $ one_smul := λ f, ext $ λ x, one_smul 𝕜 (f x) } lemma norm_smul_le (f : α →ᵇ 𝕜) (g : α →ᵇ β) : ∥f • g∥ ≤ ∥f∥ * ∥g∥ := -norm_of_normed_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _ +norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _ /- TODO: When `normed_module` has been added to `normed_space.basic`, the above facts show that the space of bounded continuous functions from `α` to `β` is naturally a normed @@ -1214,14 +1216,14 @@ In summary, if `β` is a C⋆-algebra over `𝕜`, then so is `α →ᵇ β`; n completeness is guaranteed when `β` is complete (see `bounded_continuous_function.complete`). -/ -section normed_group +section normed_add_comm_group -variables {𝕜 : Type*} [normed_field 𝕜] [star_ring 𝕜] -variables [topological_space α] [semi_normed_group β] [star_add_monoid β] [normed_star_group β] +variables {𝕜 : Type*} [normed_field 𝕜] [star_ring 𝕜] [topological_space α] + [seminormed_add_comm_group β] [star_add_monoid β] [normed_star_group β] variables [normed_space 𝕜 β] [star_module 𝕜 β] instance : star_add_monoid (α →ᵇ β) := -{ star := λ f, f.comp star star_normed_group_hom.lipschitz, +{ star := λ f, f.comp star star_normed_add_group_hom.lipschitz, star_involutive := λ f, ext $ λ x, star_star (f x), star_add := λ f g, ext $ λ x, star_add (f x) (g x) } @@ -1237,7 +1239,7 @@ instance : normed_star_group (α →ᵇ β) := instance : star_module 𝕜 (α →ᵇ β) := { star_smul := λ k f, ext $ λ x, star_smul k (f x) } -end normed_group +end normed_add_comm_group section cstar_ring @@ -1287,7 +1289,7 @@ instance : semilattice_inf (α →ᵇ β) := obtain ⟨C₁, hf⟩ := f.bounded, obtain ⟨C₂, hg⟩ := g.bounded, refine ⟨C₁ + C₂, λ x y, _⟩, - simp_rw normed_group.dist_eq at hf hg ⊢, + simp_rw normed_add_comm_group.dist_eq at hf hg ⊢, exact (norm_inf_sub_inf_le_add_norm _ _ _ _).trans (add_le_add (hf _ _) (hg _ _)), end }, inf_le_left := λ f g, continuous_map.le_def.mpr (λ _, inf_le_left), @@ -1304,7 +1306,7 @@ instance : semilattice_sup (α →ᵇ β) := obtain ⟨C₁, hf⟩ := f.bounded, obtain ⟨C₂, hg⟩ := g.bounded, refine ⟨C₁ + C₂, λ x y, _⟩, - simp_rw normed_group.dist_eq at hf hg ⊢, + simp_rw normed_add_comm_group.dist_eq at hf hg ⊢, exact (norm_sup_sub_sup_le_add_norm _ _ _ _).trans (add_le_add (hf _ _) (hg _ _)), end }, le_sup_left := λ f g, continuous_map.le_def.mpr (λ _, le_sup_left), diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 36df30b9a6952..b415647de2ea7 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -32,7 +32,8 @@ open bounded_continuous_function namespace continuous_map -variables {α β E : Type*} [topological_space α] [compact_space α] [metric_space β] [normed_group E] +variables {α β E : Type*} [topological_space α] [compact_space α] [metric_space β] + [normed_add_comm_group E] section @@ -153,7 +154,7 @@ rfl open bounded_continuous_function -instance : normed_group C(α, E) := +instance : normed_add_comm_group C(α, E) := { dist_eq := λ x y, by rw [← norm_mk_of_compact, ← dist_mk_of_compact, dist_eq_norm, mk_of_compact_sub], dist := dist, norm := norm, .. continuous_map.metric_space _ _, .. continuous_map.add_comm_group } @@ -201,7 +202,7 @@ variables {R : Type*} [normed_ring R] instance : normed_ring C(α,R) := { norm_mul := λ f g, norm_mul_le (mk_of_compact f) (mk_of_compact g), - ..(infer_instance : normed_group C(α,R)), + ..(infer_instance : normed_add_comm_group C(α,R)), .. continuous_map.ring } end @@ -308,7 +309,7 @@ end continuous_map section comp_left variables (X : Type*) {𝕜 β γ : Type*} [topological_space X] [compact_space X] [nontrivially_normed_field 𝕜] -variables [normed_group β] [normed_space 𝕜 β] [normed_group γ] [normed_space 𝕜 γ] +variables [normed_add_comm_group β] [normed_space 𝕜 β] [normed_add_comm_group γ] [normed_space 𝕜 γ] open continuous_map @@ -351,8 +352,8 @@ section comp_right /-- Precomposition by a continuous map is itself a continuous map between spaces of continuous maps. -/ -def comp_right_continuous_map {X Y : Type*} (T : Type*) - [topological_space X] [compact_space X] [topological_space Y] [compact_space Y] [normed_group T] +def comp_right_continuous_map {X Y : Type*} (T : Type*) [topological_space X] [compact_space X] + [topological_space Y] [compact_space Y] [normed_add_comm_group T] (f : C(X, Y)) : C(C(Y, T), C(X, T)) := { to_fun := λ g, g.comp f, continuous_to_fun := @@ -364,8 +365,8 @@ def comp_right_continuous_map {X Y : Type*} (T : Type*) { exact λ x, h (f x), }, end } -@[simp] lemma comp_right_continuous_map_apply {X Y : Type*} (T : Type*) - [topological_space X] [compact_space X] [topological_space Y] [compact_space Y] [normed_group T] +@[simp] lemma comp_right_continuous_map_apply {X Y : Type*} (T : Type*) [topological_space X] + [compact_space X] [topological_space Y] [compact_space Y] [normed_add_comm_group T] (f : C(X, Y)) (g : C(Y, T)) : (comp_right_continuous_map T f) g = g.comp f := rfl @@ -373,8 +374,8 @@ rfl /-- Precomposition by a homeomorphism is itself a homeomorphism between spaces of continuous maps. -/ -def comp_right_homeomorph {X Y : Type*} (T : Type*) - [topological_space X] [compact_space X] [topological_space Y] [compact_space Y] [normed_group T] +def comp_right_homeomorph {X Y : Type*} (T : Type*) [topological_space X] [compact_space X] + [topological_space Y] [compact_space Y] [normed_add_comm_group T] (f : X ≃ₜ Y) : C(Y, T) ≃ₜ C(X, T) := { to_fun := comp_right_continuous_map T f.to_continuous_map, inv_fun := comp_right_continuous_map T f.symm.to_continuous_map, @@ -415,7 +416,7 @@ section weierstrass open topological_space variables {X : Type*} [topological_space X] [t2_space X] [locally_compact_space X] -variables {E : Type*} [normed_group E] [complete_space E] +variables {E : Type*} [normed_add_comm_group E] [complete_space E] lemma summable_of_locally_summable_norm {ι : Type*} {F : ι → C(X, E)} (hF : ∀ K : compacts X, summable (λ i, ∥(F i).restrict K∥)) : @@ -442,7 +443,7 @@ Furthermore, if `α` is compact and `β` is a C⋆-ring, then `C(α, β)` is a C section normed_space variables {α : Type*} {β : Type*} -variables [topological_space α] [normed_group β] [star_add_monoid β] [normed_star_group β] +variables [topological_space α] [normed_add_comm_group β] [star_add_monoid β] [normed_star_group β] lemma _root_.bounded_continuous_function.mk_of_compact_star [compact_space α] (f : C(α, β)) : mk_of_compact (star f) = star (mk_of_compact f) := rfl diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 2c019c816e743..96c93b4f6a7fd 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -382,7 +382,7 @@ field `𝕜` whenever `β` is as well. section normed_space -variables [normed_group β] {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] +variables [normed_add_comm_group β] {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] /-- The natural inclusion `to_bcf : C₀(α, β) → (α →ᵇ β)` realized as an additive monoid homomorphism. -/ @@ -394,8 +394,8 @@ def to_bcf_add_monoid_hom : C₀(α, β) →+ (α →ᵇ β) := @[simp] lemma coe_to_bcf_add_monoid_hom (f : C₀(α, β)) : (f.to_bcf_add_monoid_hom : α → β) = f := rfl -noncomputable instance : normed_group C₀(α, β) := -normed_group.induced to_bcf_add_monoid_hom (to_bcf_injective α β) +noncomputable instance : normed_add_comm_group C₀(α, β) := +normed_add_comm_group.induced to_bcf_add_monoid_hom (to_bcf_injective α β) @[simp] lemma norm_to_bcf_eq_norm {f : C₀(α, β)} : ∥f.to_bcf∥ = ∥f∥ := rfl @@ -412,7 +412,7 @@ variables [non_unital_normed_ring β] noncomputable instance : non_unital_normed_ring C₀(α, β) := { norm_mul := λ f g, norm_mul_le f.to_bcf g.to_bcf, ..zero_at_infty_continuous_map.non_unital_ring, - ..zero_at_infty_continuous_map.normed_group } + ..zero_at_infty_continuous_map.normed_add_comm_group } end normed_ring @@ -453,7 +453,7 @@ end star section normed_star -variables [normed_group β] [star_add_monoid β] [normed_star_group β] +variables [normed_add_comm_group β] [star_add_monoid β] [normed_star_group β] instance : normed_star_group C₀(α, β) := { norm_star := λ f, (norm_star f.to_bcf : _) } diff --git a/src/topology/metric_space/algebra.lean b/src/topology/metric_space/algebra.lean index 3ed29eb490197..db5792f8f9fb5 100644 --- a/src/topology/metric_space/algebra.lean +++ b/src/topology/metric_space/algebra.lean @@ -77,8 +77,8 @@ instance has_lipschitz_mul.has_continuous_mul : has_continuous_mul β := (lipschitz_with_lipschitz_const_mul_edist ⟨x₂.unop, x₁.unop⟩ ⟨y₂.unop, y₁.unop⟩).trans_eq (congr_arg _ $ max_comm _ _)⟩ } --- this instance could be deduced from `normed_group.has_lipschitz_add`, but we prove it separately --- here so that it is available earlier in the hierarchy +-- this instance could be deduced from `normed_add_comm_group.has_lipschitz_add`, but we prove it +-- separately here so that it is available earlier in the hierarchy instance real.has_lipschitz_add : has_lipschitz_add ℝ := { lipschitz_add := ⟨2, begin rw lipschitz_with_iff_dist_le_mul, diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 45712811344d6..6f326924b6214 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -403,7 +403,7 @@ end isometric namespace continuous_linear_equiv variables {𝕜 E F : Type*} [nontrivially_normed_field 𝕜] - [normed_group E] [normed_space 𝕜 E] [normed_group F] [normed_space 𝕜 F] + [normed_add_comm_group E] [normed_space 𝕜 E] [normed_add_comm_group F] [normed_space 𝕜 F] @[simp] lemma dimH_image (e : E ≃L[𝕜] F) (s : set E) : dimH (e '' s) = dimH s := le_antisymm (e.lipschitz.dimH_image_le s) $ @@ -423,7 +423,8 @@ end continuous_linear_equiv namespace real -variables {E : Type*} [fintype ι] [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +variables {E : Type*} [fintype ι] [normed_add_comm_group E] [normed_space ℝ E] + [finite_dimensional ℝ E] theorem dimH_ball_pi (x : ι → ℝ) {r : ℝ} (hr : 0 < r) : dimH (metric.ball x r) = fintype.card ι := @@ -478,8 +479,8 @@ by rw [dimH_univ_eq_finrank ℝ, finite_dimensional.finrank_self, nat.cast_one] end real variables {E F : Type*} - [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E] - [normed_group F] [normed_space ℝ F] + [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] + [normed_add_comm_group F] [normed_space ℝ F] theorem dense_compl_of_dimH_lt_finrank {s : set E} (hs : dimH s < finrank ℝ E) : dense sᶜ := begin diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 95b5a4bbe7375..1f421ccfd5d5e 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -481,7 +481,7 @@ open topological_vector_bundle variables (B) variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] - [normed_group F] [normed_space R F] [topological_space B] + [normed_add_comm_group F] [normed_space R F] [topological_space B] [topological_space (total_space E)] [∀ x, topological_space (E x)] /-- The valid transition functions for a topological vector bundle over `B` modelled on @@ -919,7 +919,7 @@ end section variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] - [normed_group F] [normed_space R F] [topological_space B] + [normed_add_comm_group F] [normed_space R F] [topological_space B] open topological_space diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index c9fea76cf57a3..d9a29f2ce10ac 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -75,10 +75,10 @@ variables {𝕜₁ : Type*} [nontrivially_normed_field 𝕜₁] {𝕜₂ : Type* variables {B : Type*} [topological_space B] -variables (F₁ : Type*) [normed_group F₁] [normed_space 𝕜₁ F₁] +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 (F₂ : Type*) [normed_group F₂][normed_space 𝕜₂ F₂] +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₂)] diff --git a/src/topology/vector_bundle/prod.lean b/src/topology/vector_bundle/prod.lean index dc25c0f6daf91..43ced74eaa23e 100644 --- a/src/topology/vector_bundle/prod.lean +++ b/src/topology/vector_bundle/prod.lean @@ -58,11 +58,11 @@ end defs variables [nontrivially_normed_field R] [topological_space B] -variables (F₁ : Type*) [normed_group F₁] [normed_space R F₁] +variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space R F₁] (E₁ : B → Type*) [topological_space (total_space E₁)] [Π x, add_comm_monoid (E₁ x)] [Π x, module R (E₁ x)] -variables (F₂ : Type*) [normed_group F₂] [normed_space R F₂] +variables (F₂ : Type*) [normed_add_comm_group F₂] [normed_space R F₂] (E₂ : B → Type*) [topological_space (total_space E₂)] [Π x, add_comm_monoid (E₂ x)] [Π x, module R (E₂ x)] diff --git a/src/topology/vector_bundle/pullback.lean b/src/topology/vector_bundle/pullback.lean index d57ef9165af11..8074fd62b983c 100644 --- a/src/topology/vector_bundle/pullback.lean +++ b/src/topology/vector_bundle/pullback.lean @@ -66,7 +66,7 @@ begin end variables (F) [nontrivially_normed_field 𝕜] - [normed_group F] [normed_space 𝕜 F] [topological_space B] + [normed_add_comm_group F] [normed_space 𝕜 F] [topological_space B] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] lemma pullback.continuous_total_space_mk [∀ x, topological_space (E x)] diff --git a/test/calc.lean b/test/calc.lean index a527971d154fd..7220fa8fb49c7 100644 --- a/test/calc.lean +++ b/test/calc.lean @@ -8,7 +8,7 @@ section is_equivalent open_locale asymptotics -example {l : filter α} {u v w : α → β} [normed_group β] +example {l : filter α} {u v w : α → β} [normed_add_comm_group β] (huv : u ~[l] v) (hvw : v ~[l] w) : u ~[l] w := calc u ~[l] v : huv ... ~[l] w : hvw diff --git a/test/measurability.lean b/test/measurability.lean index deede74ac58d0..9b874480962ca 100644 --- a/test/measurability.lean +++ b/test/measurability.lean @@ -69,7 +69,7 @@ example [add_comm_monoid β] [has_measurable_add₂ β] {s : finset ℕ} {F : by measurability -- even with many assumptions, the tactic is not trapped by a bad lemma -example [topological_space α] [borel_space α] [normed_group β] [borel_space β] +example [topological_space α] [borel_space α] [normed_add_comm_group β] [borel_space β] [has_measurable_add₂ β] [has_measurable_sub₂ β] {s : finset ℕ} {F : ℕ → α → β} (hF : ∀ i, measurable (F i)) : ae_measurable (∑ i in s, (λ x, F (i+1) x - F i x)) μ := From 95ae6f26ec422be48447ec01d86bc5729b30260c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 24 Jul 2022 16:34:40 +0000 Subject: [PATCH 027/173] refactor(data/{finite,fintype}): move some lemmas to `data.fintype.basic` (#15626) We should have lemmas like `nonempty_fintype` in `fintype.basic` to replace `[fintype _]` by `[finite _]` in the assumptions. This PR doesn't fix any assumptions to keep it small. --- src/data/finite/basic.lean | 31 ----------------------------- src/data/fintype/basic.lean | 39 +++++++++++++++++++++++++++++++++++++ 2 files changed, 39 insertions(+), 31 deletions(-) diff --git a/src/data/finite/basic.lean b/src/data/finite/basic.lean index d5537612fb08f..34dc04a930818 100644 --- a/src/data/finite/basic.lean +++ b/src/data/finite/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ import data.fintype.basic -import data.finite.defs /-! # Finite types @@ -43,25 +42,6 @@ open_locale classical variables {α β γ : Type*} -protected lemma fintype.finite {α : Type*} (h : fintype α) : finite α := ⟨fintype.equiv_fin α⟩ - -/-- For efficiency reasons, we want `finite` instances to have higher -priority than ones coming from `fintype` instances. -/ -@[priority 900] -instance finite.of_fintype (α : Type*) [fintype α] : finite α := fintype.finite ‹_› - -lemma finite_iff_nonempty_fintype (α : Type*) : - finite α ↔ nonempty (fintype α) := -⟨λ h, let ⟨k, ⟨e⟩⟩ := @finite.exists_equiv_fin α h in ⟨fintype.of_equiv _ e.symm⟩, - λ ⟨_⟩, by exactI infer_instance⟩ - -lemma nonempty_fintype (α : Type*) [finite α] : nonempty (fintype α) := -(finite_iff_nonempty_fintype α).mp ‹_› - -/-- Noncomputably get a `fintype` instance from a `finite` instance. This is not an -instance because we want `fintype` instances to be useful for computations. -/ -def fintype.of_finite (α : Type*) [finite α] : fintype α := (nonempty_fintype α).some - lemma not_finite_iff_infinite {α : Type*} : ¬ finite α ↔ infinite α := by rw [← is_empty_fintype, finite_iff_nonempty_fintype, not_nonempty_iff] @@ -99,17 +79,6 @@ lemma exists_min [finite α] [nonempty α] [linear_order β] (f : α → β) : ∃ x₀ : α, ∀ x, f x₀ ≤ f x := by { haveI := fintype.of_finite α, exact fintype.exists_min f } -lemma of_injective {α β : Sort*} [finite β] (f : α → β) (H : function.injective f) : finite α := -begin - haveI := fintype.of_finite (plift β), - rw [← equiv.injective_comp equiv.plift f, ← equiv.comp_injective _ equiv.plift.symm] at H, - haveI := fintype.of_injective _ H, - exact finite.of_equiv _ equiv.plift, -end - -lemma of_surjective {α β : Sort*} [finite α] (f : α → β) (H : function.surjective f) : finite β := -of_injective _ $ function.injective_surj_inv H - @[priority 100] -- see Note [lower instance priority] instance of_is_empty {α : Sort*} [is_empty α] : finite α := finite.of_equiv _ equiv.plift diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 70db1e66434e9..526416a162bec 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -10,6 +10,7 @@ import data.finset.pi import data.finset.powerset import data.finset.prod import data.finset.sigma +import data.finite.defs import data.list.nodup_equiv_fin import data.sym.basic import data.ulift @@ -1016,6 +1017,44 @@ rfl end finset +/-! +### Relation to `finite` + +In this section we prove that `α : Type*` is `finite` if and only if `fintype α` is nonempty. +-/ + +@[nolint fintype_finite] +protected lemma fintype.finite {α : Type*} (h : fintype α) : finite α := ⟨fintype.equiv_fin α⟩ + +/-- For efficiency reasons, we want `finite` instances to have higher +priority than ones coming from `fintype` instances. -/ +@[nolint fintype_finite, priority 900] +instance finite.of_fintype (α : Type*) [fintype α] : finite α := fintype.finite ‹_› + +lemma finite_iff_nonempty_fintype (α : Type*) : + finite α ↔ nonempty (fintype α) := +⟨λ h, let ⟨k, ⟨e⟩⟩ := @finite.exists_equiv_fin α h in ⟨fintype.of_equiv _ e.symm⟩, + λ ⟨_⟩, by exactI infer_instance⟩ + +lemma nonempty_fintype (α : Type*) [finite α] : nonempty (fintype α) := +(finite_iff_nonempty_fintype α).mp ‹_› + +/-- Noncomputably get a `fintype` instance from a `finite` instance. This is not an +instance because we want `fintype` instances to be useful for computations. -/ +noncomputable def fintype.of_finite (α : Type*) [finite α] : fintype α := (nonempty_fintype α).some + +lemma finite.of_injective {α β : Sort*} [finite β] (f : α → β) (H : injective f) : finite α := +begin + casesI nonempty_fintype (plift β), + rw [← equiv.injective_comp equiv.plift f, ← equiv.comp_injective _ equiv.plift.symm] at H, + haveI := fintype.of_injective _ H, + exact finite.of_equiv _ equiv.plift, +end + +lemma finite.of_surjective {α β : Sort*} [finite α] (f : α → β) (H : surjective f) : + finite β := +finite.of_injective _ $ injective_surj_inv H + namespace fintype variables [fintype α] [fintype β] From 5103c95e928b3c1773beab7ce1606d57ff919c7c Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sun, 24 Jul 2022 16:34:41 +0000 Subject: [PATCH 028/173] feat(linear_algebra/tensor_product): add id_apply and comp_apply for ltensor and rtensor (#15628) --- src/linear_algebra/tensor_product.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index ca7ac3854b519..6b89a08161cca 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -872,9 +872,17 @@ def rtensor_hom : (N →ₗ[R] P) →ₗ[R] (N ⊗[R] M →ₗ[R] P ⊗[R] M) := lemma ltensor_comp : (g.comp f).ltensor M = (g.ltensor M).comp (f.ltensor M) := by { ext m n, simp only [compr₂_apply, mk_apply, comp_apply, ltensor_tmul] } +lemma ltensor_comp_apply (x : M ⊗[R] N) : + (g.comp f).ltensor M x = (g.ltensor M) ((f.ltensor M) x) := +by { rw [ltensor_comp, coe_comp], } + lemma rtensor_comp : (g.comp f).rtensor M = (g.rtensor M).comp (f.rtensor M) := by { ext m n, simp only [compr₂_apply, mk_apply, comp_apply, rtensor_tmul] } +lemma rtensor_comp_apply (x : N ⊗[R] M) : + (g.comp f).rtensor M x = (g.rtensor M) ((f.rtensor M) x) := +by { rw [rtensor_comp, coe_comp], } + lemma ltensor_mul (f g : module.End R N) : (f * g).ltensor M = (f.ltensor M) * (g.ltensor M) := ltensor_comp M f g @@ -885,8 +893,16 @@ variables (N) @[simp] lemma ltensor_id : (id : N →ₗ[R] N).ltensor M = id := map_id +-- `simp` can prove this. +lemma ltensor_id_apply (x : M ⊗[R] N) : (linear_map.id : N →ₗ[R] N).ltensor M x = x := +by {rw [ltensor_id, id_coe, id.def], } + @[simp] lemma rtensor_id : (id : N →ₗ[R] N).rtensor M = id := map_id +-- `simp` can prove this. +lemma rtensor_id_apply (x : N ⊗[R] M) : (linear_map.id : N →ₗ[R] N).rtensor M x = x := +by { rw [rtensor_id, id_coe, id.def], } + variables {N} @[simp] lemma ltensor_comp_rtensor (f : M →ₗ[R] P) (g : N →ₗ[R] Q) : From 192e8e4fc00683bc4f60b8896b75bee16853cf41 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 24 Jul 2022 16:34:42 +0000 Subject: [PATCH 029/173] feat(order/bounded_order): an order is either an `order_bot` or a `no_bot_order` (#15636) --- src/order/bounded_order.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index d21adfb9db0c7..6f31b9634e630 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -65,6 +65,18 @@ class order_top (α : Type u) [has_le α] extends has_top α := (le_top : ∀ a : α, a ≤ ⊤) section order_top + +/-- An order is (noncomputably) either an `order_top` or a `no_order_top`. Use as +`casesI bot_order_or_no_bot_order α`. -/ +noncomputable def top_order_or_no_top_order (α : Type*) [has_le α] : + psum (order_top α) (no_top_order α) := +begin + by_cases H : ∀ a : α, ∃ b, ¬ b ≤ a, + { exact psum.inr ⟨H⟩ }, + { push_neg at H, + exact psum.inl ⟨_, classical.some_spec H⟩ } +end + section has_le variables [has_le α] [order_top α] {a : α} @@ -150,6 +162,17 @@ class order_bot (α : Type u) [has_le α] extends has_bot α := section order_bot +/-- An order is (noncomputably) either an `order_bot` or a `no_order_bot`. Use as +`casesI bot_order_or_no_bot_order α`. -/ +noncomputable def bot_order_or_no_bot_order (α : Type*) [has_le α] : + psum (order_bot α) (no_bot_order α) := +begin + by_cases H : ∀ a : α, ∃ b, ¬ a ≤ b, + { exact psum.inr ⟨H⟩ }, + { push_neg at H, + exact psum.inl ⟨_, classical.some_spec H⟩ } +end + section has_le variables [has_le α] [order_bot α] {a : α} From 7de4a19a576bb9f1e918b143d563d8d9e13a80cc Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sun, 24 Jul 2022 16:34:43 +0000 Subject: [PATCH 030/173] feat(algebra/periodic): `fract_periodic` (#15660) Add the lemma that `int.fract` is periodic with period 1. Note that this goes in `algebra.periodic` alongside the definition of `periodic`, rather than `algebra.order.floor` alongside the definition of `fract`, because of import ordering (`algebra.periodic` ends up importing `algebra.order.floor`). --- src/algebra/periodic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 46ee5dbca1e23..a3b282c17f46b 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -523,3 +523,7 @@ lemma antiperiodic.div [has_add α] [division_ring β] by simp [*, neg_div_neg_eq] at * end function + +lemma int.fract_periodic (α) [linear_ordered_ring α] [floor_ring α] : + function.periodic int.fract (1 : α) := +by exact_mod_cast λ a, int.fract_add_int a 1 From 69b51a56d49ed05051ee7773ea9af2d28a839d3c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 24 Jul 2022 23:44:09 +0000 Subject: [PATCH 031/173] =?UTF-8?q?feat(representation=5Ftheory/Rep):=20Re?= =?UTF-8?q?p=20k=20G=20=E2=89=8C=20Module=20(monoid=5Falgebra=20k=20G)=20(?= =?UTF-8?q?#13713)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Antoine Labelle Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> Co-authored-by: Scott Morrison --- src/algebra/algebra/restrict_scalars.lean | 28 +++- src/algebra/module/linear_map.lean | 2 +- src/representation_theory/Rep.lean | 130 ++++++++++++++++- src/representation_theory/basic.lean | 164 ++++++++++++++++++++-- src/representation_theory/invariants.lean | 12 +- 5 files changed, 310 insertions(+), 26 deletions(-) diff --git a/src/algebra/algebra/restrict_scalars.lean b/src/algebra/algebra/restrict_scalars.lean index 217969d7b7f4e..5c3c9856ddb10 100644 --- a/src/algebra/algebra/restrict_scalars.lean +++ b/src/algebra/algebra/restrict_scalars.lean @@ -143,17 +143,37 @@ end variables [add_comm_monoid M] /-- `restrict_scalars.add_equiv` is the additive equivalence with the original module. -/ -@[simps] def restrict_scalars.add_equiv : restrict_scalars R S M ≃+ M := +def restrict_scalars.add_equiv : restrict_scalars R S M ≃+ M := add_equiv.refl M variables [comm_semiring R] [semiring S] [algebra R S] [module S M] +/-- +Note that this lemma relies on the definitional equality `restrict_scalars R S M = M`, +so usage may result in instance leakage. +`restrict_scalars.add_equiv_map_smul` is the "hygienic" version. +-/ lemma restrict_scalars_smul_def (c : R) (x : restrict_scalars R S M) : c • x = ((algebra_map R S c) • x : M) := rfl -@[simp] lemma restrict_scalars.add_equiv_map_smul (t : R) (x : restrict_scalars R S M) : - restrict_scalars.add_equiv R S M (t • x) - = (algebra_map R S t) • restrict_scalars.add_equiv R S M x := +@[simp] lemma restrict_scalars.add_equiv_map_smul (c : R) (x : restrict_scalars R S M) : + restrict_scalars.add_equiv R S M (c • x) + = (algebra_map R S c) • restrict_scalars.add_equiv R S M x := +rfl + +lemma restrict_scalars.add_equiv_symm_map_algebra_map_smul (r : R) (x : M) : + (restrict_scalars.add_equiv R S M).symm (algebra_map R S r • x) + = r • (restrict_scalars.add_equiv R S M).symm x := +rfl + +lemma restrict_scalars.add_equiv_symm_map_smul_smul (r : R) (s : S) (x : M) : + (restrict_scalars.add_equiv R S M).symm ((r • s) • x) + = r • (restrict_scalars.add_equiv R S M ).symm (s • x) := +by { rw [algebra.smul_def, mul_smul], refl, } + +lemma restrict_scalars.lsmul_apply_apply (s : S) (x : restrict_scalars R S M) : + restrict_scalars.lsmul R S M s x = + (restrict_scalars.add_equiv R S M).symm (s • (restrict_scalars.add_equiv R S M x)) := rfl end module diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index 8dc6aaf812bbc..b5df38c3555d4 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -900,7 +900,7 @@ namespace module variables (R M) [semiring R] [add_comm_monoid M] [module R M] variables [semiring S] [module S M] [smul_comm_class S R M] -/-- Each element of the monoid defines a module endomorphism. +/-- Each element of the semiring defines a module endomorphism. This is a stronger version of `distrib_mul_action.to_module_End`. -/ @[simps] diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index ad9e093c6632d..21b1c5278e5d3 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -3,6 +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 representation_theory.basic import representation_theory.Action import algebra.category.Module.abelian import algebra.category.Module.colimits @@ -18,6 +19,7 @@ Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`. Conversely, given a homomorphism `ρ : G →* (V →ₗ[k] V)`, you can construct the bundled representation as `Rep.of ρ`. +We construct the categorical equivalence `Rep k G ≌ Module (monoid_algebra k G)`. We verify that `Rep k G` is a `k`-linear abelian symmetric monoidal category with all (co)limits. -/ @@ -37,24 +39,32 @@ by apply_instance namespace Rep -variables {k G : Type u} [ring k] [monoid G] +variables {k G : Type u} [comm_ring k] [monoid G] instance : has_coe_to_sort (Rep k G) (Type u) := concrete_category.has_coe_to_sort _ -instance (V : Rep k G) : add_comm_monoid V := -by { change add_comm_monoid ((forget₂ (Rep k G) (Module k)).obj V), apply_instance, } +instance (V : Rep k G) : add_comm_group V := +by { change add_comm_group ((forget₂ (Rep k G) (Module k)).obj V), apply_instance, } instance (V : Rep k G) : module k V := by { change module k ((forget₂ (Rep k G) (Module k)).obj V), apply_instance, } --- This works well with the new design for representations: -example (V : Rep k G) : G →* (V →ₗ[k] V) := V.ρ +/-- +Specialize the existing `Action.ρ`, changing the type to `representation k G V`. +-/ +def ρ (V : Rep k G) : representation k G V := V.ρ /-- Lift an unbundled representation to `Rep`. -/ -@[simps ρ] def of {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) : Rep k G := ⟨Module.of k V, ρ⟩ +@[simp] +lemma coe_of {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) : + (of ρ : Type u) = V := rfl + +@[simp] lemma of_ρ {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) : + (of ρ).ρ = ρ := rfl + -- Verify that limits are calculated correctly. noncomputable example : preserves_limits (forget₂ (Rep k G) (Module.{u} k)) := by apply_instance @@ -63,6 +73,9 @@ by apply_instance end Rep +/-! +# The categorical equivalence `Rep k G ≌ Module.{u} (monoid_algebra k G)`. +-/ namespace Rep variables {k G : Type u} [comm_ring k] [monoid G] @@ -71,4 +84,109 @@ example : symmetric_category (Rep k G) := by apply_instance example : monoidal_preadditive (Rep k G) := by apply_instance example : monoidal_linear k (Rep k G) := by apply_instance +noncomputable theory + +/-- Auxilliary lemma for `to_Module_monoid_algebra`. -/ +lemma to_Module_monoid_algebra_map_aux {k G : Type*} [comm_ring k] [monoid G] + (V W : Type*) [add_comm_group V] [add_comm_group W] [module k V] [module k W] + (ρ : G →* V →ₗ[k] V) (σ : G →* W →ₗ[k] W) + (f : V →ₗ[k] W) (w : ∀ (g : G), f.comp (ρ g) = (σ g).comp f) + (r : monoid_algebra k G) (x : V) : + f ((((monoid_algebra.lift k G (V →ₗ[k] V)) ρ) r) x) = + (((monoid_algebra.lift k G (W →ₗ[k] W)) σ) r) (f x) := +begin + apply monoid_algebra.induction_on r, + { intro g, + simp only [one_smul, monoid_algebra.lift_single, monoid_algebra.of_apply], + exact linear_map.congr_fun (w g) x, }, + { intros g h gw hw, simp only [map_add, add_left_inj, linear_map.add_apply, hw, gw], }, + { intros r g w, + simp only [alg_hom.map_smul, w, ring_hom.id_apply, + linear_map.smul_apply, linear_map.map_smulₛₗ], } +end + +/-- Auxilliary definition for `to_Module_monoid_algebra`. -/ +def to_Module_monoid_algebra_map {V W : Rep k G} (f : V ⟶ W) : + Module.of (monoid_algebra k G) V.ρ.as_module ⟶ Module.of (monoid_algebra k G) W.ρ.as_module := +{ map_smul' := λ r x, to_Module_monoid_algebra_map_aux V.V W.V V.ρ W.ρ f.hom f.comm r x, + ..f.hom, } + +/-- Functorially convert a representation of `G` into a module over `monoid_algebra k G`. -/ +def to_Module_monoid_algebra : Rep k G ⥤ Module.{u} (monoid_algebra k G) := +{ obj := λ V, Module.of _ V.ρ.as_module , + map := λ V W f, to_Module_monoid_algebra_map f, } + +/-- Functorially convert a module over `monoid_algebra k G` into a representation of `G`. -/ +def of_Module_monoid_algebra : Module.{u} (monoid_algebra k G) ⥤ Rep k G := +{ obj := λ M, Rep.of (representation.of_module k G M), + map := λ M N f, + { hom := + { map_smul' := λ r x, f.map_smul (algebra_map k _ r) x, + ..f }, + comm' := λ g, by { ext, apply f.map_smul, }, }, }. + +lemma of_Module_monoid_algebra_obj_coe (M : Module.{u} (monoid_algebra k G)) : + (of_Module_monoid_algebra.obj M : Type u) = restrict_scalars k (monoid_algebra k G) M := rfl + +lemma of_Module_monoid_algebra_obj_ρ (M : Module.{u} (monoid_algebra k G)) : + (of_Module_monoid_algebra.obj M).ρ = representation.of_module k G M := rfl + +/-- Auxilliary definition for `equivalence_Module_monoid_algebra`. -/ +def counit_iso_add_equiv {M : Module.{u} (monoid_algebra k G)} : + ((of_Module_monoid_algebra ⋙ to_Module_monoid_algebra).obj M) ≃+ M := +begin + dsimp [of_Module_monoid_algebra, to_Module_monoid_algebra], + refine (representation.of_module k G ↥M).as_module_equiv.trans (restrict_scalars.add_equiv _ _ _), +end + +/-- Auxilliary definition for `equivalence_Module_monoid_algebra`. -/ +def unit_iso_add_equiv {V : Rep k G} : + V ≃+ ((to_Module_monoid_algebra ⋙ of_Module_monoid_algebra).obj V) := +begin + dsimp [of_Module_monoid_algebra, to_Module_monoid_algebra], + refine V.ρ.as_module_equiv.symm.trans _, + exact (restrict_scalars.add_equiv _ _ _).symm, +end + +/-- Auxilliary definition for `equivalence_Module_monoid_algebra`. -/ +def counit_iso (M : Module.{u} (monoid_algebra k G)) : + (of_Module_monoid_algebra ⋙ to_Module_monoid_algebra).obj M ≅ M := +linear_equiv.to_Module_iso' +{ map_smul' := λ r x, begin + dsimp [counit_iso_add_equiv], + simp, + end, + ..counit_iso_add_equiv, } + +lemma unit_iso_comm (V : Rep k G) (g : G) (x : V) : + unit_iso_add_equiv (((V.ρ) g).to_fun x) = + (((of_Module_monoid_algebra.obj (to_Module_monoid_algebra.obj V)).ρ) g).to_fun + (unit_iso_add_equiv x) := +begin + dsimp [unit_iso_add_equiv, of_Module_monoid_algebra, to_Module_monoid_algebra], + simp only [add_equiv.apply_eq_iff_eq, add_equiv.apply_symm_apply, + representation.as_module_equiv_symm_map_rho, representation.of_module_as_module_act], +end + +/-- Auxilliary definition for `equivalence_Module_monoid_algebra`. -/ +def unit_iso (V : Rep k G) : + V ≅ ((to_Module_monoid_algebra ⋙ of_Module_monoid_algebra).obj V) := +Action.mk_iso (linear_equiv.to_Module_iso' +{ map_smul' := λ r x, begin + dsimp [unit_iso_add_equiv], + simp only [representation.as_module_equiv_symm_map_smul, + restrict_scalars.add_equiv_symm_map_algebra_map_smul], + end, + ..unit_iso_add_equiv, }) + (λ g, by { ext, apply unit_iso_comm, }) + +/-- The categorical equivalence `Rep k G ≌ Module (monoid_algebra k G)`. -/ +def equivalence_Module_monoid_algebra : Rep k G ≌ Module.{u} (monoid_algebra k G) := +{ functor := to_Module_monoid_algebra, + inverse := of_Module_monoid_algebra, + unit_iso := nat_iso.of_components (λ V, unit_iso V) (by tidy), + counit_iso := nat_iso.of_components (λ M, counit_iso M) (by tidy), } + +-- TODO Verify that the equivalence with `Module (monoid_algebra k G)` is a monoidal functor. + end Rep diff --git a/src/representation_theory/basic.lean b/src/representation_theory/basic.lean index 456791264b02f..80d5f2988061b 100644 --- a/src/representation_theory/basic.lean +++ b/src/representation_theory/basic.lean @@ -69,29 +69,172 @@ A `k`-linear representation of `G` on `V` can be thought of as an algebra map from `monoid_algebra k G` into the `k`-linear endomorphisms of `V`. -/ noncomputable def as_algebra_hom : monoid_algebra k G →ₐ[k] (module.End k V) := - (lift k G _) ρ +(lift k G _) ρ -lemma as_algebra_hom_def : - as_algebra_hom ρ = (lift k G _) ρ := rfl +lemma as_algebra_hom_def : as_algebra_hom ρ = (lift k G _) ρ := +rfl @[simp] -lemma as_algebra_hom_single (g : G): +lemma as_algebra_hom_single (g : G) (r : k) : + (as_algebra_hom ρ (finsupp.single g r)) = r • ρ g := +by simp only [as_algebra_hom_def, monoid_algebra.lift_single] + +lemma as_algebra_hom_single_one (g : G): (as_algebra_hom ρ (finsupp.single g 1)) = ρ g := -by simp only [as_algebra_hom_def, monoid_algebra.lift_single, one_smul] +by simp -lemma as_algebra_hom_of (g : G): +lemma as_algebra_hom_of (g : G) : (as_algebra_hom ρ (of k G g)) = ρ g := -by simp only [monoid_algebra.of_apply, as_algebra_hom_single] +by simp only [monoid_algebra.of_apply, as_algebra_hom_single, one_smul] + +/-- +If `ρ : representation k G V`, then `ρ.as_module` is a type synonym for `V`, +which we equip with an instance `module (monoid_algebra k G) ρ.as_module`. + +You should use `as_module_equiv : ρ.as_module ≃+ V` to translate terms. +-/ +@[nolint unused_arguments, derive add_comm_monoid] +def as_module (ρ : representation k G V) := V + +instance : inhabited ρ.as_module := ⟨0⟩ /-- A `k`-linear representation of `G` on `V` can be thought of as a module over `monoid_algebra k G`. -/ -noncomputable def as_module : module (monoid_algebra k G) V := - module.comp_hom V (as_algebra_hom ρ).to_ring_hom +noncomputable instance as_module_module : module (monoid_algebra k G) ρ.as_module := +begin + change module (monoid_algebra k G) V, + exact module.comp_hom V (as_algebra_hom ρ).to_ring_hom, +end + +/-- +The additive equivalence from the `module (monoid_algebra k G)` to the original vector space +of the representative. + +This is just the identity, but it is helpful for typechecking and keeping track of instances. +-/ +def as_module_equiv : ρ.as_module ≃+ V := +add_equiv.refl _ + +@[simp] +lemma as_module_equiv_map_smul (r : monoid_algebra k G) (x : ρ.as_module) : + ρ.as_module_equiv (r • x) = ρ.as_algebra_hom r (ρ.as_module_equiv x) := +rfl + +@[simp] +lemma as_module_equiv_symm_map_smul (r : k) (x : V) : + ρ.as_module_equiv.symm (r • x) = + algebra_map k (monoid_algebra k G) r • (ρ.as_module_equiv.symm x) := +begin + apply_fun ρ.as_module_equiv, + simp, +end + +@[simp] +lemma as_module_equiv_symm_map_rho (g : G) (x : V) : + ρ.as_module_equiv.symm (ρ g x) = monoid_algebra.of k G g • (ρ.as_module_equiv.symm x) := +begin + apply_fun ρ.as_module_equiv, + simp, +end + +/-- +Build a `representation k G M` from a `[module (monoid_algebra k G) M]`. + +This version is not always what we want, as it relies on an existing `[module k M]` +instance, along with a `[is_scalar_tower k (monoid_algebra k G) M]` instance. + +We remedy this below in `of_module` +(with the tradeoff that the representation is defined +only on a type synonym of the original module.) +-/ +noncomputable +def of_module' (M : Type*) [add_comm_monoid M] [module k M] [module (monoid_algebra k G) M] + [is_scalar_tower k (monoid_algebra k G) M] : representation k G M := +(monoid_algebra.lift k G (M →ₗ[k] M)).symm (algebra.lsmul k M) + +section +variables (k G) (M : Type*) [add_comm_monoid M] [module (monoid_algebra k G) M] + +/-- +Build a `representation` from a `[module (monoid_algebra k G) M]`. + +Note that the representation is built on `restrict_scalars k (monoid_algebra k G) M`, +rather than on `M` itself. +-/ +noncomputable +def of_module : + representation k G (restrict_scalars k (monoid_algebra k G) M) := +(monoid_algebra.lift k G + (restrict_scalars k (monoid_algebra k G) M →ₗ[k] restrict_scalars k (monoid_algebra k G) M)).symm + (restrict_scalars.lsmul k (monoid_algebra k G) M) + +/-! +## `of_module` and `as_module` are inverses. + +This requires a little care in both directions: +this is a categorical equivalence, not an isomorphism. + +See `Rep.equivalence_Module_monoid_algebra` for the full statement. + +Starting with `ρ : representation k G V`, converting to a module and back again +we have a `representation k G (restrict_scalars k (monoid_algebra k G) ρ.as_module)`. +To compare these, we use the composition of `restrict_scalars_add_equiv` and `ρ.as_module_equiv`. + +Similarly, starting with `module (monoid_algebra k G) M`, +after we convert to a representation and back to a module, +we have `module (monoid_algebra k G) (restrict_scalars k (monoid_algebra k G) M)`. +-/ + +@[simp] lemma of_module_as_algebra_hom_apply_apply + (r : monoid_algebra k G) (m : restrict_scalars k (monoid_algebra k G) M) : + ((((of_module k G M).as_algebra_hom) r) m) = + (restrict_scalars.add_equiv _ _ _).symm (r • restrict_scalars.add_equiv _ _ _ m) := +begin + apply monoid_algebra.induction_on r, + { intros g, + simp only [one_smul, monoid_algebra.lift_symm_apply, monoid_algebra.of_apply, + representation.as_algebra_hom_single, representation.of_module, + add_equiv.apply_eq_iff_eq, restrict_scalars.lsmul_apply_apply], }, + { intros f g fw gw, + simp only [fw, gw, map_add, add_smul, linear_map.add_apply], }, + { intros r f w, + simp only [w, alg_hom.map_smul, linear_map.smul_apply, + restrict_scalars.add_equiv_symm_map_smul_smul], } +end + +@[simp] +lemma of_module_as_module_act (g : G) (x : restrict_scalars k (monoid_algebra k G) ρ.as_module) : + of_module k G (ρ.as_module) g x = + (restrict_scalars.add_equiv _ _ _).symm ((ρ.as_module_equiv).symm + (ρ g (ρ.as_module_equiv (restrict_scalars.add_equiv _ _ _ x)))) := +begin + apply_fun restrict_scalars.add_equiv _ _ ρ.as_module using + (restrict_scalars.add_equiv _ _ _).injective, + dsimp [of_module, restrict_scalars.lsmul_apply_apply], + simp, +end + +lemma smul_of_module_as_module (r : monoid_algebra k G) + (m : (of_module k G M).as_module) : + (restrict_scalars.add_equiv _ _ _) ((of_module k G M).as_module_equiv (r • m)) = + r • (restrict_scalars.add_equiv _ _ _) ((of_module k G M).as_module_equiv m) := +by { dsimp, simp only [add_equiv.apply_symm_apply, of_module_as_algebra_hom_apply_apply], } + +end end monoid_algebra +section add_comm_group + +variables {k G V : Type*} [comm_ring k] [monoid G] [I : add_comm_group V] [module k V] +variables (ρ : representation k G V) + +instance : add_comm_group ρ.as_module := I + +end add_comm_group + section mul_action variables (k : Type*) [comm_semiring k] (G : Type*) [monoid G] (H : Type*) [mul_action G H] @@ -126,7 +269,7 @@ When `G` is a group, a `k`-linear representation of `G` on `V` can be thought of a group homomorphism from `G` into the invertible `k`-linear endomorphisms of `V`. -/ def as_group_hom : G →* units (V →ₗ[k] V) := - monoid_hom.to_hom_units ρ +monoid_hom.to_hom_units ρ lemma as_group_hom_apply (g : G) : ↑(as_group_hom ρ g) = ρ g := by simp only [as_group_hom, monoid_hom.coe_to_hom_units] @@ -177,6 +320,7 @@ def lin_hom : representation k G (V →ₗ[k] W) := map_mul' := λ g h, linear_map.ext $ λ x, by simp_rw [coe_mul, coe_mk, function.comp_apply, mul_inv_rev, map_mul, mul_eq_comp, comp_assoc ]} + @[simp] lemma lin_hom_apply (g : G) (f : V →ₗ[k] W) : (lin_hom ρV ρW) g f = (ρW g) ∘ₗ f ∘ₗ (ρV g⁻¹) := rfl diff --git a/src/representation_theory/invariants.lean b/src/representation_theory/invariants.lean index d535445d622fb..063c9503e5092 100644 --- a/src/representation_theory/invariants.lean +++ b/src/representation_theory/invariants.lean @@ -100,8 +100,8 @@ noncomputable def average_map : V →ₗ[k] V := as_algebra_hom ρ (average k G) The `average_map` sends elements of `V` to the subspace of invariants. -/ theorem average_map_invariant (v : V) : average_map ρ v ∈ invariants ρ := -λ g, by rw [average_map, ←as_algebra_hom_single, ←linear_map.mul_apply, ←map_mul (as_algebra_hom ρ), - mul_average_left] +λ g, by rw [average_map, ←as_algebra_hom_single_one, ←linear_map.mul_apply, + ←map_mul (as_algebra_hom ρ), mul_average_left] /-- The `average_map` acts as the identity on the subspace of invariants. @@ -126,10 +126,12 @@ open category_theory Action variables {k : Type u} [comm_ring k] {G : Group.{u}} lemma mem_invariants_iff_comm {X Y : Rep k G} (f : X.V →ₗ[k] Y.V) (g : G) : - (lin_hom X.ρ Y.ρ) g f = f ↔ X.ρ g ≫ f = f ≫ Y.ρ g := + (lin_hom X.ρ Y.ρ) g f = f ↔ f.comp (X.ρ g) = (Y.ρ g).comp f := begin - rw [lin_hom_apply, ←ρ_Aut_apply_inv, ←linear_map.comp_assoc, ←Module.comp_def, ←Module.comp_def, - iso.inv_comp_eq, ρ_Aut_apply_hom], exact comm, + dsimp, + erw [←ρ_Aut_apply_inv], + rw [←linear_map.comp_assoc, ←Module.comp_def, ←Module.comp_def, iso.inv_comp_eq, ρ_Aut_apply_hom], + exact comm, end /-- The invariants of the representation `lin_hom X.ρ Y.ρ` correspond to the the representation From 479e11a8c3d7d7d513c704630a19ba444f2192ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 25 Jul 2022 01:36:13 +0000 Subject: [PATCH 032/173] =?UTF-8?q?refactor(set=5Ftheory/ordinal/basic):?= =?UTF-8?q?=20`ordinal.min`=20=E2=86=92=20`infi`=20(#14707)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We ditch `ordinal.min` (which is really just `infi`). Apart from this, we add some missing theorems on conditionally complete lattices with a bottom element. --- src/order/conditionally_complete_lattice.lean | 9 ++ src/set_theory/ordinal/basic.lean | 105 +++++++----------- 2 files changed, 47 insertions(+), 67 deletions(-) diff --git a/src/order/conditionally_complete_lattice.lean b/src/order/conditionally_complete_lattice.lean index 39b493810fd5f..b823acb29e85f 100644 --- a/src/order/conditionally_complete_lattice.lean +++ b/src/order/conditionally_complete_lattice.lean @@ -642,6 +642,8 @@ le_is_glb_iff (is_least_Inf hs).is_glb lemma Inf_mem (hs : s.nonempty) : Inf s ∈ s := (is_least_Inf hs).1 +lemma infi_mem [nonempty ι] (f : ι → α) : infi f ∈ range f := Inf_mem (range_nonempty f) + lemma monotone_on.map_Inf {β : Type*} [conditionally_complete_lattice β] {f : α → β} (hf : monotone_on f s) (hs : s.nonempty) : f (Inf s) = Inf (f '' s) := (hf.map_is_least (is_least_Inf hs)).cInf_eq.symm @@ -697,9 +699,16 @@ theorem le_cInf_iff'' {s : set α} {a : α} (ne : s.nonempty) : a ≤ Inf s ↔ ∀ (b : α), b ∈ s → a ≤ b := le_cInf_iff ⟨⊥, λ a _, bot_le⟩ ne +theorem le_cinfi_iff' [nonempty ι] {f : ι → α} {a : α} : + a ≤ infi f ↔ ∀ i, a ≤ f i := +le_cinfi_iff ⟨⊥, λ a _, bot_le⟩ + theorem cInf_le' {s : set α} {a : α} (h : a ∈ s) : Inf s ≤ a := cInf_le ⟨⊥, λ a _, bot_le⟩ h +theorem cinfi_le' (f : ι → α) (i : ι) : infi f ≤ f i := +cinfi_le ⟨⊥, λ a _, bot_le⟩ _ + lemma exists_lt_of_lt_cSup' {s : set α} {a : α} (h : a < Sup s) : ∃ b ∈ s, a < b := by { contrapose! h, exact cSup_le' h } diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index c5030e43062cc..0a4f52e70e0c9 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -39,7 +39,6 @@ initial segment (or, equivalently, in any way). This total order is well founded we only introduce it and prove its basic properties to deduce the fact that the order on ordinals is total (and well founded). * `succ o` is the successor of the ordinal `o`. -* `ordinal.min`: the minimal element of a nonempty indexed family of ordinals * `cardinal.ord c`: when `c` is a cardinal, `ord c` is the smallest ordinal with this cardinality. It is the canonical way to represent a cardinal with an ordinal. @@ -84,6 +83,9 @@ def well_ordering_rel : σ → σ → Prop := embedding_to_cardinal ⁻¹'o (<) instance well_ordering_rel.is_well_order : is_well_order σ well_ordering_rel := (rel_embedding.preimage _ _).is_well_order +instance is_well_order.subtype_nonempty : nonempty {r // is_well_order σ r} := +⟨⟨well_ordering_rel, infer_instance⟩⟩ + end well_ordering_thm /-! ### Definition of ordinals -/ @@ -680,6 +682,24 @@ instance : linear_order ordinal := instance : is_well_order ordinal (<) := ⟨lt_wf⟩ +instance : conditionally_complete_linear_order_bot ordinal := +is_well_order.conditionally_complete_linear_order_bot _ + +@[simp] lemma bot_eq_zero : (⊥ : ordinal) = 0 := rfl + +@[simp] lemma max_zero_left : ∀ a : ordinal, max 0 a = a := max_bot_left +@[simp] lemma max_zero_right : ∀ a : ordinal, max a 0 = a := max_bot_right +@[simp] lemma max_eq_zero {a b : ordinal} : max a b = 0 ↔ a = 0 ∧ b = 0 := max_eq_bot + +protected theorem not_lt_zero (o : ordinal) : ¬ o < 0 := +not_lt_bot + +theorem eq_zero_or_pos : ∀ a : ordinal, a = 0 ∨ 0 < a := +eq_bot_or_bot_lt + +@[simp] theorem Inf_empty : Inf (∅ : set ordinal) = 0 := +dif_neg not_nonempty_empty + private theorem succ_le_iff' {a b : ordinal} : a + 1 ≤ b ↔ a < b := ⟨lt_of_lt_of_le (induction_on a $ λ α r _, ⟨⟨⟨⟨λ x, sum.inl x, λ _ _, sum.inl.inj⟩, λ _ _, sum.lex_inl_inl⟩, @@ -815,46 +835,6 @@ theorem lift.principal_seg_top' : lift.principal_seg.{u (u+1)}.top = @type ordinal (<) _ := by simp only [lift.principal_seg_top, univ_id] -/-! ### Minimum -/ - -/-- The minimal element of a nonempty family of ordinals -/ -def min {ι} (I : nonempty ι) (f : ι → ordinal) : ordinal := -lt_wf.min (set.range f) (let ⟨i⟩ := I in ⟨_, set.mem_range_self i⟩) - -theorem min_eq {ι} (I) (f : ι → ordinal) : ∃ i, min I f = f i := -let ⟨i, e⟩ := lt_wf.min_mem (set.range f) _ in ⟨i, e.symm⟩ - -theorem min_le {ι I} (f : ι → ordinal) (i) : min I f ≤ f i := -le_of_not_gt $ lt_wf.not_lt_min (set.range f) _ (set.mem_range_self i) - -theorem le_min {ι I} {f : ι → ordinal} {a} : a ≤ min I f ↔ ∀ i, a ≤ f i := -⟨λ h i, le_trans h (min_le _ _), - λ h, let ⟨i, e⟩ := min_eq I f in e.symm ▸ h i⟩ - -@[simp] theorem lift_min {ι} (I) (f : ι → ordinal) : lift (min I f) = min I (lift ∘ f) := -le_antisymm (le_min.2 $ λ a, lift_le.2 $ min_le _ a) $ -let ⟨i, e⟩ := min_eq I (lift ∘ f) in -by rw e; exact lift_le.2 (le_min.2 $ λ j, lift_le.1 $ -by have := min_le (lift ∘ f) j; rwa e at this) - -instance : conditionally_complete_linear_order_bot ordinal := -is_well_order.conditionally_complete_linear_order_bot _ - -@[simp] lemma bot_eq_zero : (⊥ : ordinal) = 0 := rfl - -@[simp] lemma max_zero_left : ∀ a : ordinal, max 0 a = a := max_bot_left -@[simp] lemma max_zero_right : ∀ a : ordinal, max a 0 = a := max_bot_right -@[simp] lemma max_eq_zero {a b : ordinal} : max a b = 0 ↔ a = 0 ∧ b = 0 := max_eq_bot - -protected theorem not_lt_zero (o : ordinal) : ¬ o < 0 := -not_lt_bot - -theorem eq_zero_or_pos : ∀ a : ordinal, a = 0 ∨ 0 < a := -eq_bot_or_bot_lt - -@[simp] theorem Inf_empty : Inf (∅ : set ordinal) = 0 := -dif_neg not_nonempty_empty - end ordinal /-! ### Representing a cardinal with an ordinal -/ @@ -868,36 +848,27 @@ open ordinal /-- The ordinal corresponding to a cardinal `c` is the least ordinal whose cardinal is `c`. For the order-embedding version, see `ord.order_embedding`. -/ def ord (c : cardinal) : ordinal := +let F := λ α : Type u, ⨅ r : {r // is_well_order α r}, @type α r.1 r.2 in +quot.lift_on c F begin - let ι := λ α, {r // is_well_order α r}, - have : Π α, ι α := λ α, ⟨well_ordering_rel, by apply_instance⟩, - let F := λ α, ordinal.min ⟨this _⟩ (λ i:ι α, ⟦⟨α, i.1, i.2⟩⟧), - refine quot.lift_on c F _, suffices : ∀ {α β}, α ≈ β → F α ≤ F β, - from λ α β h, le_antisymm (this h) (this (setoid.symm h)), - intros α β h, cases h with f, refine ordinal.le_min.2 (λ i, _), - haveI := @rel_embedding.is_well_order _ _ - (f ⁻¹'o i.1) _ ↑(rel_iso.preimage f i.1) i.2, - rw ← show type (f ⁻¹'o i.1) = ⟦⟨β, i.1, i.2⟩⟧, from - quot.sound ⟨rel_iso.preimage f i.1⟩, - exact ordinal.min_le (λ i:ι α, ⟦⟨α, i.1, i.2⟩⟧) ⟨_, _⟩ + from λ α β h, (this h).antisymm (this (setoid.symm h)), + rintros α β ⟨f⟩, + refine le_cinfi_iff'.2 (λ i, _), + haveI := @rel_embedding.is_well_order _ _ (f ⁻¹'o i.1) _ ↑(rel_iso.preimage f i.1) i.2, + exact (cinfi_le' _ (subtype.mk (⇑f ⁻¹'o i.val) + (@rel_embedding.is_well_order _ _ _ _ ↑(rel_iso.preimage f i.1) i.2))).trans_eq + (quot.sound ⟨rel_iso.preimage f i.1⟩) end -lemma ord_eq_min (α : Type u) : ord (#α) = - @ordinal.min {r // is_well_order α r} ⟨⟨well_ordering_rel, by apply_instance⟩⟩ - (λ i, ⟦⟨α, i.1, i.2⟩⟧) := rfl - -theorem ord_eq (α) : ∃ (r : α → α → Prop) [wo : is_well_order α r], - ord (#α) = @type α r wo := -let ⟨⟨r, wo⟩, h⟩ := @ordinal.min_eq {r // is_well_order α r} - ⟨⟨well_ordering_rel, by apply_instance⟩⟩ - (λ i:{r // is_well_order α r}, ⟦⟨α, i.1, i.2⟩⟧) in -⟨r, wo, h⟩ - -theorem ord_le_type (r : α → α → Prop) [is_well_order α r] : ord (#α) ≤ ordinal.type r := -@ordinal.min_le {r // is_well_order α r} - ⟨⟨well_ordering_rel, by apply_instance⟩⟩ - (λ i:{r // is_well_order α r}, ⟦⟨α, i.1, i.2⟩⟧) ⟨r, _⟩ +lemma ord_eq_Inf (α : Type u) : ord (#α) = ⨅ r : {r // is_well_order α r}, @type α r.1 r.2 := +rfl + +theorem ord_eq (α) : ∃ (r : α → α → Prop) [wo : is_well_order α r], ord (#α) = @type α r wo := +let ⟨r, wo⟩ := infi_mem (λ r : {r // is_well_order α r}, @type α r.1 r.2) in ⟨r.1, r.2, wo.symm⟩ + +theorem ord_le_type (r : α → α → Prop) [h : is_well_order α r] : ord (#α) ≤ type r := +cinfi_le' _ (subtype.mk r h) theorem ord_le {c o} : ord c ≤ o ↔ c ≤ o.card := induction_on c $ λ α, ordinal.induction_on o $ λ β s _, From 46ca867f1dc37b129eb543285e2ab15bbc645b69 Mon Sep 17 00:00:00 2001 From: JonBannon Date: Mon, 25 Jul 2022 01:36:14 +0000 Subject: [PATCH 033/173] feat(analysis/normed_space/lp_space): construct star structures on lp spaces (#15317) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit On general `lp` spaces we construct a star add monoid and star module structure, and in the case when `1 ≤ p` we also show it is a normed star group. In addition, for `p = ∞`, we provide a star ring structure and show that it is a C⋆-ring. This establishes that the (\ell ^ \infty) direct sum of C⋆-algebras is itself a C⋆-algebra with the supremum norm. Co-authored-by: JonBannon <59937998+JonBannon@users.noreply.github.com> Co-authored-by: Jireh Loreaux --- src/analysis/normed_space/lp_space.lean | 76 ++++++++++++++++++++++++- 1 file changed, 75 insertions(+), 1 deletion(-) diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 5a5ca99356ccc..6a644b03008e6 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -30,7 +30,8 @@ The space `lp E p` is the subtype of elements of `Π i : α, E i` which satisfy * `lp E p` : elements of `Π i : α, E i` such that `mem_ℓp f p`. Defined as an `add_subgroup` of a type synonym `pre_lp` for `Π i : α, E i`, and equipped with a `normed_add_comm_group` structure. Under appropriate conditions, this is also equipped with the instances `lp.normed_space`, - `lp.complete_space`. For `p=∞`, there is also `lp.infty_normed_ring`, `lp.infty_normed_algebra`. + `lp.complete_space`. For `p=∞`, there is also `lp.infty_normed_ring`, + `lp.infty_normed_algebra`, `lp.infty_star_ring` and `lp.infty_cstar_ring`. ## Main results @@ -629,6 +630,53 @@ end end normed_space +section normed_star_group + +variables [Π i, star_add_monoid (E i)] [Π i, normed_star_group (E i)] + +lemma _root_.mem_ℓp.star_mem {f : Π i, E i} + (hf : mem_ℓp f p) : mem_ℓp (star f) p := +begin + rcases p.trichotomy with rfl | rfl | hp, + { apply mem_ℓp_zero, + simp [hf.finite_dsupport] }, + { apply mem_ℓp_infty, + simpa using hf.bdd_above }, + { apply mem_ℓp_gen, + simpa using hf.summable hp }, +end + +@[simp] lemma _root_.mem_ℓp.star_iff {f : Π i, E i} : mem_ℓp (star f) p ↔ mem_ℓp f p := +⟨λ h, star_star f ▸ mem_ℓp.star_mem h ,mem_ℓp.star_mem⟩ + +instance : has_star (lp E p) := +{ star := λ f, ⟨(star f : Π i, E i), f.property.star_mem⟩} + +@[simp] lemma coe_fn_star (f : lp E p) : ⇑(star f) = star f := rfl +@[simp] protected theorem star_apply (f : lp E p) (i : α) : star f i = star (f i) := rfl + +instance : has_involutive_star (lp E p) := { star_involutive := λ x, by {ext, simp} } + +instance : star_add_monoid (lp E p) := { star_add := λ f g, ext $ star_add _ _ } + +instance [hp : fact (1 ≤ p)] : normed_star_group (lp E p) := +{ norm_star := λ f, + begin + unfreezingI { rcases p.trichotomy with rfl | rfl | h }, + { exfalso, + have := ennreal.to_real_mono ennreal.zero_ne_top hp.elim, + norm_num at this,}, + { simp only [lp.norm_eq_csupr, lp.star_apply, norm_star] }, + { 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)] + +instance : star_module 𝕜 (lp E p) := { star_smul := λ r f, ext $ star_smul _ _ } + +end normed_star_group + section non_unital_normed_ring variables {I : Type*} {B : I → Type*} [Π i, non_unital_normed_ring (B i)] @@ -675,6 +723,32 @@ instance infty_smul_comm_class {𝕜} [normed_field 𝕜] [Π i, normed_space smul_comm_class 𝕜 (lp B ∞) (lp B ∞) := ⟨λ r f g, lp.ext $ smul_comm r ⇑f ⇑g⟩ +section star_ring + +variables [Π i, star_ring (B i)] [Π i, normed_star_group (B i)] + +instance infty_star_ring : star_ring (lp B ∞) := +{ star_mul := λ f g, ext $ star_mul (_ : Π i, B i) _, + .. (show star_add_monoid (lp B ∞), + by { letI : Π i, star_add_monoid (B i) := λ i, infer_instance, apply_instance }) } + +instance infty_cstar_ring [∀ i, cstar_ring (B i)] : cstar_ring (lp B ∞) := +{ norm_star_mul_self := λ f, + begin + apply le_antisymm, + { rw ←sq, + refine lp.norm_le_of_forall_le (sq_nonneg ∥ f ∥) (λ i, _), + simp only [lp.star_apply, cstar_ring.norm_star_mul_self, ←sq, infty_coe_fn_mul, pi.mul_apply], + refine sq_le_sq' _ (lp.norm_apply_le_norm ennreal.top_ne_zero _ _), + linarith [norm_nonneg (f i), norm_nonneg f] }, + { rw [←sq, ←real.le_sqrt (norm_nonneg _) (norm_nonneg _)], + refine lp.norm_le_of_forall_le (∥star f * f∥.sqrt_nonneg) (λ i, _), + rw [real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ←cstar_ring.norm_star_mul_self], + exact lp.norm_apply_le_norm ennreal.top_ne_zero (star f * f) i, } + end } + +end star_ring + end non_unital_normed_ring section normed_ring From 21c56ff23763e66ea7609c9d077ed3975e3e60f6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Jul 2022 01:36:15 +0000 Subject: [PATCH 034/173] chore(algebra/order/sub): Generalize lemmas (#15497) Generalize many lemmas from `canonically_ordered_add_monoid` to `has_exists_add_of_le`, and a few other generalizations. ### New lemmas * `mul_le_cancellable_one`/`add_le_cancellable_zero` * `tsub_add_le_right_comm` * `add_tsub_add_le_tsub_left` * `add_tsub_add_le_tsub_right` --- src/algebra/big_operators/multiset.lean | 5 +- src/algebra/order/monoid_lemmas.lean | 3 + src/algebra/order/ring.lean | 2 +- src/algebra/order/sub.lean | 270 ++++++++++----------- src/data/list/basic.lean | 2 +- src/data/nat/fib.lean | 2 +- src/data/real/nnreal.lean | 2 +- src/order/order_iso_nat.lean | 2 +- src/order/well_founded_set.lean | 2 +- src/ring_theory/polynomial/eisenstein.lean | 6 +- 10 files changed, 146 insertions(+), 150 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index a57e533d9e30f..15e7e6b55d0c8 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -153,10 +153,7 @@ lemma dvd_prod : a ∈ s → a ∣ s.prod := quotient.induction_on s (λ l a h, by simpa using list.dvd_prod h) a lemma prod_dvd_prod_of_le (h : s ≤ t) : s.prod ∣ t.prod := -begin - obtain ⟨z, rfl⟩ := multiset.le_iff_exists_add.1 h, - simp only [prod_add, dvd_mul_right], -end +by { obtain ⟨z, rfl⟩ := exists_add_of_le h, simp only [prod_add, dvd_mul_right] } end comm_monoid diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid_lemmas.lean index e2612bc064c16..629325700a6f4 100644 --- a/src/algebra/order/monoid_lemmas.lean +++ b/src/algebra/order/monoid_lemmas.lean @@ -1068,6 +1068,9 @@ lemma contravariant.mul_le_cancellable [has_mul α] [has_le α] [contravariant_c {a : α} : mul_le_cancellable a := λ b c, le_of_mul_le_mul_left' +@[to_additive] lemma mul_le_cancellable_one [monoid α] [has_le α] : mul_le_cancellable (1 : α) := +λ a b, by simpa only [one_mul] using id + namespace mul_le_cancellable @[to_additive] diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index b2cb5e639bc26..cc7b63bda7912 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -1514,7 +1514,7 @@ instance to_no_zero_divisors : no_zero_divisors α := instance to_covariant_mul_le : covariant_class α α (*) (≤) := begin refine ⟨λ a b c h, _⟩, - rcases le_iff_exists_add.1 h with ⟨c, rfl⟩, + rcases exists_add_of_le h with ⟨c, rfl⟩, rw mul_add, apply self_le_add_right end diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub.lean index 921e49814aad5..b0e25dd900403 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub.lean @@ -94,7 +94,7 @@ end /-! ### Preorder -/ -section ordered_add_comm_monoid +section ordered_add_comm_semigroup section preorder variables [preorder α] @@ -135,6 +135,10 @@ lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } +/-- See `tsub_add_eq_add_tsub` for the equality. -/ +lemma add_tsub_le_tsub_add : a + b - c ≤ a - c + b := +by { rw [add_comm, add_comm _ b], exact add_tsub_le_assoc } + lemma add_le_add_add_tsub : a + b ≤ (a + c) + (b - c) := by { rw [add_assoc], exact add_le_add_left le_add_tsub a } @@ -159,6 +163,23 @@ tsub_le_iff_right.2 $ calc ... ≤ a - b + (c + (b - c)) : add_le_add_left le_add_tsub _ ... = a - b + c + (b - c) : (add_assoc _ _ _).symm +/-- See `tsub_add_tsub_comm` for the equality. -/ +lemma add_tsub_add_le_tsub_add_tsub : a + b - (c + d) ≤ a - c + (b - d) := +begin + rw [add_comm c, tsub_le_iff_left, add_assoc, ←tsub_le_iff_left, ←tsub_le_iff_left], + refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, + rw [add_comm a, add_comm (a - c)], + exact add_tsub_le_assoc, +end + +/-- See `add_tsub_add_eq_tsub_left` for the equality. -/ +lemma add_tsub_add_le_tsub_left : a + b - (a + c) ≤ b - c := +by { rw [tsub_le_iff_left, add_assoc], exact add_le_add_left le_add_tsub _ } + +/-- See `add_tsub_add_eq_tsub_right` for the equality. -/ +lemma add_tsub_add_le_tsub_right : a + c - (b + c) ≤ a - b := +by { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c } + end cov /-! #### Lemmas that assume that an element is `add_le_cancellable` -/ @@ -221,16 +242,7 @@ begin { rw [tsub_le_iff_left, add_assoc, ← tsub_le_iff_left, ← tsub_le_iff_left] } end -section cov -variable [covariant_class α α (+) (≤)] - -lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := -begin - refine le_antisymm (tsub_le_iff_left.mpr _) - (tsub_le_iff_left.mpr $ tsub_le_iff_left.mpr _), - { rw [add_assoc], refine le_trans le_add_tsub (add_le_add_left le_add_tsub _) }, - { rw [← add_assoc], apply le_add_tsub } -end +lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := (tsub_tsub _ _ _).symm lemma tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := by { rw [add_comm], apply tsub_add_eq_tsub_tsub } @@ -238,17 +250,6 @@ by { rw [add_comm], apply tsub_add_eq_tsub_tsub } lemma tsub_right_comm : a - b - c = a - c - b := by simp_rw [← tsub_add_eq_tsub_tsub, add_comm] -lemma add_tsub_add_le_tsub_add_tsub : - (a + b) - (c + d) ≤ (a - c) + (b - d) := -begin - rw [add_comm c, ← tsub_tsub], - refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, - rw [add_comm a, add_comm (a - c)], - exact add_tsub_le_assoc -end - -end cov - /-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ namespace add_le_cancellable @@ -287,6 +288,12 @@ begin simpa [hc] using h, end +protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := +(hc.le_tsub_of_add_le_right h.le).lt_of_ne $ by { rintro rfl, exact h.not_le le_tsub_add } + +protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := +ha.lt_tsub_of_add_lt_right $ by rwa add_comm + end add_le_cancellable /-! #### Lemmas where addition is order-reflecting. -/ @@ -317,6 +324,14 @@ contravariant.add_le_cancellable.lt_add_of_tsub_lt_left h lemma lt_add_of_tsub_lt_right (h : a - c < b) : a < b + c := contravariant.add_le_cancellable.lt_add_of_tsub_lt_right h +/-- This lemma (and some of its corollaries) also holds for `ennreal`, but this proof doesn't work +for it. Maybe we should add this lemma as field to `has_ordered_sub`? -/ +lemma lt_tsub_of_add_lt_left : a + c < b → c < b - a := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_left + +lemma lt_tsub_of_add_lt_right : a + c < b → a < b - c := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_right + end contra section both @@ -324,12 +339,9 @@ variables [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤) lemma add_tsub_add_eq_tsub_right (a c b : α) : (a + c) - (b + c) = a - b := begin - apply le_antisymm, - { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c }, - { rw [tsub_le_iff_left, add_comm b], - apply le_of_add_le_add_right, - rw [add_assoc], - exact le_tsub_add } + refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 $ le_of_add_le_add_right _), swap, + rw add_assoc, + exact le_tsub_add, end lemma add_tsub_add_eq_tsub_left (a b c : α) : (a + b) - (a + c) = b - c := @@ -337,7 +349,7 @@ by rw [add_comm a b, add_comm a c, add_tsub_add_eq_tsub_right] end both -end ordered_add_comm_monoid +end ordered_add_comm_semigroup /-! ### Lemmas in a linearly ordered monoid. -/ section linear_order @@ -358,8 +370,6 @@ lt_iff_lt_of_le_iff_le tsub_le_iff_left lemma lt_tsub_comm : a < b - c ↔ c < b - a := lt_tsub_iff_left.trans lt_tsub_iff_right.symm - - section cov variable [covariant_class α α (+) (≤)] @@ -371,27 +381,28 @@ end cov end linear_order -/-! ### Lemmas in a canonically ordered monoid. -/ +section ordered_add_comm_monoid +variables [partial_order α] [add_comm_monoid α] [has_sub α] [has_ordered_sub α] -section canonically_ordered_add_monoid -variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} +@[simp] lemma tsub_zero (a : α) : a - 0 = a := +add_le_cancellable.tsub_eq_of_eq_add add_le_cancellable_zero (add_zero _).symm + +end ordered_add_comm_monoid + +section has_exists_add_of_le +variables [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] + [covariant_class α α (+) (≤)] [has_sub α] [has_ordered_sub α] {a b c d : α} @[simp] lemma add_tsub_cancel_of_le (h : a ≤ b) : a + (b - a) = b := begin refine le_antisymm _ le_add_tsub, - obtain ⟨c, rfl⟩ := le_iff_exists_add.1 h, + obtain ⟨c, rfl⟩ := exists_add_of_le h, exact add_le_add_left add_tsub_le_left a, end lemma tsub_add_cancel_of_le (h : a ≤ b) : b - a + a = b := by { rw [add_comm], exact add_tsub_cancel_of_le h } -lemma add_tsub_cancel_iff_le : a + (b - a) = b ↔ a ≤ b := -⟨λ h, le_iff_exists_add.mpr ⟨b - a, h.symm⟩, add_tsub_cancel_of_le⟩ - -lemma tsub_add_cancel_iff_le : b - a + a = b ↔ a ≤ b := -by { rw [add_comm], exact add_tsub_cancel_iff_le } - lemma add_le_of_le_tsub_right_of_le (h : b ≤ c) (h2 : a ≤ c - b) : a + b ≤ c := (add_le_add_right h2 b).trans_eq $ tsub_add_cancel_of_le h @@ -404,63 +415,28 @@ by rw [tsub_le_iff_right, tsub_add_cancel_of_le h] lemma tsub_left_inj (h1 : c ≤ a) (h2 : c ≤ b) : a - c = b - c ↔ a = b := by simp_rw [le_antisymm_iff, tsub_le_tsub_iff_right h1, tsub_le_tsub_iff_right h2] +lemma tsub_inj_left (h₁ : a ≤ b) (h₂ : a ≤ c) : b - a = c - a → b = c := (tsub_left_inj h₁ h₂).1 + /-- See `lt_of_tsub_lt_tsub_right` for a stronger statement in a linear order. -/ lemma lt_of_tsub_lt_tsub_right_of_le (h : c ≤ b) (h2 : a - c < b - c) : a < b := by { refine ((tsub_le_tsub_iff_right h).mp h2.le).lt_of_ne _, rintro rfl, exact h2.false } -@[simp] lemma tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := -by rw [← nonpos_iff_eq_zero, tsub_le_iff_left, add_zero] - -/-- One direction of `tsub_eq_zero_iff_le`, as a `@[simp]`-lemma. -/ -@[simp] lemma tsub_eq_zero_of_le (h : a ≤ b) : a - b = 0 := -tsub_eq_zero_iff_le.mpr h - -@[simp] lemma tsub_self (a : α) : a - a = 0 := -tsub_eq_zero_iff_le.mpr le_rfl - -@[simp] lemma tsub_le_self : a - b ≤ a := -tsub_le_iff_left.mpr $ le_add_left le_rfl - -@[simp] lemma tsub_zero (a : α) : a - 0 = a := -le_antisymm tsub_le_self $ le_add_tsub.trans_eq $ zero_add _ - -@[simp] lemma zero_tsub (a : α) : 0 - a = 0 := -tsub_eq_zero_iff_le.mpr $ zero_le a - -lemma tsub_self_add (a b : α) : a - (a + b) = 0 := -by { rw [tsub_eq_zero_iff_le], apply self_le_add_right } - -lemma tsub_inj_left (h₁ : a ≤ b) (h₂ : a ≤ c) (h₃ : b - a = c - a) : b = c := -by rw [← tsub_add_cancel_of_le h₁, ← tsub_add_cancel_of_le h₂, h₃] - -lemma tsub_pos_iff_not_le : 0 < a - b ↔ ¬ a ≤ b := -by rw [pos_iff_ne_zero, ne.def, tsub_eq_zero_iff_le] - -lemma tsub_pos_of_lt (h : a < b) : 0 < b - a := -tsub_pos_iff_not_le.mpr h.not_le - -lemma tsub_add_tsub_cancel (hab : b ≤ a) (hbc : c ≤ b) : (a - b) + (b - c) = a - c := +lemma tsub_add_tsub_cancel (hab : b ≤ a) (hcb : c ≤ b) : (a - b) + (b - c) = a - c := begin convert tsub_add_cancel_of_le (tsub_le_tsub_right hab c) using 2, - rw [tsub_tsub, add_tsub_cancel_of_le hbc], + rw [tsub_tsub, add_tsub_cancel_of_le hcb], end lemma tsub_tsub_tsub_cancel_right (h : c ≤ b) : (a - c) - (b - c) = a - b := by rw [tsub_tsub, add_tsub_cancel_of_le h] -lemma tsub_lt_of_lt (h : a < b) : a - c < b := -lt_of_le_of_lt tsub_le_self h - -/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ +/-! #### Lemmas that assume that an element is `add_le_cancellable`. -/ namespace add_le_cancellable + protected lemma eq_tsub_iff_add_eq_of_le (hc : add_le_cancellable c) (h : c ≤ b) : a = b - c ↔ a + c = b := -begin - split, - { rintro rfl, exact tsub_add_cancel_of_le h }, - { rintro rfl, exact (hc.add_tsub_cancel_right).symm } -end +⟨by { rintro rfl, exact tsub_add_cancel_of_le h }, hc.eq_tsub_of_add_eq⟩ protected lemma tsub_eq_iff_eq_add_of_le (hb : add_le_cancellable b) (h : b ≤ a) : a - b = c ↔ a = c + b := @@ -477,8 +453,7 @@ by rw [add_comm a, hb.add_tsub_assoc_of_le h, add_comm] protected lemma tsub_tsub_assoc (hbc : add_le_cancellable (b - c)) (h₁ : b ≤ a) (h₂ : c ≤ b) : a - (b - c) = a - b + c := -by rw [hbc.tsub_eq_iff_eq_add_of_le (tsub_le_self.trans h₁), add_assoc, - add_tsub_cancel_of_le h₂, tsub_add_cancel_of_le h₁] +hbc.tsub_eq_of_eq_add $ by rw [add_assoc, add_tsub_cancel_of_le h₂, tsub_add_cancel_of_le h₁] protected lemma tsub_add_tsub_comm (hb : add_le_cancellable b) (hd : add_le_cancellable d) (hba : b ≤ a) (hdc : d ≤ c) : a - b + (c - d) = a + c - (b + d) := @@ -502,17 +477,6 @@ protected lemma tsub_lt_iff_right (hb : add_le_cancellable b) (hba : b ≤ a) : a - b < c ↔ a < c + b := by { rw [add_comm], exact hb.tsub_lt_iff_left hba } -protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := -begin - apply lt_of_le_of_ne, - { rw [← add_tsub_cancel_of_le h.le, add_right_comm, add_assoc], - rw [hc.add_tsub_assoc_of_le], refine le_self_add, refine le_add_self }, - { rintro rfl, apply h.not_le, exact le_tsub_add } -end - -protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := -by { apply ha.lt_tsub_of_add_lt_right, rwa add_comm } - protected lemma tsub_lt_iff_tsub_lt (hb : add_le_cancellable b) (hc : add_le_cancellable c) (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < c ↔ a - c < b := by rw [hb.tsub_lt_iff_left h₁, hc.tsub_lt_iff_right h₂] @@ -524,12 +488,9 @@ by rw [ha.le_tsub_iff_left h₁, hc.le_tsub_iff_right h₂] protected lemma lt_tsub_iff_right_of_le (hc : add_le_cancellable c) (h : c ≤ b) : a < b - c ↔ a + c < b := begin - refine ⟨_, hc.lt_tsub_of_add_lt_right⟩, - intro h2, - refine (add_le_of_le_tsub_right_of_le h h2.le).lt_of_ne _, + refine ⟨λ h', (add_le_of_le_tsub_right_of_le h h'.le).lt_of_ne _, hc.lt_tsub_of_add_lt_right⟩, rintro rfl, - apply h2.not_le, - rw [hc.add_tsub_cancel_right] + exact h'.ne' hc.add_tsub_cancel_right, end protected lemma lt_tsub_iff_left_of_le (hc : add_le_cancellable c) (h : c ≤ b) : @@ -543,20 +504,6 @@ begin exact lt_of_add_lt_add_left (hb.lt_add_of_tsub_lt_right h), end -protected lemma tsub_le_tsub_iff_left (ha : add_le_cancellable a) (hc : add_le_cancellable c) - (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := -begin - refine ⟨_, λ h, tsub_le_tsub_left h a⟩, - rw [tsub_le_iff_left, ← hc.add_tsub_assoc_of_le h, - hc.le_tsub_iff_right (h.trans le_add_self), add_comm b], - apply ha, -end - -protected lemma tsub_right_inj (ha : add_le_cancellable a) (hb : add_le_cancellable b) - (hc : add_le_cancellable c) (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := -by simp_rw [le_antisymm_iff, ha.tsub_le_tsub_iff_left hb hba, ha.tsub_le_tsub_iff_left hc hca, - and_comm] - protected lemma tsub_lt_tsub_right_of_le (hc : add_le_cancellable c) (h : c ≤ a) (h2 : a < b) : a - c < b - c := by { apply hc.lt_tsub_of_add_lt_left, rwa [add_tsub_cancel_of_le h] } @@ -576,12 +523,11 @@ end @[simp] protected lemma add_tsub_tsub_cancel (hac : add_le_cancellable (a - c)) (h : c ≤ a) : (a + b) - (a - c) = b + c := -(hac.tsub_eq_iff_eq_add_of_le $ tsub_le_self.trans le_self_add).mpr $ - by rw [add_assoc, add_tsub_cancel_of_le h, add_comm] +hac.tsub_eq_of_eq_add $ by rw [add_assoc, add_tsub_cancel_of_le h, add_comm] protected lemma tsub_tsub_cancel_of_le (hba : add_le_cancellable (b - a)) (h : a ≤ b) : b - (b - a) = a := -by rw [hba.tsub_eq_iff_eq_add_of_le tsub_le_self, add_tsub_cancel_of_le h] +hba.tsub_eq_of_eq_add (add_tsub_cancel_of_le h).symm end add_le_cancellable @@ -620,15 +566,6 @@ contravariant.add_le_cancellable.tsub_lt_iff_left hbc lemma tsub_lt_iff_right (hbc : b ≤ a) : a - b < c ↔ a < c + b := contravariant.add_le_cancellable.tsub_lt_iff_right hbc -/-- This lemma (and some of its corollaries also holds for `ennreal`, - but this proof doesn't work for it. - Maybe we should add this lemma as field to `has_ordered_sub`? -/ -lemma lt_tsub_of_add_lt_right (h : a + c < b) : a < b - c := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_right h - -lemma lt_tsub_of_add_lt_left (h : a + c < b) : c < b - a := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_left h - lemma tsub_lt_iff_tsub_lt (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < c ↔ a - c < b := contravariant.add_le_cancellable.tsub_lt_iff_tsub_lt contravariant.add_le_cancellable h₁ h₂ @@ -648,13 +585,6 @@ lemma lt_of_tsub_lt_tsub_left_of_le [contravariant_class α α (+) (<)] (hca : c ≤ a) (h : a - b < a - c) : c < b := contravariant.add_le_cancellable.lt_of_tsub_lt_tsub_left_of_le hca h -lemma tsub_le_tsub_iff_left (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := -contravariant.add_le_cancellable.tsub_le_tsub_iff_left contravariant.add_le_cancellable h - -lemma tsub_right_inj (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := -contravariant.add_le_cancellable.tsub_right_inj contravariant.add_le_cancellable - contravariant.add_le_cancellable hba hca - lemma tsub_lt_tsub_right_of_le (h : c ≤ a) (h2 : a < b) : a - c < b - c := contravariant.add_le_cancellable.tsub_lt_tsub_right_of_le h h2 @@ -674,6 +604,72 @@ contravariant.add_le_cancellable.add_tsub_tsub_cancel h lemma tsub_tsub_cancel_of_le (h : a ≤ b) : b - (b - a) = a := contravariant.add_le_cancellable.tsub_tsub_cancel_of_le h +end contra +end has_exists_add_of_le + +/-! ### Lemmas in a canonically ordered monoid. -/ + +section canonically_ordered_add_monoid +variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma add_tsub_cancel_iff_le : a + (b - a) = b ↔ a ≤ b := +⟨λ h, le_iff_exists_add.mpr ⟨b - a, h.symm⟩, add_tsub_cancel_of_le⟩ + +lemma tsub_add_cancel_iff_le : b - a + a = b ↔ a ≤ b := +by { rw [add_comm], exact add_tsub_cancel_iff_le } + +@[simp] lemma tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := +by rw [← nonpos_iff_eq_zero, tsub_le_iff_left, add_zero] + +alias tsub_eq_zero_iff_le ↔ _ tsub_eq_zero_of_le + +attribute [simp] tsub_eq_zero_of_le + +@[simp] lemma tsub_self (a : α) : a - a = 0 := tsub_eq_zero_of_le le_rfl + +@[simp] lemma tsub_le_self : a - b ≤ a := tsub_le_iff_left.mpr $ le_add_left le_rfl + +@[simp] lemma zero_tsub (a : α) : 0 - a = 0 := tsub_eq_zero_of_le $ zero_le a + +lemma tsub_self_add (a b : α) : a - (a + b) = 0 := tsub_eq_zero_of_le $ self_le_add_right _ _ + +lemma tsub_pos_iff_not_le : 0 < a - b ↔ ¬ a ≤ b := +by rw [pos_iff_ne_zero, ne.def, tsub_eq_zero_iff_le] + +lemma tsub_pos_of_lt (h : a < b) : 0 < b - a := tsub_pos_iff_not_le.mpr h.not_le + +lemma tsub_lt_of_lt (h : a < b) : a - c < b := lt_of_le_of_lt tsub_le_self h + +namespace add_le_cancellable + +protected lemma tsub_le_tsub_iff_left (ha : add_le_cancellable a) (hc : add_le_cancellable c) + (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := +begin + refine ⟨_, λ h, tsub_le_tsub_left h a⟩, + rw [tsub_le_iff_left, ← hc.add_tsub_assoc_of_le h, hc.le_tsub_iff_right (h.trans le_add_self), + add_comm b], + apply ha, +end + +protected lemma tsub_right_inj (ha : add_le_cancellable a) (hb : add_le_cancellable b) + (hc : add_le_cancellable c) (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := +by simp_rw [le_antisymm_iff, ha.tsub_le_tsub_iff_left hb hba, ha.tsub_le_tsub_iff_left hc hca, + and_comm] + +end add_le_cancellable + +/-! #### Lemmas where addition is order-reflecting. -/ + +section contra +variable [contravariant_class α α (+) (≤)] + +lemma tsub_le_tsub_iff_left (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := +contravariant.add_le_cancellable.tsub_le_tsub_iff_left contravariant.add_le_cancellable h + +lemma tsub_right_inj (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := +contravariant.add_le_cancellable.tsub_right_inj contravariant.add_le_cancellable + contravariant.add_le_cancellable hba hca + end contra end canonically_ordered_add_monoid @@ -689,7 +685,7 @@ by rw [tsub_pos_iff_not_le, not_le] lemma tsub_eq_tsub_min (a b : α) : a - b = a - min a b := begin cases le_total a b with h h, - { rw [min_eq_left h, tsub_self, tsub_eq_zero_iff_le.mpr h] }, + { rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] }, { rw [min_eq_right h] }, end @@ -750,7 +746,7 @@ end contra lemma tsub_add_eq_max : a - b + b = max a b := begin cases le_total a b with h h, - { rw [max_eq_right h, tsub_eq_zero_iff_le.mpr h, zero_add] }, + { rw [max_eq_right h, tsub_eq_zero_of_le h, zero_add] }, { rw [max_eq_left h, tsub_add_cancel_of_le h] } end @@ -760,7 +756,7 @@ by rw [add_comm, max_comm, tsub_add_eq_max] lemma tsub_min : a - min a b = a - b := begin cases le_total a b with h h, - { rw [min_eq_left h, tsub_self, tsub_eq_zero_iff_le.mpr h] }, + { rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] }, { rw [min_eq_right h] } end diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 18f339125606f..c1f62824fa76b 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -1408,7 +1408,7 @@ lemma modify_nth_tail_modify_nth_tail_le (l.modify_nth_tail f n).modify_nth_tail g m = l.modify_nth_tail (λl, (f l).modify_nth_tail g (m - n)) n := begin - rcases le_iff_exists_add.1 h with ⟨m, rfl⟩, + rcases exists_add_of_le h with ⟨m, rfl⟩, rw [add_tsub_cancel_left, add_comm, modify_nth_tail_modify_nth_tail] end diff --git a/src/data/nat/fib.lean b/src/data/nat/fib.lean index 6f8ba43ef369d..acf6505001d9e 100644 --- a/src/data/nat/fib.lean +++ b/src/data/nat/fib.lean @@ -77,7 +77,7 @@ by rw [fib_add_two, add_tsub_cancel_right] lemma fib_lt_fib_succ {n : ℕ} (hn : 2 ≤ n) : fib n < fib (n + 1) := begin - rcases le_iff_exists_add.1 hn with ⟨n, rfl⟩, + rcases exists_add_of_le hn with ⟨n, rfl⟩, rw [← tsub_pos_iff_lt, add_comm 2, fib_add_two_sub_fib_add_one], apply fib_pos (succ_pos n), end diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 2c6e0abb5e715..c08c191aad4a2 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -357,7 +357,7 @@ ordered_comm_monoid.to_covariant_class_left ℝ≥0 lemma le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ε, 0 < ε → a ≤ b + ε) : a ≤ b := le_of_forall_le_of_dense $ assume x hxb, begin - rcases le_iff_exists_add.1 (le_of_lt hxb) with ⟨ε, rfl⟩, + rcases exists_add_of_le (le_of_lt hxb) with ⟨ε, rfl⟩, exact h _ ((lt_add_iff_pos_right b).1 hxb) end diff --git a/src/order/order_iso_nat.lean b/src/order/order_iso_nat.lean index 5a372a3386051..61b2b46975bd7 100644 --- a/src/order/order_iso_nat.lean +++ b/src/order/order_iso_nat.lean @@ -148,7 +148,7 @@ theorem exists_increasing_or_nonincreasing_subseq (r : α → α → Prop) [is_t begin obtain ⟨g, hr | hnr⟩ := exists_increasing_or_nonincreasing_subseq' r f, { refine ⟨g, or.intro_left _ (λ m n mn, _)⟩, - obtain ⟨x, rfl⟩ := le_iff_exists_add.1 (nat.succ_le_iff.2 mn), + obtain ⟨x, rfl⟩ := exists_add_of_le (nat.succ_le_iff.2 mn), induction x with x ih, { apply hr }, { apply is_trans.trans _ _ _ _ (hr _), diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index 5330367f08fff..eb1ef542aca4d 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -663,7 +663,7 @@ begin by_cases hn : n < g 0, { apply hf1.2 m n mn, rwa [if_pos hn, if_pos (mn.trans hn)] at hmn }, - { obtain ⟨n', rfl⟩ := le_iff_exists_add.1 (not_lt.1 hn), + { obtain ⟨n', rfl⟩ := exists_add_of_le (not_lt.1 hn), rw [if_neg hn, add_comm (g 0) n', add_tsub_cancel_right] at hmn, split_ifs at hmn with hm hm, { apply hf1.2 m (g n') (lt_of_lt_of_le hm (g.monotone n'.zero_le)), diff --git a/src/ring_theory/polynomial/eisenstein.lean b/src/ring_theory/polynomial/eisenstein.lean index 39eae9a2d9c55..cf28fd70a6b87 100644 --- a/src/ring_theory/polynomial/eisenstein.lean +++ b/src/ring_theory/polynomial/eisenstein.lean @@ -115,7 +115,7 @@ lemma exists_mem_adjoin_mul_eq_pow_nat_degree_le {x : S} (hx : aeval x f = 0) ∃ y ∈ adjoin R ({x} : set S), (algebra_map R S) p * y = x ^ i := begin intros i hi, - obtain ⟨k, hk⟩ := le_iff_exists_add.1 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, _, _⟩, @@ -131,7 +131,7 @@ lemma pow_nat_degree_le_of_root_of_monic_mem {x : R} (hroot : is_root f x) (hmo ∀ i, f.nat_degree ≤ i → x ^ i ∈ 𝓟 := begin intros i hi, - obtain ⟨k, hk⟩ := le_iff_exists_add.1 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 }, @@ -147,7 +147,7 @@ lemma pow_nat_degree_le_of_aeval_zero_of_monic_mem_map {x : S} (hx : aeval x f = begin suffices : x ^ (f.map (algebra_map R S)).nat_degree ∈ 𝓟.map (algebra_map R S), { intros i hi, - obtain ⟨k, hk⟩ := le_iff_exists_add.1 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, From 815b9f79b6126302bc636a0fe2627ae6092fb502 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Jul 2022 01:36:16 +0000 Subject: [PATCH 035/173] chore(algebra/module/basic): use `simp` instead of `norm_num` (#15670) --- src/algebra/module/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index b7e07cf660abf..73c64a199c293 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -5,9 +5,9 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import algebra.big_operators.basic import algebra.smul_with_zero +import data.rat.cast import group_theory.group_action.big_operators import group_theory.group_action.group -import tactic.norm_num /-! # Modules over a ring @@ -524,7 +524,7 @@ lemma nat.no_zero_smul_divisors : no_zero_smul_divisors ℕ M := ⟨by { intros c x, rw [nsmul_eq_smul_cast R, smul_eq_zero], simp }⟩ @[simp] lemma two_nsmul_eq_zero {v : M} : 2 • v = 0 ↔ v = 0 := -by { haveI := nat.no_zero_smul_divisors R M, norm_num [smul_eq_zero] } +by { haveI := nat.no_zero_smul_divisors R M, simp [smul_eq_zero] } end nat From 38d5ea5521da61834b0ca37f096c78fdcc20dd30 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Mon, 25 Jul 2022 04:21:38 +0000 Subject: [PATCH 036/173] chore(scripts): update nolints.txt (#15674) 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 c7a1ccb0979f5..e781488f41d9a 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -287,10 +287,6 @@ apply_nolint ctop.realizer.of_equiv doc_blame apply_nolint ctop.to_realizer doc_blame apply_nolint locally_finite.realizer has_inhabited_instance doc_blame --- data/finite/basic.lean -apply_nolint finite.of_fintype fintype_finite -apply_nolint fintype.finite fintype_finite - -- data/finset/noncomm_prod.lean apply_nolint add_monoid_hom.pi_ext fintype_finite apply_nolint finset.noncomm_prod_union_of_disjoint to_additive_doc From da1b94b6c86b8dc94f80b8d96bdb38a1f2f0f6a7 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 25 Jul 2022 05:10:07 +0000 Subject: [PATCH 037/173] feat(algebraic_geometry/morphisms): Basic framework for classes of morphisms between schemes (#14944) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/morphisms/basic.lean | 48 +++++++ .../limits/shapes/pullbacks.lean | 10 ++ src/category_theory/morphism_property.lean | 126 ++++++++++++++++++ 3 files changed, 184 insertions(+) create mode 100644 src/algebraic_geometry/morphisms/basic.lean create mode 100644 src/category_theory/morphism_property.lean diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean new file mode 100644 index 0000000000000..b1774464f194e --- /dev/null +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -0,0 +1,48 @@ +/- +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.AffineScheme +import algebraic_geometry.pullbacks +import category_theory.morphism_property + +/-! +# Properties of morphisms between Schemes + +We provide the basic framework for talking about properties of morphisms between Schemes. + +-/ + +universe u + +open topological_space category_theory category_theory.limits opposite + +noncomputable theory + +namespace algebraic_geometry + +/-- An `affine_target_morphism_property` is a class of morphisms from an arbitrary scheme into an +affine scheme. -/ +def affine_target_morphism_property := ∀ ⦃X Y : Scheme⦄ (f : X ⟶ Y) [is_affine Y], Prop + +/-- `is_iso` as a `morphism_property`. -/ +protected def Scheme.is_iso : morphism_property Scheme := @is_iso Scheme _ + +/-- `is_iso` as an `affine_morphism_property`. -/ +protected def Scheme.affine_target_is_iso : affine_target_morphism_property := +λ X Y f H, is_iso f + +instance : inhabited affine_target_morphism_property := ⟨Scheme.affine_target_is_iso⟩ + +/-- A `affine_target_morphism_property` can be extended to a `morphism_property` such that it +*never* holds when the target is not affine -/ +def affine_target_morphism_property.to_property (P : affine_target_morphism_property) : + morphism_property Scheme := +λ X Y f, ∃ h, @@P f h + +lemma affine_target_morphism_property.to_property_apply (P : affine_target_morphism_property) + {X Y : Scheme} (f : X ⟶ Y) [is_affine Y] : + P.to_property f ↔ P f := by { delta affine_target_morphism_property.to_property, simp [*] } + +end algebraic_geometry diff --git a/src/category_theory/limits/shapes/pullbacks.lean b/src/category_theory/limits/shapes/pullbacks.lean index d6f2e1b78619a..dfd1089f8fbf1 100644 --- a/src/category_theory/limits/shapes/pullbacks.lean +++ b/src/category_theory/limits/shapes/pullbacks.lean @@ -2209,4 +2209,14 @@ begin apply_instance end +variable {C} + +/-- Given a morphism `f : X ⟶ Y`, we can take morphisms over `Y` to morphisms over `X` via +pullbacks. This is right adjoint to `over.map` (TODO) -/ +@[simps obj_left obj_hom map_left {rhs_md := semireducible, simp_rhs := tt}] +def base_change [has_pullbacks C] {X Y : C} (f : X ⟶ Y) : over Y ⥤ over X := +{ obj := λ g, over.mk (pullback.snd : pullback g.hom f ⟶ _), + map := λ g₁ g₂ i, over.hom_mk (pullback.map _ _ _ _ i.left (𝟙 _) (𝟙 _) (by simp) (by simp)) + (by simp) } + end category_theory.limits diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean new file mode 100644 index 0000000000000..8db835f72d3d3 --- /dev/null +++ b/src/category_theory/morphism_property.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import category_theory.limits.shapes.pullbacks + +/-! +# Properties of morphisms + +We provide the basic framework for talking about properties of morphisms. +The following meta-properties are defined + +* `respects_iso`: `P` respects isomorphisms if `P f → P (e ≫ f)` and `P f → P (f ≫ e)`, where + `e` is an isomorphism. +* `stable_under_composition`: `P` is stable under composition if `P f → P g → P (f ≫ g)`. +* `stable_under_base_change`: `P` is stable under base change if `P (Y ⟶ S) → P (X ×[S] Y ⟶ X)`. + +-/ + +universes v u + +open category_theory category_theory.limits opposite + +noncomputable theory + +namespace category_theory + +variables (C : Type u) [category.{v} C] + +/-- A `morphism_property C` is a class of morphisms between objects in `C`. -/ +@[derive complete_lattice] +def morphism_property := ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Prop + +instance : inhabited (morphism_property C) := ⟨⊤⟩ + +variable {C} + +namespace morphism_property + +/-- A morphism property `respects_iso` if it still holds when composed with an isomorphism -/ +def respects_iso (P : morphism_property C) : Prop := + (∀ {X Y Z} (e : X ≅ Y) (f : Y ⟶ Z), P f → P (e.hom ≫ f)) ∧ + (∀ {X Y Z} (e : Y ≅ Z) (f : X ⟶ Y), P f → P (f ≫ e.hom)) + +/-- A morphism property is `stable_under_composition` if the composition of two such morphisms +still falls in the class. -/ +def stable_under_composition (P : morphism_property C) : Prop := + ∀ ⦃X Y Z⦄ (f : X ⟶ Y) (g : Y ⟶ Z), P f → P g → P (f ≫ g) + +/-- A morphism property is `stable_under_composition` if the base change of such a morphism +still falls in the class. -/ +def stable_under_base_change [has_pullbacks C] (P : morphism_property C) : Prop := +∀ ⦃X Y S : C⦄ (f : X ⟶ S) (g : Y ⟶ S), P g → P (pullback.fst : pullback f g ⟶ X) + +lemma stable_under_composition.respects_iso {P : morphism_property C} + (hP : stable_under_composition P) (hP' : ∀ {X Y} (e : X ≅ Y), P e.hom) : respects_iso P := +⟨λ X Y Z e f hf, hP _ _ (hP' e) hf, λ X Y Z e f hf, hP _ _ hf (hP' e)⟩ + +lemma respects_iso.cancel_left_is_iso {P : morphism_property C} + (hP : respects_iso P) {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [is_iso f] : + P (f ≫ g) ↔ P g := +⟨λ h, by simpa using hP.1 (as_iso f).symm (f ≫ g) h, hP.1 (as_iso f) g⟩ + +lemma respects_iso.cancel_right_is_iso {P : morphism_property C} + (hP : respects_iso P) {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [is_iso g] : + P (f ≫ g) ↔ P f := +⟨λ h, by simpa using hP.2 (as_iso g).symm (f ≫ g) h, hP.2 (as_iso g) f⟩ + +-- This is here to mirror `stable_under_base_change.snd`. +@[nolint unused_arguments] +lemma stable_under_base_change.fst [has_pullbacks C] {P : morphism_property C} + (hP : stable_under_base_change P) (hP' : respects_iso P) {X Y S : C} (f : X ⟶ S) + (g : Y ⟶ S) (H : P g) : P (pullback.fst : pullback f g ⟶ X) := +hP f g H + +lemma stable_under_base_change.snd [has_pullbacks C] {P : morphism_property C} + (hP : stable_under_base_change P) (hP' : respects_iso P) {X Y S : C} (f : X ⟶ S) + (g : Y ⟶ S) (H : P f) : P (pullback.snd : pullback f g ⟶ Y) := +begin + rw [← pullback_symmetry_hom_comp_fst, hP'.cancel_left_is_iso], + exact hP g f H +end + +lemma stable_under_base_change.base_change_obj [has_pullbacks C] {P : morphism_property C} + (hP : stable_under_base_change P) (hP' : respects_iso P) {S S' : C} (f : S' ⟶ S) + (X : over S) (H : P X.hom) : P ((base_change f).obj X).hom := +hP.snd hP' X.hom f H + +lemma stable_under_base_change.base_change_map [has_pullbacks C] {P : morphism_property C} + (hP : stable_under_base_change P) (hP' : respects_iso P) {S S' : C} (f : S' ⟶ S) + {X Y : over S} (g : X ⟶ Y) (H : P g.left) : P ((base_change f).map g).left := +begin + let e := pullback_right_pullback_fst_iso Y.hom f g.left ≪≫ + pullback.congr_hom (g.w.trans (category.comp_id _)) rfl, + have : e.inv ≫ pullback.snd = ((base_change f).map g).left, + { apply pullback.hom_ext; dsimp; simp }, + rw [← this, hP'.cancel_left_is_iso], + apply hP.snd hP', + exact H +end + +lemma stable_under_base_change.pullback_map [has_pullbacks C] {P : morphism_property C} + (hP : stable_under_base_change P) (hP' : respects_iso P) + (hP'' : stable_under_composition P) {S X X' Y Y' : C} + {f : X ⟶ S} {g : Y ⟶ S} {f' : X' ⟶ S} {g' : Y' ⟶ S} {i₁ : X ⟶ X'} {i₂ : Y ⟶ Y'} + (h₁ : P i₁) (h₂ : P i₂) (e₁ : f = i₁ ≫ f') (e₂ : g = i₂ ≫ g') : + P (pullback.map f g f' g' i₁ i₂ (𝟙 _) + ((category.comp_id _).trans e₁) ((category.comp_id _).trans e₂)) := +begin + have : pullback.map f g f' g' i₁ i₂ (𝟙 _) + ((category.comp_id _).trans e₁) ((category.comp_id _).trans e₂) = + ((pullback_symmetry _ _).hom ≫ + ((base_change _).map (over.hom_mk _ e₂.symm : over.mk g ⟶ over.mk g')).left) ≫ + (pullback_symmetry _ _).hom ≫ + ((base_change g').map (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f')).left, + { apply pullback.hom_ext; dsimp; simp }, + rw this, + apply hP''; rw hP'.cancel_left_is_iso, + exacts [hP.base_change_map hP' _ (over.hom_mk _ e₂.symm : over.mk g ⟶ over.mk g') h₂, + hP.base_change_map hP' _ (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f') h₁], +end + +end morphism_property + +end category_theory From 8d973301b4126255d99238950dc354c5f7c83592 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 25 Jul 2022 05:10:08 +0000 Subject: [PATCH 038/173] fix(algebraic_geometry/projective_spectrum/scheme) : fix module doc string (#15633) After renaming definitions/lemmas in the body of the `src/algebraic_geometry/projective_spectrum/scheme.lean`, the module doc string is left unchanged. This pr fix the doc string --- .../projective_spectrum/scheme.lean | 17 ++++++++--------- 1 file changed, 8 insertions(+), 9 deletions(-) diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 70c147b01c969..d99ce60d766e8 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -32,12 +32,12 @@ open sets in `Proj`, more specifically: 1. We prove that `Proj` can be covered by basic open sets at homogeneous element of positive degree. 2. We prove that for any `f : A`, `Proj.T | (pbo f)` is homeomorphic to `Spec.T A⁰_f`: - - forward direction : + - forward direction `to_Spec`: for any `x : pbo f`, i.e. a relevant homogeneous prime ideal `x`, send it to - `x ∩ span {g / 1 | g ∈ A}` (see `Top_component.forward.carrier`). This ideal is prime, the proof - is in `Top_component.forward.to_fun`. The fact that this function is continuous is found in - `Top_component.forward` - - backward direction : TBC + `x ∩ span {g / 1 | g ∈ A}` (see `Proj_iso_Spec_Top_component.to_Spec.carrier`). This ideal is + prime, the proof is in `Proj_iso_Spec_Top_component.to_Spec.to_fun`. The fact that this function + is continuous is found in `Proj_iso_Spec_Top_component.to_Spec` + - backward direction `from_Spec`: TBC ## Main Definitions and Statements @@ -45,11 +45,10 @@ open sets in `Proj`, more specifically: element of degree `n` is the subring of elements of the form `a/f^m` where `a` has degree `mn`. For a homogeneous element `f` of degree `n` -* `Top_component.forward`: `forward f` is the +* `Proj_iso_Spec_Top_component.to_Spec`: `forward f` is the continuous map between `Proj.T| pbo f` and `Spec.T A⁰_f` -* `Top_component.forward.preimage_eq`: for any `a: A`, if `a/f^m` has degree zero, then the preimage - of `sbo a/f^m` under `forward f` is `pbo f ∩ pbo a`. - +* `Proj_iso_Spec_Top_component.to_Spec.preimage_eq`: for any `a: A`, if `a/f^m` has degree zero, + then the preimage of `sbo a/f^m` under `to_Spec f` is `pbo f ∩ pbo a`. * [Robin Hartshorne, *Algebraic Geometry*][Har77]: Chapter II.2 Proposition 2.5 -/ From e3d9b532e1761a9dc06758451adbed25212cf28d Mon Sep 17 00:00:00 2001 From: Julian-Kuelshammer Date: Mon, 25 Jul 2022 06:23:38 +0000 Subject: [PATCH 039/173] feat(category_theory/preadditive/*): algebra over endofunctor preadditive and forget additive functor (#15100) This PR shows that the category of algebras over an endofunctor is preadditive and that forgetful functors from algebras over endofunctors and (co)algebras over (co)monads are additive. --- .../preadditive/additive_functor.lean | 3 + .../preadditive/eilenberg_moore.lean | 28 +++++ .../preadditive/endo_functor.lean | 102 ++++++++++++++++++ 3 files changed, 133 insertions(+) create mode 100644 src/category_theory/preadditive/endo_functor.lean diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index e517370f1cd34..8bfead2488c20 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -71,6 +71,9 @@ lemma map_neg {X Y : C} {f : X ⟶ Y} : F.map (-f) = - F.map f := lemma map_sub {X Y : C} {f g : X ⟶ Y} : F.map (f - g) = F.map f - F.map g := (F.map_add_hom : (X ⟶ Y) →+ (F.obj X ⟶ F.obj Y)).map_sub _ _ +lemma map_nsmul {X Y : C} {f : X ⟶ Y} {n : ℕ} : F.map (n • f) = n • F.map f := +(F.map_add_hom : (X ⟶ Y) →+ (F.obj X ⟶ F.obj Y)).map_nsmul _ _ + -- You can alternatively just use `functor.map_smul` here, with an explicit `(r : ℤ)` argument. lemma map_zsmul {X Y : C} {f : X ⟶ Y} {r : ℤ} : F.map (r • f) = r • F.map f := (F.map_add_hom : (X ⟶ Y) →+ (F.obj X ⟶ F.obj Y)).map_zsmul _ _ diff --git a/src/category_theory/preadditive/eilenberg_moore.lean b/src/category_theory/preadditive/eilenberg_moore.lean index 6f7c657336dc4..4ac0e531313ea 100644 --- a/src/category_theory/preadditive/eilenberg_moore.lean +++ b/src/category_theory/preadditive/eilenberg_moore.lean @@ -34,21 +34,35 @@ instance monad.algebra_preadditive : preadditive (monad.algebra T) := zero := { f := 0, h' := by simp only [functor.map_zero, zero_comp, comp_zero] }, + nsmul := λ n α, + { f := n • α.f, + h' := by rw [functor.map_nsmul, nsmul_comp, monad.algebra.hom.h, comp_nsmul] }, neg := λ α, { f := -α.f, h' := by simp only [functor.map_neg, neg_comp, monad.algebra.hom.h, comp_neg] }, sub := λ α β, { f := α.f - β.f, h' := by simp only [functor.map_sub, sub_comp, monad.algebra.hom.h, comp_sub] }, + zsmul := λ r α, + { f := r • α.f, + h' := by rw [functor.map_zsmul, zsmul_comp, monad.algebra.hom.h, comp_zsmul] }, add_assoc := by { intros, ext, apply add_assoc }, zero_add := by { intros, ext, apply zero_add }, add_zero := by { intros, ext, apply add_zero }, + nsmul_zero' := by { intros, ext, apply zero_smul }, + nsmul_succ' := by { intros, ext, apply succ_nsmul }, sub_eq_add_neg := by { intros, ext, apply sub_eq_add_neg }, + zsmul_zero' := by { intros, ext, apply zero_smul }, + zsmul_succ' := by { intros, ext, dsimp, simp only [coe_nat_zsmul, succ_nsmul], refl, }, + zsmul_neg' := by { intros, ext, simp only [zsmul_neg_succ_of_nat, neg_inj, + nsmul_eq_smul_cast ℤ] }, add_left_neg := by { intros, ext, apply add_left_neg }, add_comm := by { intros, ext, apply add_comm } }, add_comp' := by { intros, ext, apply add_comp }, comp_add' := by { intros, ext, apply comp_add } } +instance monad.forget_additive : (monad.forget T).additive := {} + variables (U : comonad C) [functor.additive (U : C ⥤ C)] /-- The category of coalgebras over an additive comonad on a preadditive category is preadditive. -/ @@ -61,19 +75,33 @@ instance comonad.coalgebra_preadditive : preadditive (comonad.coalgebra U) := zero := { f := 0, h' := by simp only [functor.map_zero, comp_zero, zero_comp] }, + nsmul := λ n α, + { f := n • α.f, + h' := by rw [functor.map_nsmul, comp_nsmul, comonad.coalgebra.hom.h, nsmul_comp] }, neg := λ α, { f := -α.f, h' := by simp only [functor.map_neg, comp_neg, comonad.coalgebra.hom.h, neg_comp] }, sub := λ α β, { f := α.f - β.f, h' := by simp only [functor.map_sub, comp_sub, comonad.coalgebra.hom.h, sub_comp] }, + zsmul := λ r α, + { f := r • α.f, + h' := by rw [functor.map_zsmul, comp_zsmul, comonad.coalgebra.hom.h, zsmul_comp] }, add_assoc := by { intros, ext, apply add_assoc }, zero_add := by { intros, ext, apply zero_add }, add_zero := by { intros, ext, apply add_zero }, + nsmul_zero' := by { intros, ext, apply zero_smul }, + nsmul_succ' := by { intros, ext, apply succ_nsmul }, sub_eq_add_neg := by { intros, ext, apply sub_eq_add_neg }, + zsmul_zero' := by { intros, ext, apply zero_smul }, + zsmul_succ' := by { intros, ext, dsimp, simp only [coe_nat_zsmul, succ_nsmul], refl, }, + zsmul_neg' := by { intros, ext, simp only [zsmul_neg_succ_of_nat, neg_inj, + nsmul_eq_smul_cast ℤ] }, add_left_neg := by { intros, ext, apply add_left_neg }, add_comm := by { intros, ext, apply add_comm } }, add_comp' := by { intros, ext, apply add_comp }, comp_add' := by { intros, ext, apply comp_add } } +instance comonad.forget_additive : (comonad.forget U).additive := {} + end category_theory diff --git a/src/category_theory/preadditive/endo_functor.lean b/src/category_theory/preadditive/endo_functor.lean new file mode 100644 index 0000000000000..8b4951081b04b --- /dev/null +++ b/src/category_theory/preadditive/endo_functor.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2022 Julian Kuelshammer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Julian Kuelshammer +-/ + +import category_theory.preadditive.default +import category_theory.endofunctor.algebra +import category_theory.preadditive.additive_functor + +/-! +# Preadditive structure on algebras over a monad + +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. +-/ + +universes v₁ u₁ -- morphism levels before object levels. See note [category_theory universes]. + +namespace category_theory +variables (C : Type u₁) [category.{v₁} C] [preadditive C] (F : C ⥤ C) + [functor.additive (F : C ⥤ C)] + +open category_theory.limits preadditive + +/-- The category of algebras over an additive endofunctor on a preadditive category is preadditive. +-/ +@[simps] +instance endofunctor.algebra_preadditive : preadditive (endofunctor.algebra F) := +{ hom_group := λ A₁ A₂, { add := λ α β, + { f := α.f + β.f, + h' := by simp only [functor.map_add, add_comp, endofunctor.algebra.hom.h, comp_add] }, + zero := + { f := 0, + h' := by simp only [functor.map_zero, zero_comp, comp_zero] }, + nsmul := λ n α, + { f := n • α.f, + h' := by rw [comp_nsmul, functor.map_nsmul, nsmul_comp, endofunctor.algebra.hom.h] }, + neg := λ α, + { f := -α.f, + h' := by simp only [functor.map_neg, neg_comp, endofunctor.algebra.hom.h, comp_neg] }, + sub := λ α β, + { f := α.f - β.f, + h' := by simp only [functor.map_sub, sub_comp, endofunctor.algebra.hom.h, comp_sub] }, + zsmul := λ r α, + { f := r • α.f, + h' := by rw [comp_zsmul, functor.map_zsmul, zsmul_comp, endofunctor.algebra.hom.h] }, + add_assoc := by { intros, ext, apply add_assoc }, + zero_add := by { intros, ext, apply zero_add }, + add_zero := by { intros, ext, apply add_zero }, + nsmul_zero' := by { intros, ext, apply zero_smul }, + nsmul_succ' := by { intros, ext, apply succ_nsmul }, + sub_eq_add_neg := by { intros, ext, apply sub_eq_add_neg }, + zsmul_zero' := by { intros, ext, apply zero_smul }, + zsmul_succ' := by { intros, ext, dsimp, simp only [coe_nat_zsmul, succ_nsmul], refl, }, + zsmul_neg' := by { intros, ext, simp only [zsmul_neg_succ_of_nat, neg_inj, + nsmul_eq_smul_cast ℤ] }, + add_left_neg := by { intros, ext, apply add_left_neg }, + add_comm := by { intros, ext, apply add_comm } }, + add_comp' := by { intros, ext, apply add_comp }, + comp_add' := by { intros, ext, apply comp_add } } + +instance algebra.forget_additive : (endofunctor.algebra.forget F).additive := {} + +@[simps] +instance endofunctor.coalgebra_preadditive : preadditive (endofunctor.coalgebra F) := +{ hom_group := λ A₁ A₂, { add := λ α β, + { f := α.f + β.f, + h' := by simp only [functor.map_add, comp_add, endofunctor.coalgebra.hom.h, add_comp] }, + zero := + { f := 0, + h' := by simp only [functor.map_zero, zero_comp, comp_zero] }, + nsmul := λ n α, + { f := n • α.f, + h' := by rw [functor.map_nsmul, comp_nsmul, endofunctor.coalgebra.hom.h, nsmul_comp] }, + neg := λ α, + { f := -α.f, + h' := by simp only [functor.map_neg, comp_neg, endofunctor.coalgebra.hom.h, neg_comp] }, + sub := λ α β, + { f := α.f - β.f, + h' := by simp only [functor.map_sub, comp_sub, endofunctor.coalgebra.hom.h, sub_comp] }, + zsmul := λ r α, + { f := r • α.f, + h' := by rw [functor.map_zsmul, comp_zsmul, endofunctor.coalgebra.hom.h, zsmul_comp] }, + add_assoc := by { intros, ext, apply add_assoc }, + zero_add := by { intros, ext, apply zero_add }, + add_zero := by { intros, ext, apply add_zero }, + nsmul_zero' := by { intros, ext, apply zero_smul }, + nsmul_succ' := by { intros, ext, apply succ_nsmul }, + sub_eq_add_neg := by { intros, ext, apply sub_eq_add_neg }, + zsmul_zero' := by { intros, ext, apply zero_smul }, + zsmul_succ' := by { intros, ext, dsimp, simp only [coe_nat_zsmul, succ_nsmul], refl, }, + zsmul_neg' := by { intros, ext, simp only [zsmul_neg_succ_of_nat, neg_inj, + nsmul_eq_smul_cast ℤ] }, + add_left_neg := by { intros, ext, apply add_left_neg }, + add_comm := by { intros, ext, apply add_comm } }, + add_comp' := by { intros, ext, apply add_comp }, + comp_add' := by { intros, ext, apply comp_add } } + +instance coalgebra.forget_additive : (endofunctor.coalgebra.forget F).additive := {} + +end category_theory From 53cd6f2337db500a217e4f832c457b3b2ef862fd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Jul 2022 09:05:27 +0000 Subject: [PATCH 040/173] feat(*/partition_of_unity): more lemmas based on the partition of unity (#15609) Add `metric.exists_continuous_real_forall_closed_ball_subset` and `metric.exists_smooth_forall_closed_ball_subset`. For the sphere eversion project, Lemma 3.6. --- src/geometry/manifold/partition_of_unity.lean | 34 +++- src/topology/metric_space/emetric_space.lean | 14 ++ .../metric_space/partition_of_unity.lean | 145 ++++++++++++++++++ 3 files changed, 191 insertions(+), 2 deletions(-) create mode 100644 src/topology/metric_space/partition_of_unity.lean diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index fc5fba78dd5e7..778490f7c8900 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -5,8 +5,7 @@ Authors: Yury G. Kudryashov -/ import geometry.manifold.algebra.structures import geometry.manifold.bump_function -import topology.paracompact -import topology.partition_of_unity +import topology.metric_space.partition_of_unity import topology.shrinking_lemma /-! @@ -502,3 +501,34 @@ lemma exists_smooth_forall_mem_convex_of_local_const (ht : ∀ x, convex ℝ (t ∃ g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯, ∀ x, g x ∈ t x := exists_smooth_forall_mem_convex_of_local I ht $ λ x, let ⟨c, hc⟩ := Hloc x in ⟨_, hc, λ _, c, smooth_on_const, λ y, id⟩ + +/-- Let `M` be a smooth σ-compact manifold with extended distance. Let `K : ι → set M` be a locally +finite family of closed sets, let `U : ι → set M` be a family of open sets such that `K i ⊆ U i` for +all `i`. Then there exists a positive smooth function `δ : M → ℝ≥0` such that for any `i` and +`x ∈ K i`, we have `emetric.closed_ball x (δ x) ⊆ U i`. -/ +lemma emetric.exists_smooth_forall_closed_ball_subset {M} [emetric_space M] [charted_space H M] + [smooth_manifold_with_corners I M] [sigma_compact_space M] {K : ι → set M} + {U : ι → set M} (hK : ∀ i, is_closed (K i)) (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) + (hfin : locally_finite K) : + ∃ δ : C^∞⟮I, M; 𝓘(ℝ, ℝ), ℝ⟯, (∀ x, 0 < δ x) ∧ + ∀ i (x ∈ K i), emetric.closed_ball x (ennreal.of_real (δ x)) ⊆ U i := +by simpa only [mem_inter_eq, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι M] + using exists_smooth_forall_mem_convex_of_local_const I + emetric.exists_forall_closed_ball_subset_aux₂ + (emetric.exists_forall_closed_ball_subset_aux₁ hK hU hKU hfin) + +/-- Let `M` be a smooth σ-compact manifold with a metric. Let `K : ι → set M` be a locally finite +family of closed sets, let `U : ι → set M` be a family of open sets such that `K i ⊆ U i` for all +`i`. Then there exists a positive smooth function `δ : M → ℝ≥0` such that for any `i` and `x ∈ K i`, +we have `metric.closed_ball x (δ x) ⊆ U i`. -/ +lemma metric.exists_smooth_forall_closed_ball_subset {M} [metric_space M] [charted_space H M] + [smooth_manifold_with_corners I M] [sigma_compact_space M] {K : ι → set M} + {U : ι → set M} (hK : ∀ i, is_closed (K i)) (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) + (hfin : locally_finite K) : + ∃ δ : C^∞⟮I, M; 𝓘(ℝ, ℝ), ℝ⟯, (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), metric.closed_ball x (δ x) ⊆ U i := +begin + rcases emetric.exists_smooth_forall_closed_ball_subset I hK hU hKU hfin with ⟨δ, hδ0, hδ⟩, + refine ⟨δ, hδ0, λ i x hx, _⟩, + rw [← metric.emetric_closed_ball (hδ0 _).le], + exact hδ i x hx +end diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index e4e9477af35ec..ced3fdf605dfc 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -232,6 +232,12 @@ emetric.mk_uniformity_basis (λ _, ennreal.coe_pos.2) (λ ε ε₀, let ⟨δ, hδ⟩ := ennreal.lt_iff_exists_nnreal_btwn.1 ε₀ in ⟨δ, ennreal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩) +theorem uniformity_basis_edist_nnreal_le : + (𝓤 α).has_basis (λ ε : ℝ≥0, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 ≤ ε}) := +emetric.mk_uniformity_basis_le (λ _, ennreal.coe_pos.2) + (λ ε ε₀, let ⟨δ, hδ⟩ := ennreal.lt_iff_exists_nnreal_btwn.1 ε₀ in + ⟨δ, ennreal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩) + theorem uniformity_basis_edist_inv_nat : (𝓤 α).has_basis (λ _, true) (λ n:ℕ, {p:α×α | edist p.1 p.2 < (↑n)⁻¹}) := emetric.mk_uniformity_basis @@ -556,6 +562,14 @@ eq_empty_iff_forall_not_mem.trans ⟨λh, le_bot_iff.1 (le_of_not_gt (λ ε0, h _ (mem_ball_self ε0))), λε0 y h, not_lt_of_le (le_of_eq ε0) (pos_of_mem_ball h)⟩ +lemma ord_connected_set_of_closed_ball_subset (x : α) (s : set α) : + ord_connected {r | closed_ball x r ⊆ s} := +⟨λ r₁ hr₁ r₂ hr₂ r hr, (closed_ball_subset_closed_ball hr.2).trans hr₂⟩ + +lemma ord_connected_set_of_ball_subset (x : α) (s : set α) : + ord_connected {r | ball x r ⊆ s} := +⟨λ r₁ hr₁ r₂ hr₂ r hr, (ball_subset_ball hr.2).trans hr₂⟩ + /-- Relation “two points are at a finite edistance” is an equivalence relation. -/ def edist_lt_top_setoid : setoid α := { r := λ x y, edist x y < ⊤, diff --git a/src/topology/metric_space/partition_of_unity.lean b/src/topology/metric_space/partition_of_unity.lean new file mode 100644 index 0000000000000..5fa49bc4e0c8a --- /dev/null +++ b/src/topology/metric_space/partition_of_unity.lean @@ -0,0 +1,145 @@ +/- +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.emetric_paracompact +import analysis.convex.partition_of_unity + +/-! +# Lemmas about (e)metric spaces that need partition of unity + +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 +positive continuous function `δ : C(X, → ℝ)` such that for any `i` and `x ∈ K i`, we have +`metric.closed_ball x (δ x) ⊆ U i`. We also formulate versions of this lemma for extended metric +spaces and for different codomains (`ℝ`, `ℝ≥0`, and `ℝ≥0∞`). + +We also prove a few auxiliary lemmas to be used later in a proof of the smooth version of this +lemma. + +## Tags + +metric space, partition of unity, locally finite +-/ + +open_locale topological_space ennreal big_operators nnreal filter +open set function filter topological_space + +variables {ι X : Type*} + +namespace emetric + +variables [emetric_space X] {K : ι → set X} {U : ι → set X} + +/-- Let `K : ι → set X` be a locally finitie family of closed sets in an emetric space. Let +`U : ι → set X` be a family of open sets such that `K i ⊆ U i` for all `i`. Then for any point +`x : X`, for sufficiently small `r : ℝ≥0∞` and for `y` sufficiently close to `x`, for all `i`, if +`y ∈ K i`, then `emetric.closed_ball y r ⊆ U i`. -/ +lemma eventually_nhds_zero_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) (x : X) : + ∀ᶠ p : ℝ≥0∞ × X in 𝓝 0 ×ᶠ 𝓝 x, ∀ i, p.2 ∈ K i → closed_ball p.2 p.1 ⊆ U i := +begin + suffices : ∀ i, x ∈ K i → ∀ᶠ p : ℝ≥0∞ × X in 𝓝 0 ×ᶠ 𝓝 x, closed_ball p.2 p.1 ⊆ U i, + { filter_upwards [tendsto_snd (hfin.Inter_compl_mem_nhds hK x), + (eventually_all_finite (hfin.point_finite x)).2 this], + rintro ⟨r, y⟩ hxy hyU i hi, + simp only [mem_Inter₂, mem_compl_iff, not_imp_not, mem_preimage] at hxy, + exact hyU _ (hxy _ hi) }, + intros i hi, + rcases nhds_basis_closed_eball.mem_iff.1 ((hU i).mem_nhds $ hKU i hi) with ⟨R, hR₀, hR⟩, + rcases ennreal.lt_iff_exists_nnreal_btwn.mp hR₀ with ⟨r, hr₀, hrR⟩, + filter_upwards [prod_mem_prod (eventually_lt_nhds hr₀) + (closed_ball_mem_nhds x (tsub_pos_iff_lt.2 hrR))] with p hp z hz, + apply hR, + calc edist z x ≤ edist z p.2 + edist p.2 x : edist_triangle _ _ _ + ... ≤ p.1 + (R - p.1) : add_le_add hz $ le_trans hp.2 $ tsub_le_tsub_left hp.1.out.le _ + ... = R : add_tsub_cancel_of_le (lt_trans hp.1 hrR).le, +end + +lemma exists_forall_closed_ball_subset_aux₁ (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) (x : X) : + ∃ r : ℝ, ∀ᶠ y in 𝓝 x, r ∈ Ioi (0 : ℝ) ∩ + ennreal.of_real ⁻¹' ⋂ i (hi : y ∈ K i), {r | closed_ball y r ⊆ U i} := +begin + have := (ennreal.continuous_of_real.tendsto' 0 0 ennreal.of_real_zero).eventually + (eventually_nhds_zero_forall_closed_ball_subset hK hU hKU hfin x).curry, + rcases this.exists_gt with ⟨r, hr0, hr⟩, + refine ⟨r, hr.mono (λ y hy, ⟨hr0, _⟩)⟩, + rwa [mem_preimage, mem_Inter₂] +end + +lemma exists_forall_closed_ball_subset_aux₂ (y : X) : + convex ℝ (Ioi (0 : ℝ) ∩ ennreal.of_real ⁻¹' ⋂ i (hi : y ∈ K i), {r | closed_ball y r ⊆ U i}) := +(convex_Ioi _).inter $ ord_connected.convex $ ord_connected.preimage_ennreal_of_real $ + ord_connected_Inter $ λ i, ord_connected_Inter $ + λ hi, ord_connected_set_of_closed_ball_subset y (U i) + +/-- Let `X` be an extended 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 positive continuous function `δ : C(X, ℝ)` such that for any `i` and `x ∈ K i`, +we have `emetric.closed_ball x (ennreal.of_real (δ x)) ⊆ U i`. -/ +lemma exists_continuous_real_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : + ∃ δ : C(X, ℝ), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (ennreal.of_real $ δ x) ⊆ U i := +by simpa only [mem_inter_eq, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι X] + using exists_continuous_forall_mem_convex_of_local_const exists_forall_closed_ball_subset_aux₂ + (exists_forall_closed_ball_subset_aux₁ hK hU hKU hfin) + +/-- Let `X` be an extended 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 positive continuous function `δ : C(X, ℝ≥0)` such that for any `i` and `x ∈ K i`, +we have `emetric.closed_ball x (δ x) ⊆ U i`. -/ +lemma exists_continuous_nnreal_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : + ∃ δ : C(X, ℝ≥0), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (δ x) ⊆ U i := +begin + rcases exists_continuous_real_forall_closed_ball_subset hK hU hKU hfin with ⟨δ, hδ₀, hδ⟩, + lift δ to C(X, ℝ≥0) using λ x, (hδ₀ x).le, + refine ⟨δ, hδ₀, λ i x hi, _⟩, + simpa only [← ennreal.of_real_coe_nnreal] using hδ i x hi +end + +/-- Let `X` be an extended 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 positive continuous function `δ : C(X, ℝ≥0∞)` such that for any `i` and `x ∈ K i`, +we have `emetric.closed_ball x (δ x) ⊆ U i`. -/ +lemma exists_continuous_ennreal_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : + ∃ δ : C(X, ℝ≥0∞), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (δ x) ⊆ U i := +let ⟨δ, hδ₀, hδ⟩ := exists_continuous_nnreal_forall_closed_ball_subset hK hU hKU hfin +in ⟨continuous_map.comp ⟨coe, ennreal.continuous_coe⟩ δ, λ x, ennreal.coe_pos.2 (hδ₀ x), hδ⟩ + +end emetric + +namespace metric + +variables [metric_space X] {K : ι → set X} {U : ι → set X} + +/-- 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 +positive continuous function `δ : C(X, ℝ≥0)` such that for any `i` and `x ∈ K i`, we have +`metric.closed_ball x (δ x) ⊆ U i`. -/ +lemma exists_continuous_nnreal_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : + ∃ δ : C(X, ℝ≥0), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (δ x) ⊆ U i := +begin + rcases emetric.exists_continuous_nnreal_forall_closed_ball_subset hK hU hKU hfin + with ⟨δ, hδ0, hδ⟩, + refine ⟨δ, hδ0, λ i x hx, _⟩, + rw [← emetric_closed_ball_nnreal], + exact hδ i x hx +end + +/-- 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 +positive continuous function `δ : C(X, ℝ)` such that for any `i` and `x ∈ K i`, we have +`metric.closed_ball x (δ x) ⊆ U i`. -/ +lemma exists_continuous_real_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) + (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : + ∃ δ : C(X, ℝ), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (δ x) ⊆ U i := +let ⟨δ, hδ₀, hδ⟩ := exists_continuous_nnreal_forall_closed_ball_subset hK hU hKU hfin +in ⟨continuous_map.comp ⟨coe, nnreal.continuous_coe⟩ δ, hδ₀, hδ⟩ + +end metric From 6c967df32b2ec034596c8c5b18185e6e9698d978 Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Mon, 25 Jul 2022 10:24:27 +0000 Subject: [PATCH 041/173] feat(data/matrix/basic): Add `alg_equiv` and `linear_equiv` instances for transpose. (#15386) `transpose` has natural bundlings as an `alg_equiv` and a `linear_equiv` for which we already have the substantial lemmas. Similarly, `conj_transpose` can be bundled as a `linear_equiv`. This also alters the other bundled versions to take explicit variables as this saves the need for many type annotations, and makes the necessary edits to fix proofs. Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com> Co-authored-by: Eric Wieser --- src/data/matrix/basic.lean | 68 +++++++++++++++++++++++++----- src/ring_theory/trace.lean | 2 +- src/topology/instances/matrix.lean | 8 ++-- 3 files changed, 63 insertions(+), 15 deletions(-) diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 92dccae4f1f00..9da0bdad20444 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -1369,6 +1369,8 @@ by ext i j; refl lemma transpose_map {f : α → β} {M : matrix m n α} : Mᵀ.map f = (M.map f)ᵀ := by { ext, refl } +variables (m n α) + /-- `matrix.transpose` as an `add_equiv` -/ @[simps apply] def transpose_add_equiv [has_add α] : matrix m n α ≃+ matrix n m α := @@ -1379,19 +1381,33 @@ def transpose_add_equiv [has_add α] : matrix m n α ≃+ matrix n m α := map_add' := transpose_add } @[simp] lemma transpose_add_equiv_symm [has_add α] : - (transpose_add_equiv : matrix m n α ≃+ matrix n m α).symm = transpose_add_equiv := rfl + (transpose_add_equiv m n α).symm = transpose_add_equiv n m α := rfl + +variables {m n α} lemma transpose_list_sum [add_monoid α] (l : list (matrix m n α)) : l.sumᵀ = (l.map transpose).sum := -(transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_list_sum l +(transpose_add_equiv m n α).to_add_monoid_hom.map_list_sum l lemma transpose_multiset_sum [add_comm_monoid α] (s : multiset (matrix m n α)) : s.sumᵀ = (s.map transpose).sum := -(transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_multiset_sum s +(transpose_add_equiv m n α).to_add_monoid_hom.map_multiset_sum s lemma transpose_sum [add_comm_monoid α] {ι : Type*} (s : finset ι) (M : ι → matrix m n α) : (∑ i in s, M i)ᵀ = ∑ i in s, (M i)ᵀ := -(transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_sum _ s +(transpose_add_equiv m n α).to_add_monoid_hom.map_sum _ s + +variables (m n R α) + +/-- `matrix.transpose` as a `linear_map` -/ +@[simps apply] +def transpose_linear_equiv [semiring R] [add_comm_monoid α] [module R α] : + matrix m n α ≃ₗ[R] matrix n m α := { map_smul' := transpose_smul, ..transpose_add_equiv m n α} + +@[simp] lemma transpose_linear_equiv_symm [semiring R] [add_comm_monoid α] [module R α] : + (transpose_linear_equiv m n R α).symm = transpose_linear_equiv n m R α := rfl + +variables {m n R α} variables (m α) @@ -1403,7 +1419,7 @@ 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 _ _), - ..transpose_add_equiv.trans mul_opposite.op_add_equiv } + ..(transpose_add_equiv m m α).trans mul_opposite.op_add_equiv } variables {m α} @@ -1415,6 +1431,20 @@ lemma transpose_list_prod [comm_semiring α] [fintype m] [decidable_eq m] (l : l l.prodᵀ = (l.map transpose).reverse.prod := (transpose_ring_equiv m α).unop_map_list_prod l +variables (R m α) + +/-- `matrix.transpose` as an `alg_equiv` to the opposite ring -/ +@[simps] +def transpose_alg_equiv [comm_semiring R] [comm_semiring α] [fintype m] [decidable_eq m] + [algebra R α] : matrix m m α ≃ₐ[R] (matrix m m α)ᵐᵒᵖ := +{ to_fun := λ M, mul_opposite.op (Mᵀ), + commutes' := λ r, by simp only [algebra_map_eq_diagonal, diagonal_transpose, + mul_opposite.algebra_map_apply], + ..(transpose_add_equiv m m α).trans mul_opposite.op_add_equiv, + ..transpose_ring_equiv m α } + +variables {R m α} + end transpose section conj_transpose @@ -1520,6 +1550,8 @@ lemma conj_transpose_map [has_star α] [has_star β] {A : matrix m n α} (f : α Aᴴ.map f = (A.map f)ᴴ := matrix.ext $ λ i j, hf _ +variables (m n α) + /-- `matrix.conj_transpose` as an `add_equiv` -/ @[simps apply] def conj_transpose_add_equiv [add_monoid α] [star_add_monoid α] : matrix m n α ≃+ matrix n m α := @@ -1530,21 +1562,37 @@ def conj_transpose_add_equiv [add_monoid α] [star_add_monoid α] : matrix m n map_add' := conj_transpose_add } @[simp] lemma conj_transpose_add_equiv_symm [add_monoid α] [star_add_monoid α] : - (conj_transpose_add_equiv : matrix m n α ≃+ matrix n m α).symm = conj_transpose_add_equiv := rfl + (conj_transpose_add_equiv m n α).symm = conj_transpose_add_equiv n m α := rfl + +variables {m n α} lemma conj_transpose_list_sum [add_monoid α] [star_add_monoid α] (l : list (matrix m n α)) : l.sumᴴ = (l.map conj_transpose).sum := -(conj_transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_list_sum l +(conj_transpose_add_equiv m n α).to_add_monoid_hom.map_list_sum l lemma conj_transpose_multiset_sum [add_comm_monoid α] [star_add_monoid α] (s : multiset (matrix m n α)) : s.sumᴴ = (s.map conj_transpose).sum := -(conj_transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_multiset_sum s +(conj_transpose_add_equiv m n α).to_add_monoid_hom.map_multiset_sum s lemma conj_transpose_sum [add_comm_monoid α] [star_add_monoid α] {ι : Type*} (s : finset ι) (M : ι → matrix m n α) : (∑ i in s, M i)ᴴ = ∑ i in s, (M i)ᴴ := -(conj_transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_sum _ s +(conj_transpose_add_equiv m n α).to_add_monoid_hom.map_sum _ s + +variables (m n R α) + +/-- `matrix.conj_transpose` as a `linear_map` -/ +@[simps apply] +def conj_transpose_linear_equiv [comm_semiring R] [star_ring R] [add_comm_monoid α] + [star_add_monoid α] [module R α] [star_module R α] : matrix m n α ≃ₗ⋆[R] matrix n m α := +{ map_smul' := conj_transpose_smul, ..conj_transpose_add_equiv m n α} + +@[simp] lemma conj_transpose_linear_equiv_symm [comm_semiring R] [star_ring R] [add_comm_monoid α] + [star_add_monoid α] [module R α] [star_module R α] : + (conj_transpose_linear_equiv m n R α).symm = conj_transpose_linear_equiv n m R α := rfl + +variables {m n R α} variables (m α) @@ -1556,7 +1604,7 @@ def conj_transpose_ring_equiv [semiring α] [star_ring α] [fintype m] : inv_fun := λ M, M.unopᴴ, map_mul' := λ M N, (congr_arg mul_opposite.op (conj_transpose_mul M N)).trans (mul_opposite.op_mul _ _), - ..conj_transpose_add_equiv.trans mul_opposite.op_add_equiv } + ..(conj_transpose_add_equiv m m α).trans mul_opposite.op_add_equiv } variables {m α} diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 31f854f1320c6..fc9b832baf97a 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -420,7 +420,7 @@ end lemma trace_matrix_of_matrix_mul_vec [fintype κ] (b : κ → B) (P : matrix κ κ A) : trace_matrix A ((P.map (algebra_map A B)).mul_vec b) = P ⬝ (trace_matrix A b) ⬝ Pᵀ := begin - refine add_equiv.injective transpose_add_equiv _, + refine add_equiv.injective (transpose_add_equiv _ _ _) _, rw [transpose_add_equiv_apply, transpose_add_equiv_apply, ← vec_mul_transpose, ← transpose_map, trace_matrix_of_matrix_vec_mul, transpose_transpose, transpose_mul, transpose_transpose, transpose_mul] diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index 2ecca9e4fa066..ad4bdca8d964b 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -267,7 +267,7 @@ variables [semiring α] [add_comm_monoid R] [topological_space R] [module α R] lemma has_sum.matrix_transpose {f : X → matrix m n R} {a : matrix m n R} (hf : has_sum f a) : has_sum (λ x, (f x)ᵀ) aᵀ := -(hf.map (@matrix.transpose_add_equiv m n R _) continuous_id.matrix_transpose : _) +(hf.map (matrix.transpose_add_equiv m n R) continuous_id.matrix_transpose : _) lemma summable.matrix_transpose {f : X → matrix m n R} (hf : summable f) : summable (λ x, (f x)ᵀ) := @@ -275,7 +275,7 @@ hf.has_sum.matrix_transpose.summable @[simp] lemma summable_matrix_transpose {f : X → matrix m n R} : summable (λ x, (f x)ᵀ) ↔ summable f := -(summable.map_iff_of_equiv (@matrix.transpose_add_equiv m n R _) +(summable.map_iff_of_equiv (matrix.transpose_add_equiv m n R) (@continuous_id (matrix m n R) _).matrix_transpose (continuous_id.matrix_transpose) : _) lemma matrix.transpose_tsum [t2_space R] {f : X → matrix m n R} : (∑' x, f x)ᵀ = ∑' x, (f x)ᵀ := @@ -289,7 +289,7 @@ end lemma has_sum.matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] {f : X → matrix m n R} {a : matrix m n R} (hf : has_sum f a) : has_sum (λ x, (f x)ᴴ) aᴴ := -(hf.map (@matrix.conj_transpose_add_equiv m n R _ _) continuous_id.matrix_conj_transpose : _) +(hf.map (matrix.conj_transpose_add_equiv m n R) continuous_id.matrix_conj_transpose : _) lemma summable.matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] {f : X → matrix m n R} (hf : summable f) : @@ -299,7 +299,7 @@ hf.has_sum.matrix_conj_transpose.summable @[simp] lemma summable_matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] {f : X → matrix m n R} : summable (λ x, (f x)ᴴ) ↔ summable f := -(summable.map_iff_of_equiv (@matrix.conj_transpose_add_equiv m n R _ _) +(summable.map_iff_of_equiv (matrix.conj_transpose_add_equiv m n R) (@continuous_id (matrix m n R) _).matrix_conj_transpose (continuous_id.matrix_conj_transpose) : _) lemma matrix.conj_transpose_tsum [star_add_monoid R] [has_continuous_star R] [t2_space R] From 07d18508d0c47e09938b0fccf6e893c44536be66 Mon Sep 17 00:00:00 2001 From: Justin Thomas Date: Mon, 25 Jul 2022 12:01:28 +0000 Subject: [PATCH 042/173] feat(linear_algebra/annihilating_polynomial): add definition of annihilating ideal and show minpoly generates in field case (#12140) adding item from trivial undergrad subjects list Co-authored-by: Patrick Massot Co-authored-by: Anne Baanen Co-authored-by: Vierkantor --- docs/undergrad.yaml | 2 +- src/data/polynomial/ring_division.lean | 6 + .../annihilating_polynomial.lean | 178 ++++++++++++++++++ 3 files changed, 185 insertions(+), 1 deletion(-) create mode 100644 src/linear_algebra/annihilating_polynomial.lean diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 4df6c8fac7a73..4b39ab97ab21b 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -51,7 +51,7 @@ Linear algebra: Gaussian elimination: 'https://en.wikipedia.org/wiki/Gaussian_elimination' row-reduced matrices: 'https://en.wikipedia.org/wiki/Row_echelon_form#Reduced_row_echelon_form' Endomorphism polynomials: - annihilating polynomials: 'https://en.wikipedia.org/wiki/Minimal_polynomial_(linear_algebra)#Formal_definition' + annihilating polynomials: 'polynomial.ann_ideal' minimal polynomial: 'minpoly' characteristic polynomial: 'matrix.charpoly' Cayley-Hamilton theorem: 'matrix.aeval_self_charpoly' diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index dc316738a09c8..3d4d974f07d0a 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -144,6 +144,12 @@ begin rw [nat_degree_mul h2.1 h2.2], exact nat.le_add_right _ _ end +lemma degree_le_of_dvd {p q : R[X]} (h1 : p ∣ q) (h2 : q ≠ 0) : degree p ≤ degree q := +begin + rcases h1 with ⟨q, rfl⟩, rw mul_ne_zero_iff at h2, + exact degree_le_mul_left p h2.2, +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₂ : polynomial R} (hp₁ : p₁ ≠ 0) (hq₁ : q₁ ≠ 0) (hp₂ : p₂ ≠ 0) (hq₂ : q₂ ≠ 0) (h_eq : p₁ * q₂ = p₂ * q₁) : diff --git a/src/linear_algebra/annihilating_polynomial.lean b/src/linear_algebra/annihilating_polynomial.lean new file mode 100644 index 0000000000000..3685247bdb834 --- /dev/null +++ b/src/linear_algebra/annihilating_polynomial.lean @@ -0,0 +1,178 @@ +/- +Copyright (c) 2022 Justin Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Justin Thomas +-/ +import data.set.basic +import field_theory.minpoly +import ring_theory.principal_ideal_domain +import ring_theory.polynomial_algebra + +/-! +# Annihilating Ideal + +Given a commutative ring `R` and an `R`-algebra `A` +Every element `a : A` defines +an ideal `polynomial.ann_ideal a ⊆ R[X]`. +Simply put, this is the set of polynomials `p` where +the polynomial evaluation `p(a)` is 0. + +## Special case where the ground ring is a field + +In the special case that `R` is a field, we use the notation `R = 𝕜`. +Here `𝕜[X]` is a PID, so there is a polynomial `g ∈ polynomial.ann_ideal a` +which generates the ideal. We show that if this generator is +chosen to be monic, then it is the minimal polynomial of `a`, +as defined in `field_theory.minpoly`. + +## Special case: endomorphism algebra + +Given an `R`-module `M` (`[add_comm_group M] [module R M]`) +there are some common specializations which may be more familiar. +* Example 1: `A = M →ₗ[R] M`, the endomorphism algebra of an `R`-module M. +* Example 2: `A = n × n` matrices with entries in `R`. +-/ + +open_locale polynomial + +namespace polynomial + +section semiring + +variables {R A : Type*} [comm_semiring R] [semiring A] [algebra R A] + +variables (R) + +/-- `ann_ideal R a` is the *annihilating ideal* of all `p : R[X]` such that `p(a) = 0`. + +The informal notation `p(a)` stand for `polynomial.aeval a p`. +Again informally, the annihilating ideal of `a` is +`{ p ∈ R[X] | p(a) = 0 }`. This is an ideal in `R[X]`. +The formal definition uses the kernel of the aeval map. -/ +noncomputable def ann_ideal (a : A) : ideal R[X] := +((aeval a).to_ring_hom : R[X] →+* A).ker + +variables {R} + +/-- It is useful to refer to ideal membership sometimes + and the annihilation condition other times. -/ +lemma mem_ann_ideal_iff_aeval_eq_zero {a : A} {p : R[X]} : + p ∈ ann_ideal R a ↔ aeval a p = 0 := +iff.rfl + +end semiring + +section field + +variables {𝕜 A : Type*} [field 𝕜] [ring A] [algebra 𝕜 A] +variable (𝕜) + +open submodule + +/-- `ann_ideal_generator 𝕜 a` is the monic generator of `ann_ideal 𝕜 a` +if one exists, otherwise `0`. + +Since `𝕜[X]` is a principal ideal domain there is a polynomial `g` such that + `span 𝕜 {g} = ann_ideal a`. This picks some generator. + We prefer the monic generator of the ideal. -/ +noncomputable def ann_ideal_generator (a : A) : 𝕜[X] := +let g := is_principal.generator $ ann_ideal 𝕜 a + in g * (C g.leading_coeff⁻¹) + +section + +variables {𝕜} + +@[simp] lemma ann_ideal_generator_eq_zero_iff {a : A} : + ann_ideal_generator 𝕜 a = 0 ↔ ann_ideal 𝕜 a = ⊥ := +by simp only [ann_ideal_generator, mul_eq_zero, is_principal.eq_bot_iff_generator_eq_zero, + polynomial.C_eq_zero, inv_eq_zero, polynomial.leading_coeff_eq_zero, or_self] +end + +/-- `ann_ideal_generator 𝕜 a` is indeed a generator. -/ +@[simp] lemma span_singleton_ann_ideal_generator (a : A) : + ideal.span {ann_ideal_generator 𝕜 a} = ann_ideal 𝕜 a := +begin + by_cases h : ann_ideal_generator 𝕜 a = 0, + { rw [h, ann_ideal_generator_eq_zero_iff.mp h, set.singleton_zero, ideal.span_zero] }, + { rw [ann_ideal_generator, ideal.span_singleton_mul_right_unit, ideal.span_singleton_generator], + apply polynomial.is_unit_C.mpr, + apply is_unit.mk0, + apply inv_eq_zero.not.mpr, + apply polynomial.leading_coeff_eq_zero.not.mpr, + apply (mul_ne_zero_iff.mp h).1 } +end + +/-- The annihilating ideal generator is a member of the annihilating ideal. -/ +lemma ann_ideal_generator_mem (a : A) : ann_ideal_generator 𝕜 a ∈ ann_ideal 𝕜 a := +ideal.mul_mem_right _ _ (submodule.is_principal.generator_mem _) + +lemma mem_iff_eq_smul_ann_ideal_generator {p : 𝕜[X]} (a : A) : + p ∈ ann_ideal 𝕜 a ↔ ∃ s : 𝕜[X], p = s • ann_ideal_generator 𝕜 a := +by simp_rw [@eq_comm _ p, ← mem_span_singleton, ← span_singleton_ann_ideal_generator 𝕜 a, + ideal.span] + +/-- The generator we chose for the annihilating ideal is monic when the ideal is non-zero. -/ +lemma monic_ann_ideal_generator (a : A) (hg : ann_ideal_generator 𝕜 a ≠ 0) : + monic (ann_ideal_generator 𝕜 a) := +monic_mul_leading_coeff_inv (mul_ne_zero_iff.mp hg).1 + +/-! We are working toward showing the generator of the annihilating ideal +in the field case is the minimal polynomial. We are going to use a uniqueness +theorem of the minimal polynomial. + +This is the first condition: it must annihilate the original element `a : A`. -/ +lemma ann_ideal_generator_aeval_eq_zero (a : A) : + aeval a (ann_ideal_generator 𝕜 a) = 0 := +mem_ann_ideal_iff_aeval_eq_zero.mp (ann_ideal_generator_mem 𝕜 a) + +variables {𝕜} + +lemma mem_iff_ann_ideal_generator_dvd {p : 𝕜[X]} {a : A} : + p ∈ ann_ideal 𝕜 a ↔ ann_ideal_generator 𝕜 a ∣ p := +by rw [← ideal.mem_span_singleton, span_singleton_ann_ideal_generator] + +/-- The generator of the annihilating ideal has minimal degree among + the non-zero members of the annihilating ideal -/ +lemma degree_ann_ideal_generator_le_of_mem (a : A) (p : 𝕜[X]) + (hp : p ∈ ann_ideal 𝕜 a) (hpn0 : p ≠ 0) : + degree (ann_ideal_generator 𝕜 a) ≤ degree p := +degree_le_of_dvd (mem_iff_ann_ideal_generator_dvd.1 hp) hpn0 + +variables (𝕜) + +/-- The generator of the annihilating ideal is the minimal polynomial. -/ +lemma ann_ideal_generator_eq_minpoly (a : A) : + ann_ideal_generator 𝕜 a = minpoly 𝕜 a := +begin + by_cases h : ann_ideal_generator 𝕜 a = 0, + { rw [h, minpoly.eq_zero], + rintro ⟨p, p_monic, (hp : aeval a p = 0)⟩, + refine p_monic.ne_zero (ideal.mem_bot.mp _), + simpa only [ann_ideal_generator_eq_zero_iff.mp h] + using mem_ann_ideal_iff_aeval_eq_zero.mpr hp }, + { exact minpoly.unique _ _ + (monic_ann_ideal_generator _ _ h) + (ann_ideal_generator_aeval_eq_zero _ _) + (λ q q_monic hq, (degree_ann_ideal_generator_le_of_mem a q + (mem_ann_ideal_iff_aeval_eq_zero.mpr hq) + q_monic.ne_zero)) } +end + +/-- If a monic generates the annihilating ideal, it must match our choice + of the annihilating ideal generator. -/ +lemma monic_generator_eq_minpoly (a : A) (p : 𝕜[X]) + (p_monic : p.monic) (p_gen : ideal.span {p} = ann_ideal 𝕜 a) : + ann_ideal_generator 𝕜 a = p := +begin + by_cases h : p = 0, + { rwa [h, ann_ideal_generator_eq_zero_iff, ← p_gen, ideal.span_singleton_eq_bot.mpr], }, + { rw [← span_singleton_ann_ideal_generator, ideal.span_singleton_eq_span_singleton] at p_gen, + rw eq_comm, + apply eq_of_monic_of_associated p_monic _ p_gen, + { apply monic_ann_ideal_generator _ _ ((associated.ne_zero_iff p_gen).mp h), }, }, +end + +end field + +end polynomial From 3822eb26bc0e9424a713ed9d7ac9cf57ade344fc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 25 Jul 2022 12:01:29 +0000 Subject: [PATCH 043/173] feat(order/compare): general cleanup (#15665) We add `swap_inj`, golf some lemmas, do some simple spacing tweaks. --- src/order/compare.lean | 28 ++++++++++++---------------- 1 file changed, 12 insertions(+), 16 deletions(-) diff --git a/src/order/compare.lean b/src/order/compare.lean index 5dd03daff8d5c..03de3a0ec82ac 100644 --- a/src/order/compare.lean +++ b/src/order/compare.lean @@ -58,23 +58,23 @@ by { cases o, exacts [iff.rfl, eq_comm, iff.rfl] } alias compares_swap ↔ compares.of_swap compares.swap +@[simp] theorem swap_inj (o₁ o₂ : ordering) : o₁.swap = o₂.swap ↔ o₁ = o₂ := +by cases o₁; cases o₂; dec_trivial + lemma swap_eq_iff_eq_swap {o o' : ordering} : o.swap = o' ↔ o = o'.swap := -⟨λ h, by rw [← swap_swap o, h], λ h, by rw [← swap_swap o', h]⟩ +by rw [←swap_inj, swap_swap] -lemma compares.eq_lt [preorder α] : - ∀ {o} {a b : α}, compares o a b → (o = lt ↔ a < b) +lemma compares.eq_lt [preorder α] : ∀ {o} {a b : α}, compares o a b → (o = lt ↔ a < b) | lt a b h := ⟨λ _, h, λ _, rfl⟩ | eq a b h := ⟨λ h, by injection h, λ h', (ne_of_lt h' h).elim⟩ | gt a b h := ⟨λ h, by injection h, λ h', (lt_asymm h h').elim⟩ -lemma compares.ne_lt [preorder α] : - ∀ {o} {a b : α}, compares o a b → (o ≠ lt ↔ b ≤ a) +lemma compares.ne_lt [preorder α] : ∀ {o} {a b : α}, compares o a b → (o ≠ lt ↔ b ≤ a) | lt a b h := ⟨absurd rfl, λ h', (not_le_of_lt h h').elim⟩ | eq a b h := ⟨λ _, ge_of_eq h, λ _ h, by injection h⟩ | gt a b h := ⟨λ _, le_of_lt h, λ _ h, by injection h⟩ -lemma compares.eq_eq [preorder α] : - ∀ {o} {a b : α}, compares o a b → (o = eq ↔ a = b) +lemma compares.eq_eq [preorder α] : ∀ {o} {a b : α}, compares o a b → (o = eq ↔ a = b) | lt a b h := ⟨λ h, by injection h, λ h', (ne_of_lt h h').elim⟩ | eq a b h := ⟨λ _, h, λ _, rfl⟩ | gt a b h := ⟨λ h, by injection h, λ h', (ne_of_gt h h').elim⟩ @@ -85,20 +85,17 @@ swap_eq_iff_eq_swap.symm.trans h.swap.eq_lt lemma compares.ne_gt [preorder α] {o} {a b : α} (h : compares o a b) : (o ≠ gt ↔ a ≤ b) := (not_congr swap_eq_iff_eq_swap.symm).trans h.swap.ne_lt -lemma compares.le_total [preorder α] {a b : α} : - ∀ {o}, compares o a b → a ≤ b ∨ b ≤ a +lemma compares.le_total [preorder α] {a b : α} : ∀ {o}, compares o a b → a ≤ b ∨ b ≤ a | lt h := or.inl (le_of_lt h) | eq h := or.inl (le_of_eq h) | gt h := or.inr (le_of_lt h) -lemma compares.le_antisymm [preorder α] {a b : α} : - ∀ {o}, compares o a b → a ≤ b → b ≤ a → a = b +lemma compares.le_antisymm [preorder α] {a b : α} : ∀ {o}, compares o a b → a ≤ b → b ≤ a → a = b | lt h _ hba := (not_le_of_lt h hba).elim | eq h _ _ := h | gt h hab _ := (not_le_of_lt h hab).elim -lemma compares.inj [preorder α] {o₁} : - ∀ {o₂} {a b : α}, compares o₁ a b → compares o₂ a b → o₁ = o₂ +lemma compares.inj [preorder α] {o₁} : ∀ {o₂} {a b : α}, compares o₁ a b → compares o₂ a b → o₁ = o₂ | lt a b h₁ h₂ := h₁.eq_lt.2 h₂ | eq a b h₁ h₂ := h₁.eq_eq.2 h₂ | gt a b h₁ h₂ := h₁.eq_gt.2 h₂ @@ -143,7 +140,7 @@ lemma ordering.compares.cmp_eq [linear_order α] {a b : α} {o : ordering} (h : cmp a b = o := (cmp_compares a b).inj h -lemma cmp_swap [preorder α] [@decidable_rel α (<)] (a b : α) : (cmp a b).swap = cmp b a := +@[simp] lemma cmp_swap [preorder α] [@decidable_rel α (<)] (a b : α) : (cmp a b).swap = cmp b a := begin unfold cmp cmp_using, by_cases a < b; by_cases h₂ : b < a; simp [h, h₂, ordering.swap], @@ -190,8 +187,7 @@ by rw cmp_eq_eq_iff variables {x y} {β : Type*} [linear_order β] {x' y' : β} lemma cmp_eq_cmp_symm : cmp x y = cmp x' y' ↔ cmp y x = cmp y' x' := -by { split, rw [←cmp_swap _ y, ←cmp_swap _ y'], cc, - rw [←cmp_swap _ x, ←cmp_swap _ x'], cc, } +⟨λ h, by rwa [←cmp_swap x', ←cmp_swap, swap_inj], λ h, by rwa [←cmp_swap y', ←cmp_swap, swap_inj]⟩ lemma lt_iff_lt_of_cmp_eq_cmp (h : cmp x y = cmp x' y') : x < y ↔ x' < y' := by rw [←cmp_eq_lt_iff, ←cmp_eq_lt_iff, h] From fa120115d3b66739eba0157458582d70e24b1112 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Jul 2022 14:27:48 +0000 Subject: [PATCH 044/173] feat(combinatorics/set_family/kleitman): Kleitman's bound (#14543) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The union of `k` intersecting families over a type of cardinality `n` has size at most `2ⁿ - 2ⁿ⁻ᵏ`. This is the main result in [Families of Non-disjoint subsets](https://reader.elsevier.com/reader/sd/pii/S0021980066800121). --- src/combinatorics/set_family/kleitman.lean | 77 ++++++++++++++++++++++ 1 file changed, 77 insertions(+) create mode 100644 src/combinatorics/set_family/kleitman.lean diff --git a/src/combinatorics/set_family/kleitman.lean b/src/combinatorics/set_family/kleitman.lean new file mode 100644 index 0000000000000..daefabc28c79d --- /dev/null +++ b/src/combinatorics/set_family/kleitman.lean @@ -0,0 +1,77 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import combinatorics.set_family.harris_kleitman +import combinatorics.set_family.intersecting + +/-! +# Kleitman's bound on the size of intersecting families + +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. +Kleitman's bound stipulates that `k` intersecting families cover at most `2ⁿ - 2ⁿ⁻ᵏ` sets. + +## Main declarations + +* `finset.card_bUnion_le_of_intersecting`: Kleitman's theorem. + +## References + +* [D. J. Kleitman, *Families of non-disjoint subsets*][kleitman1966] +-/ + +open finset fintype (card) + +variables {ι α : Type*} [fintype α] [decidable_eq α] [nonempty α] + +/-- **Kleitman's theorem**. An intersecting family on `n` elements contains at most `2ⁿ⁻¹` sets, and +each further intersecting family takes at most half of the sets that are in no previous family. -/ +lemma finset.card_bUnion_le_of_intersecting (s : finset ι) (f : ι → finset (finset α)) + (hf : ∀ i ∈ s, (f i : set (finset α)).intersecting) : + (s.bUnion f).card ≤ 2 ^ card α - 2 ^ (card α - s.card) := +begin + obtain hs | hs := le_total (card α) s.card, + { rw [tsub_eq_zero_of_le hs, pow_zero], + refine (card_le_of_subset $ bUnion_subset.2 $ λ i hi a ha, mem_compl.2 $ not_mem_singleton.2 $ + (hf _ hi).ne_bot ha).trans_eq _, + rw [card_compl, fintype.card_finset, card_singleton] }, + induction s using finset.cons_induction with i s hi ih generalizing f, + { simp }, + classical, + set f' : ι → finset (finset α) := λ j, + if hj : j ∈ cons i s hi then (hf j hj).exists_card_eq.some else ∅ with hf', + have hf₁ : ∀ j, j ∈ cons i s hi → + f j ⊆ f' j ∧ 2 * (f' j).card = 2 ^ card α ∧ (f' j : set (finset α)).intersecting, + { rintro j hj, + simp_rw [hf', dif_pos hj, ←fintype.card_finset], + exact classical.some_spec (hf j hj).exists_card_eq }, + have hf₂ : ∀ j, j ∈ cons i s hi → is_upper_set (f' j : set (finset α)), + { refine λ j hj, (hf₁ _ hj).2.2.is_upper_set' ((hf₁ _ hj).2.2.is_max_iff_card_eq.2 _), + rw fintype.card_finset, + exact (hf₁ _ hj).2.1 }, + refine (card_le_of_subset $ bUnion_mono $ λ j hj, (hf₁ _ hj).1).trans _, + nth_rewrite 0 cons_eq_insert i, + rw bUnion_insert, + refine (card_mono $ @le_sup_sdiff _ (f' i) _ _).trans ((card_union_le _ _).trans _), + rw [union_sdiff_left, sdiff_eq_inter_compl], + refine le_of_mul_le_mul_left _ (pow_pos zero_lt_two $ card α + 1), + rw [pow_succ', mul_add, mul_assoc, mul_comm _ 2, mul_assoc], + refine (add_le_add ((mul_le_mul_left $ pow_pos two_pos _).2 + (hf₁ _ $ mem_cons_self _ _).2.2.card_le) $ (mul_le_mul_left two_pos).2 $ + is_upper_set.card_inter_le_finset _ _).trans _, + { rw coe_bUnion, + exact is_upper_set_Union₂ (λ i hi, hf₂ _ $ subset_cons _ hi) }, + { rw coe_compl, + exact (hf₂ _ $ mem_cons_self _ _).compl }, + rw [mul_tsub, card_compl, fintype.card_finset, mul_left_comm, mul_tsub, + (hf₁ _ $ mem_cons_self _ _).2.1, two_mul, add_tsub_cancel_left, ←mul_tsub, ←mul_two, mul_assoc, + ←add_mul, mul_comm], + refine mul_le_mul_left' _ _, + refine (add_le_add_left (ih ((card_le_of_subset $ subset_cons _).trans hs) _ $ λ i hi, + (hf₁ _ $ subset_cons _ hi).2.2) _).trans _, + rw [mul_tsub, two_mul, ←pow_succ, ←add_tsub_assoc_of_le (pow_le_pow' (@one_le_two ℕ _ _ _ _ _) + tsub_le_self), tsub_add_eq_add_tsub hs, card_cons, add_tsub_add_eq_tsub_right], +end From a193cd769d9d6ebe2a8daaa06a8b57f29cfff9d3 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Mon, 25 Jul 2022 14:27:50 +0000 Subject: [PATCH 045/173] refactor(measure_theory/function/uniform_integrable): change `uniform_integrable` to only require `ae_strongly_measurable` (#15623) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The L¹ version of the strong LLN does not require `strongly_measurable` but the assumption in `uniform_integrable` forces it to have this condition if not requiring an extra step to relax the condition (see #15392). This PR relaxes the definition of `uniform_integrable` so it only requires `ae_strongly_measurable`. --- .../function/conditional_expectation.lean | 2 +- .../function/uniform_integrable.lean | 138 +++++++++++++----- 2 files changed, 100 insertions(+), 40 deletions(-) diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index 245ad2e843b1e..2fd940dcc4024 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -2385,7 +2385,7 @@ begin (strongly_measurable_condexp.mono (hℱ n)).measurable.nnnorm, have hg : mem_ℒp g 1 μ := mem_ℒp_one_iff_integrable.2 hint, refine uniform_integrable_of le_rfl ennreal.one_ne_top - (λ n, strongly_measurable_condexp.mono (hℱ n)) (λ ε hε, _), + (λ n, (strongly_measurable_condexp.mono (hℱ n)).ae_strongly_measurable) (λ ε hε, _), by_cases hne : snorm g 1 μ = 0, { rw snorm_eq_zero_iff hg.1 one_ne_zero at hne, refine ⟨0, λ n, (le_of_eq $ (snorm_eq_zero_iff diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index ff7dac8012669..865face4a0deb 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -65,21 +65,24 @@ snorm (s.indicator (f i)) p μ ≤ ennreal.of_real ε uniformly integrable in the measure theory sense and is uniformly bounded. -/ def uniform_integrable {m : measurable_space α} (f : ι → α → β) (p : ℝ≥0∞) (μ : measure α) : Prop := -(∀ i, strongly_measurable (f i)) ∧ unif_integrable f p μ ∧ ∃ C : ℝ≥0, ∀ i, snorm (f i) p μ ≤ C +(∀ i, ae_strongly_measurable (f i) μ) ∧ unif_integrable f p μ ∧ ∃ C : ℝ≥0, ∀ i, snorm (f i) p μ ≤ C -lemma uniform_integrable.strongly_measurable {f : ι → α → β} {p : ℝ≥0∞} - (hf : uniform_integrable f p μ) (i : ι) : strongly_measurable (f i) := +namespace uniform_integrable + +protected lemma ae_strongly_measurable {f : ι → α → β} {p : ℝ≥0∞} + (hf : uniform_integrable f p μ) (i : ι) : ae_strongly_measurable (f i) μ := hf.1 i -lemma uniform_integrable.unif_integrable {f : ι → α → β} {p : ℝ≥0∞} +protected lemma unif_integrable {f : ι → α → β} {p : ℝ≥0∞} (hf : uniform_integrable f p μ) : unif_integrable f p μ := hf.2.1 -lemma uniform_integrable.mem_ℒp {f : ι → α → β} {p : ℝ≥0∞} +protected lemma mem_ℒp {f : ι → α → β} {p : ℝ≥0∞} (hf : uniform_integrable f p μ) (i : ι) : mem_ℒp (f i) p μ := -⟨(hf.1 i).ae_strongly_measurable, -let ⟨_, _, hC⟩ := hf.2 in lt_of_le_of_lt (hC i) ennreal.coe_lt_top⟩ +⟨hf.1 i, let ⟨_, _, hC⟩ := hf.2 in lt_of_le_of_lt (hC i) ennreal.coe_lt_top⟩ + +end uniform_integrable section unif_integrable @@ -420,7 +423,7 @@ begin end /-- This lemma is less general than `measure_theory.unif_integrable_fintype` which applies to -all sequences indexed by a fintype. -/ +all sequences indexed by a finite type. -/ lemma unif_integrable_fin (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) {n : ℕ} {f : fin n → α → β} (hf : ∀ i, mem_ℒp (f i) p μ) : unif_integrable f p μ := @@ -447,16 +450,19 @@ begin end /-- A finite sequence of Lp functions is uniformly integrable. -/ -lemma unif_integrable_fintype [fintype ι] (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) +lemma unif_integrable_finite [finite ι] (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) {f : ι → α → β} (hf : ∀ i, mem_ℒp (f i) p μ) : unif_integrable f p μ := begin + obtain ⟨n, hn⟩ := finite.exists_equiv_fin ι, intros ε hε, - set g : fin (fintype.card ι) → α → β := f ∘ (fintype.equiv_fin ι).symm, + set g : fin n → α → β := f ∘ hn.some.symm with hgeq, have hg : ∀ i, mem_ℒp (g i) p μ := λ _, hf _, obtain ⟨δ, hδpos, hδ⟩ := unif_integrable_fin μ hp_one hp_top hg hε, - exact ⟨δ, hδpos, λ i s hs hμs, - equiv.symm_apply_apply (fintype.equiv_fin ι) i ▸ hδ (fintype.equiv_fin ι i) s hs hμs⟩, + refine ⟨δ, hδpos, λ i s hs hμs, _⟩, + specialize hδ (hn.some i) s hs hμs, + simp_rw [hgeq, function.comp_app, equiv.symm_apply_apply] at hδ, + assumption, end end @@ -708,14 +714,25 @@ begin end lemma unif_integrable_of (hp : 1 ≤ p) (hp' : p ≠ ∞) {f : ι → α → β} - (hf : ∀ i, strongly_measurable (f i)) + (hf : ∀ i, ae_strongly_measurable (f i) μ) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : unif_integrable f p μ := begin - refine unif_integrable_of' μ hp hp' hf (λ ε hε, _), + set g : ι → α → β := λ i, (hf i).some, + refine (unif_integrable_of' μ hp hp' (λ i, (Exists.some_spec $hf i).1) (λ ε hε, _)).ae_eq + (λ i, (Exists.some_spec $ hf i).2.symm), obtain ⟨C, hC⟩ := h ε hε, - refine ⟨max C 1, lt_max_of_lt_right one_pos, λ i, le_trans (snorm_mono (λ x, _)) (hC i)⟩, + have hCg : ∀ i, snorm ({x | C ≤ ∥g i x∥₊}.indicator (g i)) p μ ≤ ennreal.of_real ε, + { intro i, + refine le_trans (le_of_eq $ snorm_congr_ae _) (hC i), + filter_upwards [(Exists.some_spec $ hf i).2] with x hx, + by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + { rw [indicator_of_mem hfx, indicator_of_mem, hx], + rwa [mem_set_of, hx] at hfx }, + { rw [indicator_of_not_mem hfx, indicator_of_not_mem], + rwa [mem_set_of, hx] at hfx } }, + refine ⟨max C 1, lt_max_of_lt_right one_pos, λ i, le_trans (snorm_mono (λ x, _)) (hCg i)⟩, rw [norm_indicator_eq_indicator_norm, norm_indicator_eq_indicator_norm], exact indicator_le_indicator_of_subset (λ x hx, le_trans (le_max_left _ _) hx) (λ _, norm_nonneg _) _, @@ -737,34 +754,35 @@ In this section, we will develope some API for `uniform_integrable` and prove th variables {p : ℝ≥0∞} {f : ι → α → β} -lemma uniform_integrable_zero_meas [measurable_space α] (hf : ∀ i, strongly_measurable (f i)) : +lemma uniform_integrable_zero_meas [measurable_space α] : uniform_integrable f p (0 : measure α) := -⟨hf, unif_integrable_zero_meas, 0, λ i, snorm_measure_zero.le⟩ +⟨λ n, ae_strongly_measurable_zero_measure _, + unif_integrable_zero_meas, 0, λ i, snorm_measure_zero.le⟩ lemma uniform_integrable.ae_eq {g : ι → α → β} - (hf : uniform_integrable f p μ) (hg : ∀ i, strongly_measurable (g i)) (hfg : ∀ n, f n =ᵐ[μ] g n) : + (hf : uniform_integrable f p μ) (hfg : ∀ n, f n =ᵐ[μ] g n) : uniform_integrable g p μ := begin - obtain ⟨-, hunif, C, hC⟩ := hf, - refine ⟨hg, (unif_integrable_congr_ae hfg).1 hunif, C, λ i, _⟩, + obtain ⟨hfm, hunif, C, hC⟩ := hf, + refine ⟨λ i, (hfm i).congr (hfg i), (unif_integrable_congr_ae hfg).1 hunif, C, λ i, _⟩, rw ← snorm_congr_ae (hfg i), exact hC i end lemma uniform_integrable_congr_ae {g : ι → α → β} - (hf : ∀ i, strongly_measurable (f i)) (hg : ∀ i, strongly_measurable (g i)) (hfg : ∀ n, f n =ᵐ[μ] g n) : uniform_integrable f p μ ↔ uniform_integrable g p μ := -⟨λ h, h.ae_eq hg hfg, λ h, h.ae_eq hf (λ i, (hfg i).symm)⟩ +⟨λ h, h.ae_eq hfg, λ h, h.ae_eq (λ i, (hfg i).symm)⟩ /-- A finite sequence of Lp functions is uniformly integrable in the probability sense. -/ -lemma uniform_integrable_fintype [fintype ι] (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) - (hf : ∀ i, strongly_measurable (f i)) (hf' : ∀ i, mem_ℒp (f i) p μ) : +lemma uniform_integrable_finite [finite ι] (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) + (hf : ∀ i, mem_ℒp (f i) p μ) : uniform_integrable f p μ := begin - refine ⟨hf, unif_integrable_fintype μ hp_one hp_top hf', _⟩, + casesI nonempty_fintype ι, + refine ⟨λ n, (hf n).1, unif_integrable_finite μ hp_one hp_top hf, _⟩, by_cases hι : nonempty ι, - { choose ae_meas hf using hf', + { choose ae_meas hf using hf, set C := (finset.univ.image (λ i : ι, snorm (f i) p μ)).max' ⟨snorm (f hι.some) p μ, finset.mem_image.2 ⟨hι.some, finset.mem_univ _, rfl⟩⟩, refine ⟨C.to_nnreal, λ i, _⟩, @@ -779,26 +797,27 @@ end /-- A single function is uniformly integrable in the probability sense. -/ lemma uniform_integrable_subsingleton [subsingleton ι] (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) - (hf : ∀ i, strongly_measurable (f i)) (hf' : ∀ i, mem_ℒp (f i) p μ) : + (hf : ∀ i, mem_ℒp (f i) p μ) : uniform_integrable f p μ := -uniform_integrable_fintype hp_one hp_top hf hf' +uniform_integrable_finite hp_one hp_top hf /-- A constant sequence of functions is uniformly integrable in the probability sense. -/ lemma uniform_integrable_const {g : α → β} (hp : 1 ≤ p) (hp_ne_top : p ≠ ∞) - (hgm : strongly_measurable g) (hg : mem_ℒp g p μ) : + (hg : mem_ℒp g p μ) : uniform_integrable (λ n : ι, g) p μ := -⟨λ i, hgm, unif_integrable_const μ hp hp_ne_top hg, +⟨λ i, hg.1, unif_integrable_const μ hp hp_ne_top hg, ⟨(snorm g p μ).to_nnreal, λ i, le_of_eq (ennreal.coe_to_nnreal hg.2.ne).symm⟩⟩ -/-- A sequene of functions `(fₙ)` is uniformly integrable in the probability sense if for all -`ε > 0`, there exists some `C` such that `∫ x in {|fₙ| ≥ C}, fₙ x ∂μ ≤ ε` for all `n`. -/ -lemma uniform_integrable_of [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) +/-- This lemma is superceded by `uniform_integrable_of` which only requires +`ae_strongly_measurable`. -/ +lemma uniform_integrable_of' [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) (hf : ∀ i, strongly_measurable (f i)) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : uniform_integrable f p μ := begin - refine ⟨hf, unif_integrable_of μ hp hp' hf h, _⟩, + refine ⟨λ i, (hf i).ae_strongly_measurable, + unif_integrable_of μ hp hp' (λ i, (hf i).ae_strongly_measurable) h, _⟩, obtain ⟨C, hC⟩ := h 1 one_pos, refine ⟨(C * (μ univ ^ (p.to_real⁻¹)) + 1 : ℝ≥0∞).to_nnreal, λ i, _⟩, calc snorm (f i) p μ ≤ snorm ({x : α | ∥f i x∥₊ < C}.indicator (f i)) p μ + @@ -835,11 +854,35 @@ begin end end -lemma uniform_integrable.spec (hp : p ≠ 0) (hp' : p ≠ ∞) +/-- A sequene of functions `(fₙ)` is uniformly integrable in the probability sense if for all +`ε > 0`, there exists some `C` such that `∫ x in {|fₙ| ≥ C}, fₙ x ∂μ ≤ ε` for all `n`. -/ +lemma uniform_integrable_of [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) + (hf : ∀ i, ae_strongly_measurable (f i) μ) + (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, + ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : + uniform_integrable f p μ := +begin + set g : ι → α → β := λ i, (hf i).some, + have hgmeas : ∀ i, strongly_measurable (g i) := λ i, (Exists.some_spec $ hf i).1, + have hgeq : ∀ i, g i =ᵐ[μ] f i := λ i, (Exists.some_spec $ hf i).2.symm, + refine (uniform_integrable_of' hp hp' hgmeas $ λ ε hε, _).ae_eq hgeq, + obtain ⟨C, hC⟩ := h ε hε, + refine ⟨C, λ i, le_trans (le_of_eq $ snorm_congr_ae _) (hC i)⟩, + filter_upwards [(Exists.some_spec $ hf i).2] with x hx, + by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + { rw [indicator_of_mem hfx, indicator_of_mem, hx], + rwa [mem_set_of, hx] at hfx }, + { rw [indicator_of_not_mem hfx, indicator_of_not_mem], + rwa [mem_set_of, hx] at hfx } +end + +/-- This lemma is superceded by `uniform_integrable.spec` which does not require measurability. -/ +lemma uniform_integrable.spec' (hp : p ≠ 0) (hp' : p ≠ ∞) + (hf : ∀ i, strongly_measurable (f i)) (hfu : uniform_integrable f p μ) {ε : ℝ} (hε : 0 < ε) : ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := begin - obtain ⟨hf₀, hfu, M, hM⟩ := hfu, + obtain ⟨-, hfu, M, hM⟩ := hfu, obtain ⟨δ, hδpos, hδ⟩ := hfu hε, obtain ⟨C, hC⟩ : ∃ C : ℝ≥0, ∀ i, μ {x | C ≤ ∥f i x∥₊} ≤ ennreal.of_real δ, { by_contra hcon, push_neg at hcon, @@ -857,7 +900,7 @@ begin ... ≤ snorm ({x | C ≤ ∥f (ℐ C) x∥₊}.indicator (f (ℐ C))) p μ : begin refine snorm_indicator_ge_of_bdd_below hp hp' _ - (measurable_set_le measurable_const (hf₀ _).nnnorm.measurable) + (measurable_set_le measurable_const (hf _).nnnorm.measurable) (eventually_of_forall $ λ x hx, _), rwa [nnnorm_indicator_eq_indicator_nnnorm, indicator_of_mem hx], end @@ -871,13 +914,30 @@ begin (le_max_left M 1)) (lt_of_lt_of_le _ this)).ne rfl, rw [← ennreal.coe_one, ← with_top.coe_max, ← ennreal.coe_mul, ennreal.coe_lt_coe], exact lt_two_mul_self (lt_max_of_lt_right one_pos) }, - exact ⟨C, λ i, hδ i _ (measurable_set_le measurable_const (hf₀ i).nnnorm.measurable) (hC i)⟩, + exact ⟨C, λ i, hδ i _ (measurable_set_le measurable_const (hf i).nnnorm.measurable) (hC i)⟩, +end + +lemma uniform_integrable.spec (hp : p ≠ 0) (hp' : p ≠ ∞) + (hfu : uniform_integrable f p μ) {ε : ℝ} (hε : 0 < ε) : + ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := +begin + set g : ι → α → β := λ i, (hfu.1 i).some, + have hgmeas : ∀ i, strongly_measurable (g i) := λ i, (Exists.some_spec $ hfu.1 i).1, + have hgunif : uniform_integrable g p μ := hfu.ae_eq (λ i, (Exists.some_spec $ hfu.1 i).2), + obtain ⟨C, hC⟩ := hgunif.spec' hp hp' hgmeas hε, + refine ⟨C, λ i, le_trans (le_of_eq $ snorm_congr_ae _) (hC i)⟩, + filter_upwards [(Exists.some_spec $ hfu.1 i).2] with x hx, + by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + { rw [indicator_of_mem hfx, indicator_of_mem, hx], + rwa [mem_set_of, hx] at hfx }, + { rw [indicator_of_not_mem hfx, indicator_of_not_mem], + rwa [mem_set_of, hx] at hfx } end /-- The definition of uniform integrable in mathlib is equivalent to the definition commonly found in literature. -/ lemma uniform_integrable_iff [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) : - uniform_integrable f p μ ↔ (∀ i, strongly_measurable (f i)) ∧ + 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'⟩, From b278a51a17e8c11d9795028ad4df613f08523e02 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 25 Jul 2022 14:27:52 +0000 Subject: [PATCH 046/173] feat(data/sign): `left.sign_neg`, `right.sign_neg` (#15652) Add lemmas that the sign of `-a` is the negation of the sign of `a`, for two kinds of ordered additive groups (corresponding to the type classes on the `left` and `right` variants of the `neg_pos_iff` and `neg_neg_iff` lemmas used). --- src/data/sign.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/data/sign.lean b/src/data/sign.lean index e170570845907..214ca9403e458 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -202,3 +202,29 @@ def sign_hom : α →*₀ sign_type := mul_neg_of_pos_of_neg, mul_pos] } end linear_ordered_ring + +section add_group + +variables [add_group α] [preorder α] [decidable_rel ((<) : α → α → Prop)] + +lemma left.sign_neg [covariant_class α α (+) (<)] (a : α) : sign (-a) = - sign a := +begin + simp_rw [sign_apply, left.neg_pos_iff, left.neg_neg_iff], + split_ifs with h h', + { exact false.elim (lt_asymm h h') }, + { simp }, + { simp }, + { simp } +end + +lemma right.sign_neg [covariant_class α α (function.swap (+)) (<)] (a : α) : sign (-a) = - sign a := +begin + simp_rw [sign_apply, right.neg_pos_iff, right.neg_neg_iff], + split_ifs with h h', + { exact false.elim (lt_asymm h h') }, + { simp }, + { simp }, + { simp } +end + +end add_group From d087b2608a5bf0cf84b2b16e5460eb8e4b036ce0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 25 Jul 2022 14:27:54 +0000 Subject: [PATCH 047/173] feat(algebra/lie/of_associative): add `commute.lie_eq` (#15675) --- src/algebra/lie/of_associative.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/lie/of_associative.lean b/src/algebra/lie/of_associative.lean index 484d5d780aba3..030bde6d006ea 100644 --- a/src/algebra/lie/of_associative.lean +++ b/src/algebra/lie/of_associative.lean @@ -49,6 +49,10 @@ lemma lie_def (x y : A) : ⁅x, y⁆ = x*y - y*x := rfl end ring +lemma commute_iff_lie_eq {x y : A} : commute x y ↔ ⁅x, y⁆ = 0 := sub_eq_zero.symm + +lemma commute.lie_eq {x y : A} (h : commute x y) : ⁅x, y⁆ = 0 := sub_eq_zero_of_eq h + namespace lie_ring /-- An associative ring gives rise to a Lie ring by taking the bracket to be the ring commutator. -/ From 0f49e14b3cd615b9fed57044c68f5b7e3ec6e6a5 Mon Sep 17 00:00:00 2001 From: kkytola Date: Mon, 25 Jul 2022 18:14:50 +0000 Subject: [PATCH 048/173] feat(measure_theory/measure/finite_measure_weak_convergence): Add normalization of finite measures to probability measures and characterize weak convergence in terms of it. (#15528) This PR adds the definition and basic API about normalization of finite measures to probability measures (divide by the total mass, return junk when total mass vanishes). The weak convergence of finite measures is then characterized in terms of convergence of the total mass and the convergence of the probability normalized measures. This characterization allows to obtain results about the weak convergence of finite measures from the often more convenient considerations of weak convergence of probability measures (some implications of portmanteau theorem, in particular). Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- .../finite_measure_weak_convergence.lean | 390 +++++++++++++++--- 1 file changed, 344 insertions(+), 46 deletions(-) diff --git a/src/measure_theory/measure/finite_measure_weak_convergence.lean b/src/measure_theory/measure/finite_measure_weak_convergence.lean index 2cc828c7d7917..aed6f9a293539 100644 --- a/src/measure_theory/measure/finite_measure_weak_convergence.lean +++ b/src/measure_theory/measure/finite_measure_weak_convergence.lean @@ -5,6 +5,7 @@ Authors: Kalle Kytölä -/ import measure_theory.measure.measure_space import measure_theory.integral.set_integral +import measure_theory.integral.average import topology.continuous_function.bounded import topology.algebra.module.weak_dual import topology.metric_space.thickened_indicator @@ -24,8 +25,11 @@ TODOs: ## Main definitions The main definitions are the - * types `finite_measure α` and `probability_measure α` with topologies of weak convergence; - * `to_weak_dual_bcnn : finite_measure α → (weak_dual ℝ≥0 (α →ᵇ ℝ≥0))` + * types `measure_theory.finite_measure α` and `measure_theory.probability_measure α` with + the topologies of weak convergence; + * `measure_theory.finite_measure.normalize`, normalizing a finite measure to a probability measure + (returns junk for the zero measure); + * `measure_theory.finite_measure.to_weak_dual_bcnn : finite_measure α → (weak_dual ℝ≥0 (α →ᵇ ℝ≥0))` allowing to interpret a finite measure as a continuous linear functional on the space of bounded continuous nonnegative functions on `α`. This is used for the definition of the topology of weak convergence. @@ -34,15 +38,23 @@ The main definitions are the * Finite measures `μ` on `α` give rise to continuous linear functionals on the space of bounded continuous nonnegative functions on `α` via integration: - `to_weak_dual_bcnn : finite_measure α → (weak_dual ℝ≥0 (α →ᵇ ℝ≥0))`. - * `tendsto_iff_forall_lintegral_tendsto`: Convergence of finite measures and probability measures - is characterized by the convergence of integrals of all bounded continuous (nonnegative) - functions. This essentially shows that the given definition of topology corresponds to the - common textbook definition of weak convergence of measures. + `measure_theory.finite_measure.to_weak_dual_bcnn : finite_measure α → (weak_dual ℝ≥0 (α →ᵇ ℝ≥0))` + * `measure_theory.finite_measure.tendsto_iff_forall_integral_tendsto` and + `measure_theory.probability_measure.tendsto_iff_forall_integral_tendsto`: Convergence of finite + measures and probability measures is characterized by the convergence of integrals of all + bounded continuous functions. This shows that the chosen definition of topology coincides with + the common textbook definition of weak convergence of measures. + Similar characterizations by the convergence of integrals (in the `measure_theory.lintegral` + sense) of all bounded continuous nonnegative functions are + `measure_theory.finite_measure.tendsto_iff_forall_lintegral_tendsto` and + `measure_theory.probability_measure.tendsto_iff_forall_lintegral_tendsto`. + * `measure_theory.finite_measure.tendsto_normalize_iff_tendsto`: The convergence of finite + measures to a nonzero limit is characterized by the convergence of the probability-normalized + versions and of the total masses. TODO: * Portmanteau theorem: - * `finite_measure.limsup_measure_closed_le_of_tendsto` proves one implication. + * `measure_theory.finite_measure.limsup_measure_closed_le_of_tendsto` proves one implication. The current formulation assumes `pseudo_emetric_space`. The only reason is to have bounded continuous pointwise approximations to the indicator function of a closed set. Clearly for example metrizability or pseudo-emetrizability would be sufficient assumptions. The @@ -57,22 +69,25 @@ No new notation is introduced. ## Implementation notes The topology of weak convergence of finite Borel measures will be defined using a mapping from -`finite_measure α` to `weak_dual ℝ≥0 (α →ᵇ ℝ≥0)`, inheriting the topology from the latter. +`measure_theory.finite_measure α` to `weak_dual ℝ≥0 (α →ᵇ ℝ≥0)`, inheriting the topology from the +latter. -The current implementation of `finite_measure α` and `probability_measure α` is directly as -subtypes of `measure α`, and the coercion to a function is the composition `ennreal.to_nnreal` -and the coercion to function of `measure α`. Another alternative would be to use a bijection -with `vector_measure α ℝ≥0` as an intermediate step. The choice of implementation should not have -drastic downstream effects, so it can be changed later if appropriate. +The current implementation of `measure_theory.finite_measure α` and +`measure_theory.probability_measure α` is directly as subtypes of `measure_theory.measure α`, and +the coercion to a function is the composition `ennreal.to_nnreal` and the coercion to function +of `measure_theory.measure α`. Another alternative would be to use a bijection +with `measure_theory.vector_measure α ℝ≥0` as an intermediate step. The choice of implementation +should not have drastic downstream effects, so it can be changed later if appropriate. Potential advantages of using the `nnreal`-valued vector measure alternative: * The coercion to function would avoid need to compose with `ennreal.to_nnreal`, the `nnreal`-valued API could be more directly available. + Potential drawbacks of the vector measure alternative: * The coercion to function would lose monotonicity, as non-measurable sets would be defined to have measure 0. - * No integration theory directly. E.g., the topology definition requires `lintegral` w.r.t. - a coercion to `measure α` in any case. + * No integration theory directly. E.g., the topology definition requires + `measure_theory.lintegral` w.r.t. a coercion to `measure_theory.measure α` in any case. ## References @@ -155,6 +170,12 @@ begin rwa [← ennreal_mass, ennreal.coe_eq_zero], end +lemma mass_nonzero_iff (μ : finite_measure α) : μ.mass ≠ 0 ↔ μ ≠ 0 := +begin + rw not_iff_not, + exact finite_measure.mass_zero_iff μ, +end + @[ext] lemma extensionality (μ ν : finite_measure α) (h : ∀ (s : set α), measurable_set s → μ s = ν s) : μ = ν := @@ -252,6 +273,15 @@ begin exact lintegral_mono (λ x, ennreal.coe_mono (f_le_g x)), end +@[simp] lemma test_against_nn_zero (μ : finite_measure α) : μ.test_against_nn 0 = 0 := +by simpa only [zero_mul] using μ.test_against_nn_const 0 + +@[simp] lemma test_against_nn_one (μ : finite_measure α) : μ.test_against_nn 1 = μ.mass := +begin + simp only [test_against_nn, coe_one, pi.one_apply, ennreal.coe_one, lintegral_one], + refl, +end + @[simp] lemma zero.test_against_nn_apply (f : α →ᵇ ℝ≥0) : (0 : finite_measure α).test_against_nn f = 0 := by simp only [test_against_nn, coe_zero, lintegral_zero_measure, ennreal.zero_to_nnreal] @@ -330,7 +360,7 @@ begin end /-- Finite measures yield elements of the `weak_dual` of bounded continuous nonnegative -functions via `finite_measure.test_against_nn`, i.e., integration. -/ +functions via `measure_theory.finite_measure.test_against_nn`, i.e., integration. -/ def to_weak_dual_bcnn (μ : finite_measure α) : weak_dual ℝ≥0 (α →ᵇ ℝ≥0) := { to_fun := λ f, μ.test_against_nn f, @@ -344,8 +374,9 @@ def to_weak_dual_bcnn (μ : finite_measure α) : @[simp] lemma to_weak_dual_bcnn_apply (μ : finite_measure α) (f : α →ᵇ ℝ≥0) : μ.to_weak_dual_bcnn f = (∫⁻ x, f x ∂(μ : measure α)).to_nnreal := rfl -/-- The topology of weak convergence on `finite_measures α` is inherited (induced) from the weak-* -topology on `weak_dual ℝ≥0 (α →ᵇ ℝ≥0)` via the function `finite_measures.to_weak_dual_bcnn`. -/ +/-- The topology of weak convergence on `measure_theory.finite_measure α` is inherited (induced) +from the weak-* topology on `weak_dual ℝ≥0 (α →ᵇ ℝ≥0)` via the function +`measure_theory.finite_measure.to_weak_dual_bcnn`. -/ instance : topological_space (finite_measure α) := topological_space.induced to_weak_dual_bcnn infer_instance @@ -360,17 +391,68 @@ lemma continuous_test_against_nn_eval (f : α →ᵇ ℝ≥0) : (by apply (weak_bilin.eval_continuous _ _).comp to_weak_dual_bcnn_continuous : continuous ((λ φ : weak_dual ℝ≥0 (α →ᵇ ℝ≥0), φ f) ∘ to_weak_dual_bcnn)) +/-- The total mass of a finite measure depends continuously on the measure. -/ +lemma continuous_mass : continuous (λ (μ : finite_measure α), μ.mass) := +by { simp_rw ←test_against_nn_one, exact continuous_test_against_nn_eval 1, } + +/-- Convergence of finite measures implies the convergence of their total masses. -/ +lemma _root_.filter.tendsto.mass {γ : Type*} {F : filter γ} + {μs : γ → finite_measure α} {μ : finite_measure α} (h : tendsto μs F (𝓝 μ)) : + tendsto (λ i, (μs i).mass) F (𝓝 μ.mass) := +(continuous_mass.tendsto μ).comp h + lemma tendsto_iff_weak_star_tendsto {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} {μ : finite_measure α} : tendsto μs F (𝓝 μ) ↔ tendsto (λ i, (μs(i)).to_weak_dual_bcnn) F (𝓝 μ.to_weak_dual_bcnn) := inducing.tendsto_nhds_iff ⟨rfl⟩ -theorem tendsto_iff_forall_test_against_nn_tendsto {γ : Type*} {F : filter γ} - {μs : γ → finite_measure α} {μ : finite_measure α} : +theorem tendsto_iff_forall_to_weak_dual_bcnn_tendsto + {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} {μ : finite_measure α} : tendsto μs F (𝓝 μ) ↔ ∀ (f : α →ᵇ ℝ≥0), tendsto (λ i, (μs i).to_weak_dual_bcnn f) F (𝓝 (μ.to_weak_dual_bcnn f)) := by { rw [tendsto_iff_weak_star_tendsto, tendsto_iff_forall_eval_tendsto_top_dual_pairing], refl, } +theorem tendsto_iff_forall_test_against_nn_tendsto + {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} {μ : finite_measure α} : + tendsto μs F (𝓝 μ) ↔ + ∀ (f : α →ᵇ ℝ≥0), tendsto (λ i, (μs i).test_against_nn f) F (𝓝 (μ.test_against_nn f)) := +by { rw finite_measure.tendsto_iff_forall_to_weak_dual_bcnn_tendsto, refl, } + +/-- If the total masses of finite measures tend to zero, then the measures tend to +zero. This formulation concerns the associated functionals on bounded continuous +nonnegative test functions. See `finite_measure.tendsto_zero_of_tendsto_zero_mass` for +a formulation stating the weak convergence of measures. -/ +lemma tendsto_zero_test_against_nn_of_tendsto_zero_mass + {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} + (mass_lim : tendsto (λ i, (μs i).mass) F (𝓝 0)) (f : α →ᵇ ℝ≥0) : + tendsto (λ i, (μs i).test_against_nn f) F (𝓝 0) := +begin + apply tendsto_iff_dist_tendsto_zero.mpr, + have obs := λ i, (μs i).test_against_nn_lipschitz_estimate f 0, + simp_rw [test_against_nn_zero, zero_add] at obs, + simp_rw (show ∀ i, dist ((μs i).test_against_nn f) 0 = (μs i).test_against_nn f, + by simp only [dist_nndist, nnreal.nndist_zero_eq_val', eq_self_iff_true, + implies_true_iff]), + refine squeeze_zero (λ i, nnreal.coe_nonneg _) obs _, + simp_rw nnreal.coe_mul, + have lim_pair : tendsto (λ i, (⟨nndist f 0, (μs i).mass⟩ : ℝ × ℝ)) F (𝓝 (⟨nndist f 0, 0⟩)), + { refine (prod.tendsto_iff _ _).mpr ⟨tendsto_const_nhds, _⟩, + exact (nnreal.continuous_coe.tendsto 0).comp mass_lim, }, + have key := tendsto_mul.comp lim_pair, + rwa mul_zero at key, +end + +/-- If the total masses of finite measures tend to zero, then the measures tend to zero. -/ +lemma tendsto_zero_of_tendsto_zero_mass {γ : Type*} {F : filter γ} + {μs : γ → finite_measure α} (mass_lim : tendsto (λ i, (μs i).mass) F (𝓝 0)) : + tendsto μs F (𝓝 0) := +begin + rw tendsto_iff_forall_test_against_nn_tendsto, + intro f, + convert tendsto_zero_test_against_nn_of_tendsto_zero_mass mass_lim f, + rw [zero.test_against_nn_apply], +end + /-- A characterization of weak convergence in terms of integrals of bounded continuous nonnegative functions. -/ theorem tendsto_iff_forall_lintegral_tendsto {γ : Type*} {F : filter γ} @@ -379,7 +461,7 @@ theorem tendsto_iff_forall_lintegral_tendsto {γ : Type*} {F : filter γ} ∀ (f : α →ᵇ ℝ≥0), tendsto (λ i, (∫⁻ x, (f x) ∂(μs(i) : measure α))) F (𝓝 ((∫⁻ x, (f x) ∂(μ : measure α)))) := begin - rw tendsto_iff_forall_test_against_nn_tendsto, + rw tendsto_iff_forall_to_weak_dual_bcnn_tendsto, simp_rw [to_weak_dual_bcnn_apply _ _, ←test_against_nn_coe_eq, ennreal.tendsto_coe, ennreal.to_nnreal_coe], end @@ -401,7 +483,8 @@ This formulation assumes: * the functions tend to a limit along a countably generated filter; * the limit is in the almost everywhere sense; * boundedness holds almost everywhere; - * integration is `lintegral`, i.e., the functions and their integrals are `ℝ≥0∞`-valued. + * integration is `measure_theory.lintegral`, i.e., the functions and their integrals are + `ℝ≥0∞`-valued. -/ lemma tendsto_lintegral_nn_filter_of_le_const {ι : Type*} {L : filter ι} [L.is_countably_generated] (μ : measure α) [is_finite_measure μ] {fs : ι → (α →ᵇ ℝ≥0)} {c : ℝ≥0} @@ -418,10 +501,11 @@ end /-- A bounded convergence theorem for a finite measure: If a sequence of bounded continuous non-negative functions are uniformly bounded by a constant -and tend pointwise to a limit, then their integrals (`lintegral`) against the finite measure tend -to the integral of the limit. +and tend pointwise to a limit, then their integrals (`measure_theory.lintegral`) against the finite +measure tend to the integral of the limit. -A related result with more general assumptions is `tendsto_lintegral_nn_filter_of_le_const`. +A related result with more general assumptions is +`measure_theory.finite_measure.tendsto_lintegral_nn_filter_of_le_const`. -/ lemma tendsto_lintegral_nn_of_le_const (μ : finite_measure α) {fs : ℕ → (α →ᵇ ℝ≥0)} {c : ℝ≥0} (fs_le_const : ∀ n a, fs n a ≤ c) {f : α → ℝ≥0} @@ -437,9 +521,11 @@ This formulation assumes: * the functions tend to a limit along a countably generated filter; * the limit is in the almost everywhere sense; * boundedness holds almost everywhere; - * integration is the pairing against non-negative continuous test functions (`test_against_nn`). + * integration is the pairing against non-negative continuous test functions + (`measure_theory.finite_measure.test_against_nn`). -A related result using `lintegral` for integration is `tendsto_lintegral_nn_filter_of_le_const`. +A related result using `measure_theory.lintegral` for integration is +`measure_theory.finite_measure.tendsto_lintegral_nn_filter_of_le_const`. -/ lemma tendsto_test_against_nn_filter_of_le_const {ι : Type*} {L : filter ι} [L.is_countably_generated] {μ : finite_measure α} {fs : ι → (α →ᵇ ℝ≥0)} {c : ℝ≥0} @@ -453,13 +539,15 @@ begin end /-- A bounded convergence theorem for a finite measure: -If a sequence of bounded continuous non-negative functions are uniformly bounded by a constant -and tend pointwise to a limit, then their integrals (`test_against_nn`) against the finite measure -tend to the integral of the limit. +If a sequence of bounded continuous non-negative functions are uniformly bounded by a constant and +tend pointwise to a limit, then their integrals (`measure_theory.finite_measure.test_against_nn`) +against the finite measure tend to the integral of the limit. Related results: - * `tendsto_test_against_nn_filter_of_le_const`: more general assumptions - * `tendsto_lintegral_nn_of_le_const`: using `lintegral` for integration. + * `measure_theory.finite_measure.tendsto_test_against_nn_filter_of_le_const`: + more general assumptions + * `measure_theory.finite_measure.tendsto_lintegral_nn_of_le_const`: + using `measure_theory.lintegral` for integration. -/ lemma tendsto_test_against_nn_of_le_const {μ : finite_measure α} {fs : ℕ → (α →ᵇ ℝ≥0)} {c : ℝ≥0} (fs_le_const : ∀ n a, fs n a ≤ c) {f : α →ᵇ ℝ≥0} @@ -585,13 +673,14 @@ end finite_measure -- namespace section probability_measure /-! ### Probability measures -In this section we define the `Type*` of probability measures on a measurable space `α`, denoted by -`probability_measure α`. TODO: Probability measures form a convex space. +In this section we define the type of probability measures on a measurable space `α`, denoted by +`measure_theory.probability_measure α`. TODO: Probability measures form a convex space. If `α` is moreover a topological space and the sigma algebra on `α` is finer than the Borel sigma -algebra (i.e. `[opens_measurable_space α]`), then `probability_measure α` is equipped with the -topology of weak convergence of measures. Since every probability measure is a finite measure, this -is implemented as the induced topology from the coercion `probability_measure.to_finite_measure`. +algebra (i.e. `[opens_measurable_space α]`), then `measure_theory.probability_measure α` is +equipped with the topology of weak convergence of measures. Since every probability measure is a +finite measure, this is implemented as the induced topology from the coercion +`measure_theory.probability_measure.to_finite_measure`. -/ /-- Probability measures are defined as the subtype of measures that have the property of being @@ -650,15 +739,24 @@ end @[simp] lemma mass_to_finite_measure (μ : probability_measure α) : μ.to_finite_measure.mass = 1 := μ.coe_fn_univ +lemma to_finite_measure_nonzero (μ : probability_measure α) : + μ.to_finite_measure ≠ 0 := +begin + intro maybe_zero, + have mass_zero := (finite_measure.mass_zero_iff _).mpr maybe_zero, + rw μ.mass_to_finite_measure at mass_zero, + exact one_ne_zero mass_zero, +end + variables [topological_space α] [opens_measurable_space α] lemma test_against_nn_lipschitz (μ : probability_measure α) : lipschitz_with 1 (λ (f : α →ᵇ ℝ≥0), μ.to_finite_measure.test_against_nn f) := μ.mass_to_finite_measure ▸ μ.to_finite_measure.test_against_nn_lipschitz -/-- The topology of weak convergence on `probability_measures α`. This is inherited (induced) from -the weak-* topology on `weak_dual ℝ≥0 (α →ᵇ ℝ≥0)` via the function -`probability_measures.to_weak_dual_bcnn`. -/ +/-- The topology of weak convergence on `measure_theory.probability_measure α`. This is inherited +(induced) from the topology of weak convergence of finite measures via the inclusion +`measure_theory.probability_measure.to_finite_measure`. -/ instance : topological_space (probability_measure α) := topological_space.induced to_finite_measure infer_instance @@ -667,7 +765,7 @@ lemma to_finite_measure_continuous : continuous_induced_dom /-- Probability measures yield elements of the `weak_dual` of bounded continuous nonnegative -functions via `finite_measure.test_against_nn`, i.e., integration. -/ +functions via `measure_theory.finite_measure.test_against_nn`, i.e., integration. -/ def to_weak_dual_bcnn : probability_measure α → weak_dual ℝ≥0 (α →ᵇ ℝ≥0) := finite_measure.to_weak_dual_bcnn ∘ to_finite_measure @@ -730,13 +828,212 @@ end probability_measure -- namespace end probability_measure -- section +section normalize_finite_measure +/-! ### Normalization of finite measures to probability measures + +This section is about normalizing finite measures to probability measures. + +The weak convergence of finite measures to nonzero limit measures is characterized by +the convergence of the total mass and the convergence of the normalized probability +measures. +-/ + +namespace finite_measure + +variables {α : Type*} [nonempty α] {m0 : measurable_space α} (μ : finite_measure α) + +/-- Normalize a finite measure so that it becomes a probability measure, i.e., divide by the +total mass. -/ +def normalize : probability_measure α := +if zero : μ.mass = 0 then ⟨measure.dirac ‹nonempty α›.some, measure.dirac.is_probability_measure⟩ + else { val := (μ.mass)⁻¹ • μ, + property := begin + refine ⟨_⟩, + simp only [mass, measure.coe_nnreal_smul_apply, + ←ennreal_coe_fn_eq_coe_fn_to_measure μ univ], + norm_cast, + exact inv_mul_cancel zero, + end } + +@[simp] lemma self_eq_mass_mul_normalize (s : set α) : μ s = μ.mass * μ.normalize s := +begin + by_cases μ = 0, + { rw h, + simp only [zero.mass, coe_fn_zero, pi.zero_apply, zero_mul], }, + have mass_nonzero : μ.mass ≠ 0, by rwa μ.mass_nonzero_iff, + simp only [(show μ ≠ 0, from h), mass_nonzero, normalize, not_false_iff, dif_neg], + change μ s = μ.mass * ((μ.mass)⁻¹ • μ) s, + rw coe_fn_smul_apply, + simp only [mass_nonzero, algebra.id.smul_eq_mul, mul_inv_cancel_left₀, ne.def, not_false_iff], +end + +lemma self_eq_mass_smul_normalize : μ = μ.mass • μ.normalize.to_finite_measure := +begin + ext s s_mble, + rw [μ.self_eq_mass_mul_normalize s, coe_fn_smul_apply], + refl, +end + +lemma normalize_eq_of_nonzero (nonzero : μ ≠ 0) (s : set α) : + μ.normalize s = (μ.mass)⁻¹ * (μ s) := +by simp only [μ.self_eq_mass_mul_normalize, μ.mass_nonzero_iff.mpr nonzero, + inv_mul_cancel_left₀, ne.def, not_false_iff] + +lemma normalize_eq_inv_mass_smul_of_nonzero (nonzero : μ ≠ 0) : + μ.normalize.to_finite_measure = (μ.mass)⁻¹ • μ := +begin + nth_rewrite 2 μ.self_eq_mass_smul_normalize, + rw ← smul_assoc, + simp only [μ.mass_nonzero_iff.mpr nonzero, algebra.id.smul_eq_mul, + inv_mul_cancel, ne.def, not_false_iff, one_smul], +end + +lemma coe_normalize_eq_of_nonzero (nonzero : μ ≠ 0) : (μ.normalize : measure α) = (μ.mass)⁻¹ • μ := +begin + ext1 s s_mble, + simp only [← μ.normalize.ennreal_coe_fn_eq_coe_fn_to_measure s, + μ.normalize_eq_of_nonzero nonzero s, ennreal.coe_mul, + ennreal_coe_fn_eq_coe_fn_to_measure, measure.coe_nnreal_smul_apply], +end + +@[simp] lemma _root_.probability_measure.to_finite_measure_normalize_eq_self + {m0 : measurable_space α} (μ : probability_measure α) : + μ.to_finite_measure.normalize = μ := +begin + ext s s_mble, + rw μ.to_finite_measure.normalize_eq_of_nonzero μ.to_finite_measure_nonzero s, + simp only [probability_measure.mass_to_finite_measure, inv_one, one_mul], + refl, +end + +/-- Averaging with respect to a finite measure is the same as integraing against +`measure_theory.finite_measure.normalize`. -/ +lemma average_eq_integral_normalize + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + (nonzero : μ ≠ 0) (f : α → E) : + average (μ : measure α) f = ∫ x, f x ∂(μ.normalize : measure α) := +begin + rw [μ.coe_normalize_eq_of_nonzero nonzero, average], + congr, + simp only [ring_hom.to_fun_eq_coe, ennreal.coe_of_nnreal_hom, + ennreal.coe_inv (μ.mass_nonzero_iff.mpr nonzero), ennreal_mass], +end + +variables [topological_space α] + +lemma test_against_nn_eq_mass_mul (f : α →ᵇ ℝ≥0) : + μ.test_against_nn f = μ.mass * μ.normalize.to_finite_measure.test_against_nn f := +begin + nth_rewrite 0 μ.self_eq_mass_smul_normalize, + rw μ.normalize.to_finite_measure.smul_test_against_nn_apply μ.mass f, + refl, +end + +lemma normalize_test_against_nn (nonzero : μ ≠ 0) (f : α →ᵇ ℝ≥0) : + μ.normalize.to_finite_measure.test_against_nn f = (μ.mass)⁻¹ * μ.test_against_nn f := +by simp [μ.test_against_nn_eq_mass_mul, μ.mass_nonzero_iff.mpr nonzero] + +variables [opens_measurable_space α] + +variables {μ} + +lemma tendsto_test_against_nn_of_tendsto_normalize_test_against_nn_of_tendsto_mass + {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} + (μs_lim : tendsto (λ i, (μs i).normalize) F (𝓝 μ.normalize)) + (mass_lim : tendsto (λ i, (μs i).mass) F (𝓝 μ.mass)) (f : α →ᵇ ℝ≥0) : + tendsto (λ i, (μs i).test_against_nn f) F (𝓝 (μ.test_against_nn f)) := +begin + by_cases h_mass : μ.mass = 0, + { simp only [μ.mass_zero_iff.mp h_mass, zero.test_against_nn_apply, + zero.mass, eq_self_iff_true] at *, + exact tendsto_zero_test_against_nn_of_tendsto_zero_mass mass_lim f, }, + simp_rw [(λ i, (μs i).test_against_nn_eq_mass_mul f), μ.test_against_nn_eq_mass_mul f], + rw probability_measure.tendsto_nhds_iff_to_finite_measures_tendsto_nhds at μs_lim, + rw tendsto_iff_forall_test_against_nn_tendsto at μs_lim, + have lim_pair : tendsto + (λ i, (⟨(μs i).mass, (μs i).normalize.to_finite_measure.test_against_nn f⟩ : ℝ≥0 × ℝ≥0)) + F (𝓝 (⟨μ.mass, μ.normalize.to_finite_measure.test_against_nn f⟩)), + from (prod.tendsto_iff _ _).mpr ⟨mass_lim, μs_lim f⟩, + exact tendsto_mul.comp lim_pair, +end + +lemma tendsto_normalize_test_against_nn_of_tendsto {γ : Type*} {F : filter γ} + {μs : γ → finite_measure α} (μs_lim : tendsto μs F (𝓝 μ)) (nonzero : μ ≠ 0) (f : α →ᵇ ℝ≥0) : + tendsto (λ i, (μs i).normalize.to_finite_measure.test_against_nn f) F + (𝓝 (μ.normalize.to_finite_measure.test_against_nn f)) := +begin + have lim_mass := μs_lim.mass, + have aux : {(0 : ℝ≥0)}ᶜ ∈ 𝓝 (μ.mass), + from is_open_compl_singleton.mem_nhds (μ.mass_nonzero_iff.mpr nonzero), + have eventually_nonzero : ∀ᶠ i in F, μs i ≠ 0, + { simp_rw ← mass_nonzero_iff, + exact lim_mass aux, }, + have eve : ∀ᶠ i in F, + (μs i).normalize.to_finite_measure.test_against_nn f + = ((μs i).mass)⁻¹ * (μs i).test_against_nn f, + { filter_upwards [eventually_iff.mp eventually_nonzero], + intros i hi, + apply normalize_test_against_nn _ hi, }, + simp_rw [tendsto_congr' eve, μ.normalize_test_against_nn nonzero], + have lim_pair : tendsto + (λ i, (⟨((μs i).mass)⁻¹, (μs i).test_against_nn f⟩ : ℝ≥0 × ℝ≥0)) + F (𝓝 (⟨(μ.mass)⁻¹, μ.test_against_nn f⟩)), + { refine (prod.tendsto_iff _ _).mpr ⟨_, _⟩, + { exact (continuous_on_inv₀.continuous_at aux).tendsto.comp lim_mass, }, + { exact tendsto_iff_forall_test_against_nn_tendsto.mp μs_lim f, }, }, + exact tendsto_mul.comp lim_pair, +end + +/-- If the normalized versions of finite measures converge weakly and their total masses +also converge, then the finite measures themselves converge weakly. -/ +lemma tendsto_of_tendsto_normalize_test_against_nn_of_tendsto_mass + {γ : Type*} {F : filter γ} {μs : γ → finite_measure α} + (μs_lim : tendsto (λ i, (μs i).normalize) F (𝓝 μ.normalize)) + (mass_lim : tendsto (λ i, (μs i).mass) F (𝓝 μ.mass)) : + tendsto μs F (𝓝 μ) := +begin + rw tendsto_iff_forall_test_against_nn_tendsto, + exact λ f, tendsto_test_against_nn_of_tendsto_normalize_test_against_nn_of_tendsto_mass + μs_lim mass_lim f, +end + +/-- If finite measures themselves converge weakly to a nonzero limit measure, then their +normalized versions also converge weakly. -/ +lemma tendsto_normalize_of_tendsto {γ : Type*} {F : filter γ} + {μs : γ → finite_measure α} (μs_lim : tendsto μs F (𝓝 μ)) (nonzero : μ ≠ 0) : + tendsto (λ i, (μs i).normalize) F (𝓝 (μ.normalize)) := +begin + rw [probability_measure.tendsto_nhds_iff_to_finite_measures_tendsto_nhds, + tendsto_iff_forall_test_against_nn_tendsto], + exact λ f, tendsto_normalize_test_against_nn_of_tendsto μs_lim nonzero f, +end + +/-- The weak convergence of finite measures to a nonzero limit can be characterized by the weak +convergence of both their normalized versions (probability measures) and their total masses. -/ +theorem tendsto_normalize_iff_tendsto {γ : Type*} {F : filter γ} + {μs : γ → finite_measure α} (nonzero : μ ≠ 0) : + tendsto (λ i, (μs i).normalize) F (𝓝 (μ.normalize)) ∧ tendsto (λ i, (μs i).mass) F (𝓝 (μ.mass)) + ↔ tendsto μs F (𝓝 μ) := +begin + split, + { rintros ⟨normalized_lim, mass_lim⟩, + exact tendsto_of_tendsto_normalize_test_against_nn_of_tendsto_mass normalized_lim mass_lim, }, + { intro μs_lim, + refine ⟨tendsto_normalize_of_tendsto μs_lim nonzero, μs_lim.mass⟩, }, +end + +end finite_measure --namespace + +end normalize_finite_measure -- section + section convergence_implies_limsup_closed_le /-! ### Portmanteau implication: weak convergence implies a limsup condition for closed sets In this section we prove, under the assumption that the underlying topological space `α` is -pseudo-emetrizable, that the weak convergence of measures on `finite_measure α` implies that for -any closed set `F` in `α` the limsup of the measures of `F` is at most the limit measure of `F`. -This is one implication of the portmanteau theorem characterizing weak convergence of measures. +pseudo-emetrizable, that the weak convergence of measures on `measure_theory.finite_measure α` +implies that for any closed set `F` in `α` the limsup of the measures of `F` is at most the +limit measure of `F`. This is one implication of the portmanteau theorem characterizing weak +convergence of measures. -/ variables {α : Type*} [measurable_space α] @@ -768,7 +1065,8 @@ end the functions are uniformly bounded, then their integrals against a finite measure tend to the measure of the set. -A similar result with more general assumptions is `measure_of_cont_bdd_of_tendsto_filter_indicator`. +A similar result with more general assumptions is +`measure_theory.measure_of_cont_bdd_of_tendsto_filter_indicator`. -/ lemma measure_of_cont_bdd_of_tendsto_indicator [topological_space α] [opens_measurable_space α] From db6168d925262992aa5171d0dc3f12b500e52e03 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Jul 2022 19:35:23 +0000 Subject: [PATCH 049/173] feat(data/set/countable): add `iff` versions of some lemmas (#15671) --- src/data/set/countable.lean | 35 +++++++++---------- .../constructions/borel_space.lean | 2 +- 2 files changed, 17 insertions(+), 20 deletions(-) diff --git a/src/data/set/countable.lean b/src/data/set/countable.lean index fd84f0ff87cc4..e6a26e0cba795 100644 --- a/src/data/set/countable.lean +++ b/src/data/set/countable.lean @@ -158,30 +158,27 @@ lemma countable_Union {t : ι → set α} [countable ι] (ht : ∀ i, (t i).coun (⋃a, t a).countable := by { haveI := λ a, (ht a).to_subtype, rw Union_eq_range_psigma, apply countable_range } -lemma countable.bUnion - {s : set α} {t : Π x ∈ s, set β} (hs : s.countable) (ht : ∀a∈s, (t a ‹_›).countable) : - (⋃a∈s, t a ‹_›).countable := -begin - rw bUnion_eq_Union, - haveI := hs.to_subtype, - exact countable_Union (by simpa using ht) -end +@[simp] lemma countable_Union_iff [countable ι] {t : ι → set α} : + (⋃ i, t i).countable ↔ ∀ i, (t i).countable := +⟨λ h i, h.mono $ subset_Union _ _, countable_Union⟩ -lemma countable.sUnion {s : set (set α)} (hs : s.countable) (h : ∀a∈s, (a : _).countable) : - (⋃₀ s).countable := -by rw sUnion_eq_bUnion; exact hs.bUnion h +lemma countable.bUnion_iff {s : set α} {t : Π a ∈ s, set β} (hs : s.countable) : + (⋃ a ∈ s, t a ‹_›).countable ↔ ∀ a ∈ s, (t a ‹_›).countable := +by { haveI := hs.to_subtype, rw [bUnion_eq_Union, countable_Union_iff, set_coe.forall'] } -lemma countable_Union_Prop {p : Prop} {t : p → set β} (ht : ∀h:p, (t h).countable) : - (⋃h:p, t h).countable := -by by_cases p; simp [h, ht] +lemma countable.sUnion_iff {s : set (set α)} (hs : s.countable) : + (⋃₀ s).countable ↔ ∀ a ∈ s, (a : _).countable := +by rw [sUnion_eq_bUnion, hs.bUnion_iff] -lemma countable.union - {s₁ s₂ : set α} (h₁ : s₁.countable) (h₂ : s₂.countable) : (s₁ ∪ s₂).countable := -by rw union_eq_Union; exact -countable_Union (bool.forall_bool.2 ⟨h₂, h₁⟩) +alias countable.bUnion_iff ↔ _ countable.bUnion +alias countable.sUnion_iff ↔ _ countable.sUnion @[simp] lemma countable_union {s t : set α} : (s ∪ t).countable ↔ s.countable ∧ t.countable := -⟨λ h, ⟨h.mono (subset_union_left s t), h.mono (subset_union_right _ _)⟩, λ h, h.1.union h.2⟩ +by simp [union_eq_Union, and.comm] + +lemma countable.union {s t : set α} (hs : s.countable) (ht : t.countable) : + (s ∪ t).countable := +countable_union.2 ⟨hs, ht⟩ @[simp] lemma countable_insert {s : set α} {a : α} : (insert a s).countable ↔ s.countable := by simp only [insert_eq, countable_union, countable_singleton, true_and] diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 4f172c2639b7d..1f9cb59f3c2a4 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -634,7 +634,7 @@ begin rcases exists_countable_dense_bot_top α with ⟨s, hsc, hsd, hsb, hst⟩, have : (⋃ (l ∈ s) (u ∈ s) (h : l < u), {Ico l u} : set (set α)).countable, from hsc.bUnion (λ l hl, hsc.bUnion - (λ u hu, countable_Union_Prop $ λ _, countable_singleton _)), + (λ u hu, countable_Union $ λ _, countable_singleton _)), simp only [← set_of_eq_eq_singleton, ← set_of_exists] at this, refine measure.ext_of_generate_from_of_cover_subset (borel_space.measurable_eq.trans (borel_eq_generate_from_Ico α)) From 1dd7b30b5a10f4b7f73a5acd5dab30d947784763 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 25 Jul 2022 21:48:30 +0000 Subject: [PATCH 050/173] chore(tactic/fix_reflect_string): delete file (#15537) The file `tactic/fix_reflect_string` has a header from @gebner explaining that it was written as a workaround for the Lean issue [144](https://github.com/leanprover-community/lean/issues/144): ```quote The default `has_reflect string` instance in Lean only work for strings up to few thousand characters. Anything larger than that will trigger a stack overflow because the string is represented as a very deeply nested expression ``` This issue was fixed in [185](https://github.com/leanprover-community/lean/pull/185) and at least one file which imports it, `tactic/doc_commands`, no longer uses its contents. Can we delete the file forever now? Or should we keep it because the workaround might be useful in the future? Co-authored-by: Rob Lewis --- src/meta/expr.lean | 6 +++++ src/tactic/doc_commands.lean | 1 - src/tactic/fix_reflect_string.lean | 42 ------------------------------ 3 files changed, 6 insertions(+), 43 deletions(-) delete mode 100644 src/tactic/fix_reflect_string.lean diff --git a/src/meta/expr.lean b/src/meta/expr.lean index 7c2e71ea37dd4..1dac4a859d036 100644 --- a/src/meta/expr.lean +++ b/src/meta/expr.lean @@ -1218,3 +1218,9 @@ end declaration meta instance pexpr.decidable_eq {elab} : decidable_eq (expr elab) := unchecked_cast expr.has_decidable_eq + +section +local attribute [semireducible] reflected +meta instance {α} [has_reflect α] : has_reflect (thunk α) | a := +expr.lam `x binder_info.default (reflect unit) (reflect $ a ()) +end diff --git a/src/tactic/doc_commands.lean b/src/tactic/doc_commands.lean index 721a9e477b84b..e548836ee6306 100644 --- a/src/tactic/doc_commands.lean +++ b/src/tactic/doc_commands.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ -import tactic.fix_reflect_string /-! # Documentation commands diff --git a/src/tactic/fix_reflect_string.lean b/src/tactic/fix_reflect_string.lean deleted file mode 100644 index 02b4a038288ae..0000000000000 --- a/src/tactic/fix_reflect_string.lean +++ /dev/null @@ -1,42 +0,0 @@ -/- -Copyright (c) 2020 Gabriel Ebner. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner --/ - -/-! - -# Workaround for stack overflows with `has_reflect string` - -The default `has_reflect string` instance in Lean only work for strings up to -few thousand characters. Anything larger than that will trigger a stack overflow because -the string is represented as a very deeply nested expression: -https://github.com/leanprover-community/lean/issues/144 - -This file adds a higher-priority instance for `has_reflect string`, which -behaves exactly the same for small strings (up to 256 characters). Larger -strings are carefully converted into a call to `string.join`. - --/ - -/-- -Splits a string into chunks of at most `size` characters. --/ -meta def string.to_chunks (size : ℕ) : string → opt_param (list string) [] → list string | s acc := -if s.length ≤ size then s :: acc else -string.to_chunks (s.popn_back size) (s.backn size :: acc) - -section -local attribute [semireducible] reflected -meta instance {α} [has_reflect α] : has_reflect (thunk α) | a := -expr.lam `x binder_info.default (reflect unit) (reflect $ a ()) -end - -@[priority 2000] -meta instance : has_reflect string | s := -let chunk_size := 256 in -if s.length ≤ chunk_size then reflect s else -have ts : list (thunk string), from (s.to_chunks chunk_size).map (λ s _, s), -have h : s = string.join (ts.map (λ t, t ())), from undefined, -suffices reflected _ (string.join $ ts.map (λ t, t ())), by rwa h, -`(string.join $ list.map _ _) From e2c2654489de1dc730261b948a2074b92da2c33c Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Tue, 26 Jul 2022 02:21:15 +0000 Subject: [PATCH 051/173] feat(big_operators/fin): sum over elements of vector equal to `a` equals count `a` (#15360) The sum over `i : fin n` of 1 for vector elements equal to some `a` is just count applied to `a`. Co-authored-by: Eric Wieser --- src/data/fintype/fin.lean | 27 +++++++++++++++++++++++++-- 1 file changed, 25 insertions(+), 2 deletions(-) diff --git a/src/data/fintype/fin.lean b/src/data/fintype/fin.lean index ba0a641d266d1..ab9824e81ff66 100644 --- a/src/data/fintype/fin.lean +++ b/src/data/fintype/fin.lean @@ -18,7 +18,9 @@ open fintype namespace fin -@[simp] lemma Ioi_zero_eq_map {n : ℕ} : +variables {α β : Type*} {n : ℕ} + +@[simp] lemma Ioi_zero_eq_map : Ioi (0 : fin n.succ) = univ.map (fin.succ_embedding _).to_embedding := begin ext i, @@ -31,7 +33,7 @@ begin exact succ_pos _ }, end -@[simp] lemma Ioi_succ {n : ℕ} (i : fin n) : +@[simp] lemma Ioi_succ (i : fin n) : Ioi i.succ = (Ioi i).map (fin.succ_embedding _).to_embedding := begin ext i, @@ -45,4 +47,25 @@ begin { rintro ⟨i, hi, rfl⟩, simpa }, 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 + rw [fin.univ_succ, filter_cons, card_disj_union, map_filter, card_map], + split_ifs; simp, +end + +lemma card_filter_univ_succ (p : fin (n + 1) → Prop) [decidable_pred p] : + (univ.filter p).card = + if p 0 then (univ.filter (p ∘ fin.succ)).card + 1 else (univ.filter (p ∘ fin.succ)).card := +(card_filter_univ_succ' p).trans (by split_ifs; simp [add_comm 1]) + +lemma card_filter_univ_eq_vector_nth_eq_count [decidable_eq α] (a : α) (v : vector α n) : + (univ.filter $ λ i, a = v.nth i).card = v.to_list.count a := +begin + induction v using vector.induction_on with n x xs hxs, + { simp }, + { simp_rw [card_filter_univ_succ', vector.nth_cons_zero, vector.to_list_cons, + function.comp, vector.nth_cons_succ, hxs, list.count_cons', add_comm (ite (a = x) 1 0)] } +end + end fin From b8b642e0620d1a383bf4116a0bf8551606ab1169 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Tue, 26 Jul 2022 04:02:42 +0000 Subject: [PATCH 052/173] chore(scripts): update nolints.txt (#15685) 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 e781488f41d9a..80821b398ffc3 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -893,10 +893,6 @@ apply_nolint is_countably_spanning.pi fintype_finite apply_nolint measure_theory.measure.prod_sum fintype_finite apply_nolint measure_theory.measure.sum_prod fintype_finite --- measure_theory/function/uniform_integrable.lean -apply_nolint measure_theory.unif_integrable_fintype fintype_finite -apply_nolint measure_theory.uniform_integrable_fintype fintype_finite - -- measure_theory/group/prod.lean apply_nolint measure_theory.map_prod_inv_mul_eq to_additive_doc apply_nolint measure_theory.map_prod_inv_mul_eq_swap to_additive_doc From a96672db8df9de9a5b4365d5c2c41f79b994960d Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Tue, 26 Jul 2022 04:56:22 +0000 Subject: [PATCH 053/173] doc(tactic/compute_degree): fix tactic tags, add examples to docs (#15680) Co-authored-by: Rob Lewis --- src/tactic/compute_degree.lean | 25 +++++++++++++++++++++++-- 1 file changed, 23 insertions(+), 2 deletions(-) diff --git a/src/tactic/compute_degree.lean b/src/tactic/compute_degree.lean index 9c17f06052ab8..aaf76eaa10564 100644 --- a/src/tactic/compute_degree.lean +++ b/src/tactic/compute_degree.lean @@ -139,7 +139,28 @@ open compute_degree polynomial where `f : R[X]` and `d : ℕ` or `d : with_bot ℕ`. If the given degree `d` is smaller than the one that the tactic computes, -then the tactic suggests the degree that it computed. -/ +then the tactic suggests the degree that it computed. + +Examples: + +```lean +open polynomial +open_locale polynomial + +variables {R : Type*} [semiring R] {a b c d e : R} + +example {F} [ring F] {a : F} {n : ℕ} (h : n ≤ 10) : + nat_degree (X ^ n + C a * X ^ 10 : F[X]) ≤ 10 := +by compute_degree_le + +example : nat_degree (7 * X : R[X]) ≤ 1 := +by compute_degree_le + +example {p : R[X]} {n : ℕ} {p0 : p.nat_degree = 0} : + (p ^ n).nat_degree ≤ 0 := +by compute_degree_le +``` +-/ meta def compute_degree_le : tactic unit := do t ← target, try $ refine ``(degree_le_nat_degree.trans (with_bot.coe_le_coe.mpr _)), @@ -159,7 +180,7 @@ add_tactic_doc { name := "compute_degree_le", category := doc_category.tactic, decl_names := [`tactic.interactive.compute_degree_le], - tags := ["arithmetic, finishing"] } + tags := ["arithmetic", "finishing"] } end interactive From ec474f4bd09520968679201d86e3775ae8515c89 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Tue, 26 Jul 2022 05:38:28 +0000 Subject: [PATCH 054/173] feat(algebra/jordan): Introduce Jordan rings (#11073) Introduces linear Jordan rings; linearises the Jordan axiom; shows that every associative ring has a symmertised product with respect to which it is a commutative Jordan ring. Linearising the Jordan axiom is an important step towards showing that (commutative) Jordan algebras are power associative. Co-authored-by: Christopher Hoskin --- docs/references.bib | 111 +++++++++++++++++ src/algebra/jordan/basic.lean | 223 ++++++++++++++++++++++++++++++++++ 2 files changed, 334 insertions(+) create mode 100644 src/algebra/jordan/basic.lean diff --git a/docs/references.bib b/docs/references.bib index 5ba5771623542..73b46727a3e5e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -43,6 +43,18 @@ @Article{ alfseneffros1972 zbl = {0248.46019} } +@Book{ alfsenshultz2003, + author = {Erik M. {Alfsen} and Frederic W. {Shultz}}, + title = {{Geometry of state spaces of operator algebras}}, + isbn = {0-8176-4319-2}, + pages = {xiii + 467}, + year = {2003}, + publisher = {Boston, MA: Birkh\"auser}, + language = {English}, + msc2010 = {46-02 46L05 46L10 46L70 46L30 17C65}, + zbl = {1042.46001} +} + @Book{ aluffi2016, title = {Algebra: Chapter 0}, author = {Aluffi, Paolo}, @@ -310,6 +322,25 @@ @Book{ bourbaki1987 url = {https://doi.org/10.1007/978-3-642-61715-7} } +@Book{ cabreragarciarodriguezpalacios2014, + author = {Miguel {Cabrera Garc\'{\i}a} and \'Angel {Rodr\'{\i}guez + Palacios}}, + title = {{Non-associative normed algebras. Volume 1. The + Vidav-Palmer and Gelfand-Naimark theorems}}, + fjournal = {{Encyclopedia of Mathematics and Its Applications}}, + journal = {{Encycl. Math. Appl.}}, + issn = {0953-4806}, + volume = {154}, + isbn = {978-1-107-04306-0; 978-1-107-33776-3}, + pages = {xxii + 712}, + year = {2014}, + publisher = {Cambridge: Cambridge University Press}, + language = {English}, + doi = {10.1017/CBO9781107337763}, + msc2010 = {46-02 17-02 46H70 46K70 46L70 17A15 17A80 17C65}, + zbl = {1322.46003} +} + @Article{ cadiou1972, title = {Recursive definitions of partial functions and their computations}, @@ -388,6 +419,22 @@ @InProceedings{ Chou1994 isbn = {978-3-540-48803-3} } +@Book{ chu2012, + author = {Cho-Ho {Chu}}, + title = {{Jordan structures in geometry and analysis}}, + fjournal = {{Cambridge Tracts in Mathematics}}, + journal = {{Camb. Tracts Math.}}, + issn = {0950-6284}, + volume = {190}, + isbn = {978-1-107-01617-0}, + pages = {x + 261}, + year = {2012}, + publisher = {Cambridge: Cambridge University Press}, + language = {English}, + msc2010 = {17-02 17C65 17C37 46H70 53C35 46K70 32M15}, + zbl = {1238.17001} +} + @InProceedings{ CL21, author = {Commelin, Johan and Lewis, Robert Y.}, title = {Formalizing the Ring of Witt Vectors}, @@ -647,6 +694,23 @@ @Book{ fremlin_vol4 year = {2003} } +@Book{ friedmanscarr2005, + author = {Yaakov {Friedman}}, + title = {{Physical applications of homogeneous balls. With the + assistance of Tzvi Scarr}}, + fjournal = {{Progress in Mathematical Physics}}, + journal = {{Prog. Math. Phys.}}, + issn = {1544-9998}, + volume = {40}, + isbn = {0-8176-3339-1}, + pages = {xxiv + 279}, + year = {2005}, + publisher = {Boston, MA: Birkh\"auser}, + language = {English}, + msc2010 = {46-02 17C65 46L60 46G20 46L70 83A05}, + zbl = {1080.46001} +} + @Book{ fuchs1963, author = {Fuchs, L.}, title = {Partially ordered algebraic systems}, @@ -990,6 +1054,20 @@ @Article{ huneke2002 url = {https://doi.org/10.1080/00029890.2002.11919853} } +@Book{ iordanescu2003, + author = {Radu {Iord\u{a}nescu}}, + title = {{Jordan structures in geometry and physics. With an + appendix on Jordan structures in analysis}}, + isbn = {973-27-0956-1}, + pages = {201}, + year = {2003}, + publisher = {Bucharest: Editura Academiei Rom\^ane}, + language = {English}, + msc2010 = {17C50 17-02 17C65 32M15 35Q58 51A35 53C35 46H70 46K70 + 81R12 81R50}, + zbl = {1073.17014} +} + @Article{ izhakian2016, title = {Supertropical quadratic forms I}, journal = {Journal of Pure and Applied Algebra}, @@ -1195,6 +1273,22 @@ @InProceedings{ mcbride1996 organization = {Springer} } +@Book{ mccrimmon2004, + author = {Kevin {McCrimmon}}, + title = {{A taste of Jordan algebras}}, + fjournal = {{Universitext}}, + journal = {{Universitext}}, + issn = {0172-5939}, + isbn = {0-387-95447-3}, + pages = {xxvi + 562}, + year = {2004}, + publisher = {New York, NY: Springer}, + language = {English}, + doi = {10.1007/b97489}, + msc2010 = {17-01 17Cxx}, + zbl = {1044.17001} +} + @Misc{ melikhov2011, title = {Metrizable uniform spaces}, author = {Sergey A. Melikhov}, @@ -1670,6 +1764,23 @@ @Article{ tochiori_bertrand url = {https://www.chart.co.jp/subject/sugaku/suken_tsushin/76/76-8.pdf} } +@Book{ upmeier1987, + author = {Harald {Upmeier}}, + title = {{Jordan algebras in analysis, operator theory, and quantum + mechanics}}, + fjournal = {{Regional Conference Series in Mathematics}}, + journal = {{Reg. Conf. Ser. Math.}}, + issn = {0160-7642}, + volume = {67}, + isbn = {0-8218-0717-X}, + pages = {viii + 85}, + year = {1987}, + publisher = {Providence, RI: American Mathematical Society (AMS)}, + language = {English}, + msc2010 = {17-02 46-02 17C65 46H70 32M15 46G20 46L70 47B35 81Q99}, + zbl = {0608.17013} +} + @Article{ Vaisala_2003, author = {Jussi Väisälä}, title = {A Proof of the Mazur-Ulam Theorem}, diff --git a/src/algebra/jordan/basic.lean b/src/algebra/jordan/basic.lean new file mode 100644 index 0000000000000..0acbea42db6e6 --- /dev/null +++ b/src/algebra/jordan/basic.lean @@ -0,0 +1,223 @@ +/- +Copyright (c) 2021 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ +import algebra.ring.basic +import algebra.lie.of_associative +import data.real.basic +import linear_algebra.basic + +/-! +# Jordan rings + +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`, +``` +(a * b) * a^2 = a * (b * a^2) +``` +i.e. the operators of multiplication by `a` and `a^2` commute. + +A more general concept of a (non-commutative) Jordan ring can also be defined, as a +(non-commutative, non-associative) ring `A` where, for each `a` in `A`, the operators of left and +right multiplication by `a` and `a^2` commute. + +A real Jordan algebra `A` can be introduced by +```lean +variables {A : Type*} [non_unital_non_assoc_ring A] [module ℝ A] [smul_comm_class ℝ A A] + [is_scalar_tower ℝ A A] [is_comm_jordan A] +``` + +## Main results + +- `two_nsmul_lie_lmul_lmul_add_add_eq_zero` : Linearisation of the commutative Jordan axiom + +## Implementation notes + +We shall primarily be interested in linear Jordan algebras (i.e. over rings of characteristic not +two) leaving quadratic algebras to those better versed in that theory. + +The conventional way to linearise the Jordan axiom is to equate coefficients (more formally, assume +that the axiom holds in all field extensions). For simplicity we use brute force algebraic expansion +and substitution instead. + +## Motivation + +Every Jordan algebra `A` has a triple product defined, for `a` `b` and `c` in `A` by +$$ +{a\,b\,c} = (a * b) * c - (a * c) * b + a * (b * c). +$$ +Via this triple product Jordan algebras are related to a number of other mathematical structures: +Jordan triples, partial Jordan triples, Jordan pairs and quadratic Jordan algebras. In addition to +their considerable algebraic interest ([mccrimmon2004]) these structures have been shown to have +deep connections to mathematical physics, functional analysis and differential geometry. For more +information about these connections the interested reader is referred to [alfsenshultz2003], +[chu2012], [friedmanscarr2005], [iordanescu2003] and [upmeier1987]. + +Every associative algebra can be equipped with a symmetrized multiplication (characterized by +`sym_alg.sym_mul_sym`) making it into a commutative Jordan algebra. Jordan algebras arising this way +are said to be special. + +There are also exceptional Jordan algebras which can be shown not to be the symmetrization of any +associative algebra. The 3x3 matrices of octonions is the canonical example. + +Non-commutative Jordan algebras have connections to the Vidav-Palmer theorem +[cabreragarciarodriguezpalacios2014]. + +## References + +* [Cabrera García and Rodríguez Palacios, Non-associative normed algebras. Volume 1] + [cabreragarciarodriguezpalacios2014] +* [Hanche-Olsen and Størmer, Jordan Operator Algebras][hancheolsenstormer1984] +* [McCrimmon, A taste of Jordan algebras][mccrimmon2004] + +-/ + +variables (A : Type*) + +/-- A (non-commutative) Jordan multiplication. -/ +class is_jordan [has_mul A] := +(lmul_comm_rmul : ∀ a b : A, (a * b) * a = a * (b * a)) +(lmul_lmul_comm_lmul : ∀ a b : A, (a * a) * (a * b) = a * ((a * a) * b)) +(lmul_lmul_comm_rmul : ∀ a b : A, (a * a) * (b * a) = ((a * a) * b) * a) +(lmul_comm_rmul_rmul : ∀ a b : A, (a * b) * (a * a) = a * (b * (a * a))) +(rmul_comm_rmul_rmul : ∀ a b : A, (b * a) * (a * a) = (b * (a * a)) * a) + +/-- A commutative Jordan multipication -/ +class is_comm_jordan [has_mul A] := +(mul_comm : ∀ a b : A, a * b = b * a) +(lmul_comm_rmul_rmul : ∀ a b : A, (a * b) * (a * a) = a * (b * (a * a))) + +/-- A (commutative) Jordan multiplication is also a Jordan multipication -/ +@[priority 100] -- see Note [lower instance priority] +instance is_comm_jordan.to_is_jordan [has_mul A] [is_comm_jordan A] : is_jordan A := +{ lmul_comm_rmul := λ a b, by rw [is_comm_jordan.mul_comm, is_comm_jordan.mul_comm a b], + lmul_lmul_comm_lmul := λ a b, by rw [is_comm_jordan.mul_comm (a * a) (a * b), + is_comm_jordan.lmul_comm_rmul_rmul, is_comm_jordan.mul_comm b (a * a)], + lmul_comm_rmul_rmul := is_comm_jordan.lmul_comm_rmul_rmul, + lmul_lmul_comm_rmul := λ a b, by rw [is_comm_jordan.mul_comm (a * a) (b * a), + is_comm_jordan.mul_comm b a, is_comm_jordan.lmul_comm_rmul_rmul, is_comm_jordan.mul_comm, + is_comm_jordan.mul_comm b (a * a)], + rmul_comm_rmul_rmul := λ a b, by rw [is_comm_jordan.mul_comm b a, + is_comm_jordan.lmul_comm_rmul_rmul, is_comm_jordan.mul_comm], } + +/-- Semigroup multiplication satisfies the (non-commutative) Jordan axioms-/ +@[priority 100] -- see Note [lower instance priority] +instance semigroup.is_jordan [semigroup A] : is_jordan A := +{ lmul_comm_rmul := λ a b, by rw mul_assoc, + lmul_lmul_comm_lmul := λ a b, by rw [mul_assoc, mul_assoc], + lmul_comm_rmul_rmul := λ a b, by rw [mul_assoc], + lmul_lmul_comm_rmul := λ a b, by rw [←mul_assoc], + rmul_comm_rmul_rmul := λ a b, by rw [← mul_assoc, ← mul_assoc], } + +@[priority 100] -- see Note [lower instance priority] +instance comm_semigroup.is_comm_jordan [comm_semigroup A] : is_comm_jordan A := +{ mul_comm := mul_comm, + lmul_comm_rmul_rmul := λ a b, mul_assoc _ _ _, } + +local notation `L` := add_monoid.End.mul_left +local notation `R` := add_monoid.End.mul_right + +/-! +The Jordan axioms can be expressed in terms of commuting multiplication operators. +-/ +section commute +variables {A} [non_unital_non_assoc_ring A] [is_jordan A] + +@[simp] lemma commute_lmul_rmul (a : A) : commute (L a) (R a) := +add_monoid_hom.ext $ λ b, (is_jordan.lmul_comm_rmul _ _).symm + +@[simp] lemma commute_lmul_lmul_sq (a : A) : commute (L a) (L (a * a)) := +add_monoid_hom.ext $ λ b, (is_jordan.lmul_lmul_comm_lmul _ _).symm + +@[simp] lemma commute_lmul_rmul_sq (a : A) : commute (L a) (R (a * a)) := +add_monoid_hom.ext $ λ b, (is_jordan.lmul_comm_rmul_rmul _ _).symm + +@[simp] lemma commute_lmul_sq_rmul (a : A) : commute (L (a * a)) (R a) := +add_monoid_hom.ext $ λ b, (is_jordan.lmul_lmul_comm_rmul _ _) + +@[simp] lemma commute_rmul_rmul_sq (a : A) : commute (R a) (R (a * a)) := +add_monoid_hom.ext $ λ b, (is_jordan.rmul_comm_rmul_rmul _ _).symm + +end commute + +variables {A} [non_unital_non_assoc_ring A] [is_comm_jordan A] + +/-! +The endomorphisms on an additive monoid `add_monoid.End` form a `ring`, and this may be equipped +with a Lie Bracket via `ring.has_bracket`. +-/ + +lemma two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add (a b : A) : + 2•(⁅L a, L (a * b)⁆ + ⁅L b, L (b * a)⁆) = ⁅L (a * a), L b⁆ + ⁅L (b * b), L a⁆ := +begin + suffices : 2 • ⁅L a, L (a * b)⁆ + 2 • ⁅L b, L (b * a)⁆ + ⁅L b, L (a * a)⁆ + ⁅L a, L (b * b)⁆ = 0, + { rwa [← sub_eq_zero, ← sub_sub, sub_eq_add_neg, sub_eq_add_neg, lie_skew, lie_skew, nsmul_add] }, + convert (commute_lmul_lmul_sq (a + b)).lie_eq, + simp only [add_mul, mul_add, map_add, lie_add, add_lie, is_comm_jordan.mul_comm b a, + (commute_lmul_lmul_sq a).lie_eq, (commute_lmul_lmul_sq b).lie_eq], + abel, +end + +lemma two_nsmul_lie_lmul_lmul_add_add_eq_zero (a b c : A) : + 2•(⁅L a, L (b * c)⁆ + ⁅L b, L (c * a)⁆ + ⁅L c, L (a * b)⁆) = 0 := +begin + symmetry, + calc 0 = ⁅L (a + b + c), L ((a + b + c) * (a + b + c))⁆ : + by rw (commute_lmul_lmul_sq (a + b + c)).lie_eq + ... = ⁅L a + L b + L c, + L (a * a) + L (a * b) + L (a * c) + (L (b * a) + L (b * b) + L (b * c)) + + (L (c * a) + L (c * b) + L (c * c))⁆ : + by rw [add_mul, add_mul, mul_add, mul_add, mul_add, mul_add, mul_add, mul_add, + map_add, map_add, map_add, map_add, map_add, map_add, map_add, map_add, map_add, map_add] + ... = ⁅L a + L b + L c, + L (a * a) + L (a * b) + L (c * a) + (L (a * b) + L (b * b) + L (b * c)) + + (L (c * a) + L (b * c) + L (c * c))⁆ : + by rw [is_comm_jordan.mul_comm b a, is_comm_jordan.mul_comm c a, is_comm_jordan.mul_comm c b] + ... = ⁅L a + L b + L c, L (a * a) + L (b * b) + L (c * c) + 2•L (a * b) + 2•L (c * a) + + 2•L (b * c) ⁆ : + by {rw [two_smul, two_smul, two_smul], + simp only [lie_add, add_lie, commute_lmul_lmul_sq, zero_add, add_zero], abel} + ... = ⁅L a, L (a * a)⁆ + ⁅L a, L (b * b)⁆ + ⁅L a, L (c * c)⁆ + ⁅L a, 2•L (a * b)⁆ + + ⁅L a, 2•L(c * a)⁆ + ⁅L a, 2•L (b * c)⁆ + + (⁅L b, L (a * a)⁆ + ⁅L b, L (b * b)⁆ + ⁅L b, L (c * c)⁆ + ⁅L b, 2•L (a * b)⁆ + + ⁅L b, 2•L (c * a)⁆ + ⁅L b, 2•L (b * c)⁆) + + (⁅L c, L (a * a)⁆ + ⁅L c, L (b * b)⁆ + ⁅L c, L (c * c)⁆ + ⁅L c, 2•L (a * b)⁆ + + ⁅L c, 2•L (c * a)⁆ + ⁅L c, 2•L (b * c)⁆) : + by rw [add_lie, add_lie, lie_add, lie_add, lie_add, lie_add, lie_add, lie_add, lie_add, lie_add, + lie_add, lie_add, lie_add, lie_add, lie_add, lie_add, lie_add] + ... = ⁅L a, L (b * b)⁆ + ⁅L a, L (c * c)⁆ + ⁅L a, 2•L (a * b)⁆ + ⁅L a, 2•L (c * a)⁆ + + ⁅L a, 2•L (b * c)⁆ + + (⁅L b, L (a * a)⁆ + ⁅L b, L (c * c)⁆ + ⁅L b, 2•L (a * b)⁆ + ⁅L b, 2•L (c * a)⁆ + + ⁅L b, 2•L (b * c)⁆) + + (⁅L c, L (a * a)⁆ + ⁅L c, L (b * b)⁆ + ⁅L c, 2•L (a * b)⁆ + ⁅L c, 2•L (c * a)⁆ + + ⁅L c, 2•L (b * c)⁆) : + by rw [(commute_lmul_lmul_sq a).lie_eq, (commute_lmul_lmul_sq b).lie_eq, + (commute_lmul_lmul_sq c).lie_eq, zero_add, add_zero, add_zero] + ... = ⁅L a, L (b * b)⁆ + ⁅L a, L (c * c)⁆ + 2•⁅L a, L (a * b)⁆ + 2•⁅L a, L (c * a)⁆ + + 2•⁅L a, L (b * c)⁆ + + (⁅L b, L (a * a)⁆ + ⁅L b, L (c * c)⁆ + 2•⁅L b, L (a * b)⁆ + 2•⁅L b, L (c * a)⁆ + + 2•⁅L b, L (b * c)⁆) + + (⁅L c, L (a * a)⁆ + ⁅L c, L (b * b)⁆ + 2•⁅L c, L (a * b)⁆ + 2•⁅L c, L (c * a)⁆ + + 2•⁅L c, L (b * c)⁆) : + by simp only [lie_nsmul] + ... = (⁅L a, L (b * b)⁆+ ⁅L b, L (a * a)⁆ + 2•(⁅L a, L (a * b)⁆ + ⁅L b, L (a * b)⁆)) + + (⁅L a, L (c * c)⁆ + ⁅L c, L (a * a)⁆ + 2•(⁅L a, L (c * a)⁆ + ⁅L c, L (c * a)⁆)) + + (⁅L b, L (c * c)⁆ + ⁅L c, L (b * b)⁆ + 2•(⁅L b, L (b * c)⁆ + ⁅L c, L (b * c)⁆)) + + (2•⁅L a, L (b * c)⁆ + 2•⁅L b, L (c * a)⁆ + 2•⁅L c, L (a * b)⁆) : by abel + ... = 2•⁅L a, L (b * c)⁆ + 2•⁅L b, L (c * a)⁆ + 2•⁅L c, L (a * b)⁆ : + by begin + rw add_left_eq_self, + nth_rewrite 1 is_comm_jordan.mul_comm a b, + nth_rewrite 0 is_comm_jordan.mul_comm c a, + nth_rewrite 1 is_comm_jordan.mul_comm b c, + rw [two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add, + two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add, + two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add, + ← lie_skew (L (a * a)), ← lie_skew (L (b * b)), ← lie_skew (L (c * c)), + ← lie_skew (L (a * a)), ← lie_skew (L (b * b)), ← lie_skew (L (c * c))], + abel, + end + ... = 2•(⁅L a, L (b * c)⁆ + ⁅L b, L (c * a)⁆ + ⁅L c, L (a * b)⁆) : by rw [nsmul_add, nsmul_add] +end From 7236c304918f7e3c43b625b4700a32321ef5b401 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 26 Jul 2022 08:05:37 +0000 Subject: [PATCH 055/173] refactor(category_theory/*): use simps in the old parts of the library (#14236) Co-authored-by: Scott Morrison --- src/algebraic_topology/cech_nerve.lean | 6 +++--- src/category_theory/functor/basic.lean | 5 +++-- src/category_theory/functor/category.lean | 13 ++----------- src/category_theory/functor/const.lean | 14 +++----------- src/category_theory/functor/currying.lean | 19 ++----------------- src/category_theory/functor/hom.lean | 6 +----- src/category_theory/limits/colimit_limit.lean | 4 ++-- ...iltered_colimit_commutes_finite_limit.lean | 6 +++--- src/category_theory/limits/final.lean | 2 +- src/category_theory/limits/fubini.lean | 2 +- .../limits/shapes/equalizers.lean | 4 ++-- src/category_theory/natural_isomorphism.lean | 7 +------ src/category_theory/sites/left_exact.lean | 2 +- src/category_theory/structured_arrow.lean | 8 ++++---- src/group_theory/nielsen_schreier.lean | 2 +- src/topology/category/CompHaus/default.lean | 2 +- src/topology/sheaves/presheaf.lean | 12 +++++------- .../sheaf_condition/equalizer_products.lean | 2 +- 18 files changed, 37 insertions(+), 79 deletions(-) diff --git a/src/algebraic_topology/cech_nerve.lean b/src/algebraic_topology/cech_nerve.lean index 49f1e00f84407..ba1748f435241 100644 --- a/src/algebraic_topology/cech_nerve.lean +++ b/src/algebraic_topology/cech_nerve.lean @@ -122,7 +122,7 @@ def equivalence_left_to_right (X : simplicial_object.augmented C) (F : arrow C) (λ i, X.left.map (simplex_category.const x.unop i).op ≫ G.left) (λ i, by { dsimp, erw [category.assoc, arrow.w, augmented.to_arrow_obj_hom, nat_trans.naturality_assoc, - functor.const.obj_map, category.id_comp] } ), + functor.const_obj_map, category.id_comp] } ), naturality' := begin intros x y f, ext, @@ -131,7 +131,7 @@ def equivalence_left_to_right (X : simplicial_object.augmented C) (F : arrow C) rw [← category.assoc, ← X.left.map_comp], refl }, { dsimp, - simp only [functor.const.obj_map, nat_trans.naturality_assoc, + simp only [functor.const_obj_map, nat_trans.naturality_assoc, wide_pullback.lift_base, category.assoc], erw category.id_comp } end }, @@ -296,7 +296,7 @@ def equivalence_right_to_left (F : arrow C) (X : cosimplicial_object.augmented C rw [category.assoc, ←X.right.map_comp], refl }, { dsimp, - simp only [functor.const.obj_map, ←nat_trans.naturality, + simp only [functor.const_obj_map, ←nat_trans.naturality, wide_pushout.head_desc_assoc, wide_pushout.head_desc, category.assoc], erw category.id_comp } end }, diff --git a/src/category_theory/functor/basic.lean b/src/category_theory/functor/basic.lean index 377e00e5936bf..aaef887374d5f 100644 --- a/src/category_theory/functor/basic.lean +++ b/src/category_theory/functor/basic.lean @@ -60,6 +60,7 @@ section variables (C : Type u₁) [category.{v₁} C] /-- `𝟭 C` is the identity functor on a category `C`. -/ +-- We don't use `@[simps]` here because we want `C` implicit for the simp lemmas. protected def id : C ⥤ C := { obj := λ X, X, map := λ _ _ f, f } @@ -72,6 +73,7 @@ variable {C} @[simp] lemma id_obj (X : C) : (𝟭 C).obj X = X := rfl @[simp] lemma id_map {X Y : C} (f : X ⟶ Y) : (𝟭 C).map f = f := rfl + end section @@ -82,13 +84,12 @@ variables {C : Type u₁} [category.{v₁} C] /-- `F ⋙ G` is the composition of a functor `F` and a functor `G` (`F` first, then `G`). -/ -def comp (F : C ⥤ D) (G : D ⥤ E) : C ⥤ E := +@[simps obj] def comp (F : C ⥤ D) (G : D ⥤ E) : C ⥤ E := { obj := λ X, G.obj (F.obj X), map := λ _ _ f, G.map (F.map f) } infixr ` ⋙ `:80 := comp -@[simp] lemma comp_obj (F : C ⥤ D) (G : D ⥤ E) (X : C) : (F ⋙ G).obj X = G.obj (F.obj X) := rfl @[simp] lemma comp_map (F : C ⥤ D) (G : D ⥤ E) {X Y : C} (f : X ⟶ Y) : (F ⋙ G).map f = G.map (F.map f) := rfl diff --git a/src/category_theory/functor/category.lean b/src/category_theory/functor/category.lean index baa5ded9706ee..d19024ee0f9a5 100644 --- a/src/category_theory/functor/category.lean +++ b/src/category_theory/functor/category.lean @@ -75,7 +75,7 @@ lemma epi_app_of_epi (α : F ⟶ G) [∀ (X : C), epi (α.app X)] : epi α := ⟨λ H g h eq, by { ext X, rw [←cancel_epi (α.app X), ←comp_app, eq, comp_app] }⟩ /-- `hcomp α β` is the horizontal composition of natural transformations. -/ -def hcomp {H I : D ⥤ E} (α : F ⟶ G) (β : H ⟶ I) : (F ⋙ H) ⟶ (G ⋙ I) := +@[simps] def hcomp {H I : D ⥤ E} (α : F ⟶ G) (β : H ⟶ I) : (F ⋙ H) ⟶ (G ⋙ I) := { app := λ X : C, (β.app (F.obj X)) ≫ (I.map (α.app X)), naturality' := λ X Y f, begin @@ -85,9 +85,6 @@ def hcomp {H I : D ⥤ E} (α : F ⟶ G) (β : H ⟶ I) : (F ⋙ H) ⟶ (G ⋙ I infix ` ◫ `:80 := hcomp -@[simp] lemma hcomp_app {H I : D ⥤ E} (α : F ⟶ G) (β : H ⟶ I) (X : C) : - (α ◫ β).app X = (β.app (F.obj X)) ≫ (I.map (α.app X)) := rfl - @[simp] lemma hcomp_id_app {H : D ⥤ E} (α : F ⟶ G) (X : C) : (α ◫ 𝟙 H).app X = H.map (α.app X) := by {dsimp, simp} -- See note [dsimp, simp]. @@ -107,7 +104,7 @@ open nat_trans namespace functor /-- Flip the arguments of a bifunctor. See also `currying.lean`. -/ -protected def flip (F : C ⥤ (D ⥤ E)) : D ⥤ (C ⥤ E) := +@[simps] protected def flip (F : C ⥤ (D ⥤ E)) : D ⥤ (C ⥤ E) := { obj := λ k, { obj := λ j, (F.obj j).obj k, map := λ j j' f, (F.map f).app k, @@ -116,12 +113,6 @@ protected def flip (F : C ⥤ (D ⥤ E)) : D ⥤ (C ⥤ E) := map := λ c c' f, { app := λ j, (F.obj j).map f } }. -@[simp] lemma flip_obj_obj (F : C ⥤ (D ⥤ E)) (c) (d) : (F.flip.obj d).obj c = (F.obj c).obj d := rfl -@[simp] lemma flip_obj_map (F : C ⥤ (D ⥤ E)) {c c' : C} (f : c ⟶ c') (d : D) : - (F.flip.obj d).map f = (F.map f).app d := rfl -@[simp] lemma flip_map_app (F : C ⥤ (D ⥤ E)) {d d' : D} (f : d ⟶ d') (c : C) : - (F.flip.map f).app c = (F.obj c).map f := rfl - end functor @[simp, reassoc] lemma map_hom_inv_app (F : C ⥤ D ⥤ E) {X Y : C} (e : X ≅ Y) (Z : D) : diff --git a/src/category_theory/functor/const.lean b/src/category_theory/functor/const.lean index 13211cfe93cca..e1b08fce78845 100644 --- a/src/category_theory/functor/const.lean +++ b/src/category_theory/functor/const.lean @@ -29,7 +29,7 @@ variables {C : Type u₂} [category.{v₂} C] /-- The functor sending `X : C` to the constant functor `J ⥤ C` sending everything to `X`. -/ -def const : C ⥤ (J ⥤ C) := +@[simps] def const : C ⥤ (J ⥤ C) := { obj := λ X, { obj := λ j, X, map := λ j j' f, 𝟙 X }, @@ -40,22 +40,15 @@ open opposite variables {J} -@[simp] lemma obj_obj (X : C) (j : J) : ((const J).obj X).obj j = X := rfl -@[simp] lemma obj_map (X : C) {j j' : J} (f : j ⟶ j') : ((const J).obj X).map f = 𝟙 X := rfl -@[simp] lemma map_app {X Y : C} (f : X ⟶ Y) (j : J) : ((const J).map f).app j = f := rfl - /-- The contant functor `Jᵒᵖ ⥤ Cᵒᵖ` sending everything to `op X` is (naturally isomorphic to) the opposite of the constant functor `J ⥤ C` sending everything to `X`. -/ -def op_obj_op (X : C) : +@[simps] def op_obj_op (X : C) : (const Jᵒᵖ).obj (op X) ≅ ((const J).obj X).op := { hom := { app := λ j, 𝟙 _ }, inv := { app := λ j, 𝟙 _ } } -@[simp] lemma op_obj_op_hom_app (X : C) (j : Jᵒᵖ) : (op_obj_op X).hom.app j = 𝟙 _ := rfl -@[simp] lemma op_obj_op_inv_app (X : C) (j : Jᵒᵖ) : (op_obj_op X).inv.app j = 𝟙 _ := rfl - /-- The contant functor `Jᵒᵖ ⥤ C` sending everything to `unop X` is (naturally isomorphic to) the opposite of @@ -74,9 +67,8 @@ rfl @[simp] lemma unop_functor_op_obj_map (X : Cᵒᵖ) {j₁ j₂ : J} (f : j₁ ⟶ j₂) : (unop ((functor.op (const J)).obj X)).map f = 𝟙 (unop X) := rfl -end const - +end const section variables {D : Type u₃} [category.{v₃} D] diff --git a/src/category_theory/functor/currying.lean b/src/category_theory/functor/currying.lean index 6285dff51ee58..b62d53fbcbad9 100644 --- a/src/category_theory/functor/currying.lean +++ b/src/category_theory/functor/currying.lean @@ -24,6 +24,7 @@ variables {C : Type u₁} [category.{v₁} C] /-- The uncurrying functor, taking a functor `C ⥤ (D ⥤ E)` and producing a functor `(C × D) ⥤ E`. -/ +@[simps] def uncurry : (C ⥤ (D ⥤ E)) ⥤ ((C × D) ⥤ E) := { obj := λ F, { obj := λ X, (F.obj X.1).obj X.2, @@ -58,6 +59,7 @@ def curry_obj (F : (C × D) ⥤ E) : C ⥤ (D ⥤ E) := /-- The currying functor, taking a functor `(C × D) ⥤ E` and producing a functor `C ⥤ (D ⥤ E)`. -/ +@[simps obj_obj_obj obj_obj_map obj_map_app map_app_app] def curry : ((C × D) ⥤ E) ⥤ (C ⥤ (D ⥤ E)) := { obj := λ F, curry_obj F, map := λ F G T, @@ -74,23 +76,6 @@ def curry : ((C × D) ⥤ E) ⥤ (C ⥤ (D ⥤ E)) := rw nat_trans.naturality, end } }. -@[simp] lemma uncurry.obj_obj {F : C ⥤ (D ⥤ E)} {X : C × D} : - (uncurry.obj F).obj X = (F.obj X.1).obj X.2 := rfl -@[simp] lemma uncurry.obj_map {F : C ⥤ (D ⥤ E)} {X Y : C × D} {f : X ⟶ Y} : - (uncurry.obj F).map f = ((F.map f.1).app X.2) ≫ ((F.obj Y.1).map f.2) := rfl -@[simp] lemma uncurry.map_app {F G : C ⥤ (D ⥤ E)} {α : F ⟶ G} {X : C × D} : - (uncurry.map α).app X = (α.app X.1).app X.2 := rfl -@[simp] lemma curry.obj_obj_obj - {F : (C × D) ⥤ E} {X : C} {Y : D} : - ((curry.obj F).obj X).obj Y = F.obj (X, Y) := rfl -@[simp] lemma curry.obj_obj_map - {F : (C × D) ⥤ E} {X : C} {Y Y' : D} {g : Y ⟶ Y'} : - ((curry.obj F).obj X).map g = F.map (𝟙 X, g) := rfl -@[simp] lemma curry.obj_map_app {F : (C × D) ⥤ E} {X X' : C} {f : X ⟶ X'} {Y} : - ((curry.obj F).map f).app Y = F.map (f, 𝟙 Y) := rfl -@[simp] lemma curry.map_app_app {F G : (C × D) ⥤ E} {α : F ⟶ G} {X} {Y} : - ((curry.map α).app X).app Y = α.app (X, Y) := rfl - /-- The equivalence of functor categories given by currying/uncurrying. -/ diff --git a/src/category_theory/functor/hom.lean b/src/category_theory/functor/hom.lean index 05b6f7a0e7047..ee307cb400b05 100644 --- a/src/category_theory/functor/hom.lean +++ b/src/category_theory/functor/hom.lean @@ -21,12 +21,8 @@ variables (C : Type u) [category.{v} C] /-- `functor.hom` is the hom-pairing, sending `(X, Y)` to `X ⟶ Y`, contravariant in `X` and covariant in `Y`. -/ -definition hom : Cᵒᵖ × C ⥤ Type v := +@[simps] def hom : Cᵒᵖ × C ⥤ Type v := { obj := λ p, unop p.1 ⟶ p.2, map := λ X Y f, λ h, f.1.unop ≫ h ≫ f.2 } -@[simp] lemma hom_obj (X : Cᵒᵖ × C) : (hom C).obj X = (unop X.1 ⟶ X.2) := rfl -@[simp] lemma hom_pairing_map {X Y : Cᵒᵖ × C} (f : X ⟶ Y) : - (hom C).map f = λ h, f.1.unop ≫ h ≫ f.2 := rfl - end category_theory.functor diff --git a/src/category_theory/limits/colimit_limit.lean b/src/category_theory/limits/colimit_limit.lean index 960039ce81970..50e24cd03ff4f 100644 --- a/src/category_theory/limits/colimit_limit.lean +++ b/src/category_theory/limits/colimit_limit.lean @@ -65,7 +65,7 @@ limit.lift ((curry.obj F) ⋙ colim) begin dsimp, intros k k' f, - simp only [functor.comp_map, curry.obj_map_app, limits.lim_map_π_assoc, swap_map, + simp only [functor.comp_map, curry_obj_map_app, limits.lim_map_π_assoc, swap_map, category.comp_id, map_id_left_eq_curry_map, colimit.w], end }, }, naturality' := @@ -73,7 +73,7 @@ limit.lift ((curry.obj F) ⋙ colim) dsimp, intros j j' f, ext k, - simp only [limits.colimit.ι_map, curry.obj_map_app, limits.colimit.ι_desc_assoc, + simp only [limits.colimit.ι_map, curry_obj_map_app, limits.colimit.ι_desc_assoc, limits.colimit.ι_desc, category.id_comp, category.assoc, map_id_right_eq_curry_swap_map, limit.w_assoc], end } } 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 de151200a2e8b..5ac3bd400749c 100644 --- a/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean +++ b/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean @@ -124,7 +124,7 @@ begin ext, -- Now it's just a calculation using `W` and `w`. - simp only [functor.comp_map, limit.map_π_apply, curry.obj_map_app, swap_map], + simp only [functor.comp_map, limit.map_π_apply, curry_obj_map_app, swap_map], rw ←W _ _ (fH j), rw ←W _ _ (gH j), simp [w], @@ -293,7 +293,7 @@ begin -- and as each component is an equation in a colimit, we can verify it by -- pointing out the morphism which carries one representative to the other: - simp only [←e, colimit_eq_iff.{v v}, curry.obj_obj_map, limit.π_mk', + simp only [←e, colimit_eq_iff.{v v}, curry_obj_obj_map, limit.π_mk', bifunctor.map_id_comp, id.def, types_comp_apply, limits.ι_colimit_limit_to_limit_colimit_π_apply], refine ⟨k'', 𝟙 k'', g j ≫ gf (𝟙 j) ≫ i (𝟙 j), _⟩, @@ -378,7 +378,7 @@ begin congr' 1, simp only [← category.assoc, iso.comp_inv_eq, limits.colimit_obj_iso_colimit_comp_evaluation_ι_app_hom, - limits.has_colimit.iso_of_nat_iso_ι_hom, nat_iso.of_components.hom_app], + limits.has_colimit.iso_of_nat_iso_ι_hom, nat_iso.of_components_hom_app], dsimp, simp, end diff --git a/src/category_theory/limits/final.lean b/src/category_theory/limits/final.lean index aae5fb0bc91f6..338ad5712481b 100644 --- a/src/category_theory/limits/final.lean +++ b/src/category_theory/limits/final.lean @@ -266,7 +266,7 @@ begin dsimp [is_colimit_whisker_equiv], apply P.hom_ext, intro j, - dsimp, simp, dsimp, simp, -- See library note [dsimp, simp]. + dsimp, simp, end instance colimit_pre_is_iso [has_colimit G] : diff --git a/src/category_theory/limits/fubini.lean b/src/category_theory/limits/fubini.lean index b66f019afc4c2..3dad12f3e9385 100644 --- a/src/category_theory/limits/fubini.lean +++ b/src/category_theory/limits/fubini.lean @@ -127,7 +127,7 @@ def cone_of_cone_uncurry_is_limit dsimp, slice_rhs 3 4 { rw ←nat_trans.naturality, }, slice_rhs 2 3 { rw ←(D.obj j).π.naturality, }, - simp only [functor.const.obj_map, category.id_comp, category.assoc], + simp only [functor.const_obj_map, category.id_comp, category.assoc], have w := (D.map fj).w k', dsimp at w, rw ←w, diff --git a/src/category_theory/limits/shapes/equalizers.lean b/src/category_theory/limits/shapes/equalizers.lean index b2d4fc9da76f2..1126c479f3887 100644 --- a/src/category_theory/limits/shapes/equalizers.lean +++ b/src/category_theory/limits/shapes/equalizers.lean @@ -524,7 +524,7 @@ def fork.ext {s t : fork f g} (i : s.X ≅ t.X) (w : i.hom ≫ t.ι = s.ι) : s /-- Every fork is isomorphic to one of the form `fork.of_ι _ _`. -/ def fork.iso_fork_of_ι (c : fork f g) : c ≅ fork.of_ι c.ι c.condition := -fork.ext (by simp only [fork.of_ι_X, functor.const.obj_obj]) (by simp) +fork.ext (by simp only [fork.of_ι_X, functor.const_obj_obj]) (by simp) /-- Helper function for constructing morphisms between coequalizer coforks. @@ -557,7 +557,7 @@ def cofork.ext {s t : cofork f g} (i : s.X ≅ t.X) (w : s.π ≫ i.hom = t.π) /-- Every cofork is isomorphic to one of the form `cofork.of_π _ _`. -/ def cofork.iso_cofork_of_π (c : cofork f g) : c ≅ cofork.of_π c.π c.condition := -cofork.ext (by simp only [cofork.of_π_X, functor.const.obj_obj]) (by dsimp; simp) +cofork.ext (by simp only [cofork.of_π_X, functor.const_obj_obj]) (by dsimp; simp) variables (f g) diff --git a/src/category_theory/natural_isomorphism.lean b/src/category_theory/natural_isomorphism.lean index db6947c5a6bcb..2c73960fdb04b 100644 --- a/src/category_theory/natural_isomorphism.lean +++ b/src/category_theory/natural_isomorphism.lean @@ -163,7 +163,7 @@ by { ext, simp, } Construct a natural isomorphism between functors by giving object level isomorphisms, and checking naturality only in the forward direction. -/ -def of_components (app : ∀ X : C, F.obj X ≅ G.obj X) +@[simps] def of_components (app : ∀ X : C, F.obj X ≅ G.obj X) (naturality : ∀ {X Y : C} (f : X ⟶ Y), F.map f ≫ (app Y).hom = (app X).hom ≫ G.map f) : F ≅ G := { hom := { app := λ X, (app X).hom }, @@ -179,11 +179,6 @@ def of_components (app : ∀ X : C, F.obj X ≅ G.obj X) @[simp] lemma of_components.app (app' : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X) : (of_components app' naturality).app X = app' X := by tidy -@[simp] lemma of_components.hom_app (app : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X) : - (of_components app naturality).hom.app X = (app X).hom := rfl -@[simp] lemma of_components.inv_app (app : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X) : - (of_components app naturality).inv.app X = (app X).inv := -by simp [of_components] /-- A natural transformation is an isomorphism if all its components are isomorphisms. diff --git a/src/category_theory/sites/left_exact.lean b/src/category_theory/sites/left_exact.lean index 6081c333e841b..74b59a7f4d69a 100644 --- a/src/category_theory/sites/left_exact.lean +++ b/src/category_theory/sites/left_exact.lean @@ -192,7 +192,7 @@ begin dsimp, simp only [category.assoc], rw ι_colimit_limit_iso_limit_π_assoc, - simp only [nat_iso.of_components.inv_app, + simp only [nat_iso.of_components_inv_app, colimit_obj_iso_colimit_comp_evaluation_ι_app_hom, iso.symm_inv], dsimp [is_limit.cone_point_unique_up_to_iso], rw [← category.assoc, ← nat_trans.comp_app, limit.lift_π], diff --git a/src/category_theory/structured_arrow.lean b/src/category_theory/structured_arrow.lean index e4b6497d49b47..5b765d7c15d3b 100644 --- a/src/category_theory/structured_arrow.lean +++ b/src/category_theory/structured_arrow.lean @@ -266,7 +266,7 @@ def to_costructured_arrow (F : C ⥤ D) (d : D) : map := λ X Y f, costructured_arrow.hom_mk (f.unop.right.op) begin dsimp, - rw [← op_comp, ← f.unop.w, functor.const.obj_map], + rw [← op_comp, ← f.unop.w, functor.const_obj_map], erw category.id_comp, end } @@ -283,7 +283,7 @@ def to_costructured_arrow' (F : C ⥤ D) (d : D) : begin dsimp, rw [← quiver.hom.unop_op (F.map (quiver.hom.unop f.unop.right)), ← unop_comp, ← F.op_map, - ← f.unop.w, functor.const.obj_map], + ← f.unop.w, functor.const_obj_map], erw category.id_comp, end } @@ -303,7 +303,7 @@ def to_structured_arrow (F : C ⥤ D) (d : D) : map := λ X Y f, structured_arrow.hom_mk f.unop.left.op begin dsimp, - rw [← op_comp, f.unop.w, functor.const.obj_map], + rw [← op_comp, f.unop.w, functor.const_obj_map], erw category.comp_id, end } @@ -320,7 +320,7 @@ def to_structured_arrow' (F : C ⥤ D) (d : D) : begin dsimp, rw [← quiver.hom.unop_op (F.map f.unop.left.unop), ← unop_comp, ← F.op_map, - f.unop.w, functor.const.obj_map], + f.unop.w, functor.const_obj_map], erw category.comp_id, end } diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index 4bdabc07279dd..6d863349658e4 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -248,7 +248,7 @@ begin let F : G ⥤ 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], + rw [functor.const_obj_map, id_as_one, difference_functor_map, mul_inv_eq_one], apply congr_arg free_group.of, apply (weakly_connected_component.eq _ _).mpr, exact ⟨hom.to_path (sum.inr e)⟩, diff --git a/src/topology/category/CompHaus/default.lean b/src/topology/category/CompHaus/default.lean index cecaffef17560..8c58caf89086c 100644 --- a/src/topology/category/CompHaus/default.lean +++ b/src/topology/category/CompHaus/default.lean @@ -190,7 +190,7 @@ def limit_cone {J : Type v} [small_category J] (F : J ⥤ CompHaus.{max v u}) : π := { app := λ j, (Top.limit_cone (F ⋙ CompHaus_to_Top)).π.app j, naturality' := by { intros _ _ _, ext ⟨x, hx⟩, - simp only [comp_apply, functor.const.obj_map, id_apply], exact (hx f).symm, } } } + simp only [comp_apply, functor.const_obj_map, id_apply], exact (hx f).symm, } } } /-- The limit cone `CompHaus.limit_cone F` is indeed a limit cone. -/ def limit_cone_is_limit {J : Type v} [small_category J] (F : J ⥤ CompHaus.{max v u}) : diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index d892a09e2471d..b27deb2fb58cb 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -208,13 +208,11 @@ lemma id_inv_app (U : opens Y) : (id ℱ).inv.app (op U) = colimit.ι (Lan.diagram (opens.map (𝟙 Y)).op ℱ (op U)) (@costructured_arrow.mk _ _ _ _ _ (op U) _ (eq_to_hom (by simp))) := begin - dsimp[id], simp, dsimp[colimit_of_diagram_terminal], - delta Lan.diagram, - refine eq.trans _ (category.id_comp _), - rw ← ℱ.map_id, - congr, - any_goals { apply subsingleton.helim }, - all_goals { simp } + rw [← category.id_comp ((id ℱ).inv.app (op U)), ← nat_iso.app_inv, iso.comp_inv_eq], + dsimp [id], + rw colimit.ι_desc_assoc, + dsimp, + rw [← ℱ.map_comp, ← ℱ.map_id], refl, end end pullback diff --git a/src/topology/sheaves/sheaf_condition/equalizer_products.lean b/src/topology/sheaves/sheaf_condition/equalizer_products.lean index 12a18d8ac582a..a6285e62b494d 100644 --- a/src/topology/sheaves/sheaf_condition/equalizer_products.lean +++ b/src/topology/sheaves/sheaf_condition/equalizer_products.lean @@ -241,7 +241,7 @@ begin simp only [res, diagram.iso_of_open_embedding, discrete.nat_iso_inv_app, functor.map_iso_inv, limit.lift_π, cones.postcompose_obj_π, functor.comp_map, fork_π_app_walking_parallel_pair_zero, pi_opens.iso_of_open_embedding, - nat_iso.of_components.inv_app, functor.map_iso_refl, functor.op_map, limit.lift_map, + nat_iso.of_components_inv_app, functor.map_iso_refl, functor.op_map, limit.lift_map, fan.mk_π_app, nat_trans.comp_app, quiver.hom.unop_op, category.assoc, lim_map_eq_lim_map], dsimp, rw [category.comp_id, ←F.map_comp], From 7543a9b20e3e46e12cc47c90c5ae0a8161eadb13 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 26 Jul 2022 10:45:55 +0000 Subject: [PATCH 056/173] chore(algebra/jordan/basic): remove unused imports (#15686) We don't need any linear algebra or the real numbers here --- src/algebra/jordan/basic.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/algebra/jordan/basic.lean b/src/algebra/jordan/basic.lean index 0acbea42db6e6..4dd058d744d60 100644 --- a/src/algebra/jordan/basic.lean +++ b/src/algebra/jordan/basic.lean @@ -3,10 +3,8 @@ Copyright (c) 2021 Christopher Hoskin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ -import algebra.ring.basic import algebra.lie.of_associative -import data.real.basic -import linear_algebra.basic +import algebra.ring.basic /-! # Jordan rings From 60e4d7a2de24f754b075ce1b3168291623b81047 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:00 +0000 Subject: [PATCH 057/173] feat(order/filter/*): `filter.pi` is countably generated (#15632) * rename `filter.has_basis_infi` to `filter.has_basis_infi'`, add new `filter.has_basis_infi`; * move `prod.is_countably_generated`, golf, add `coprod.is_countably_generated`; * `is_countably_generated_seq` is no longer an instance, `infi.is_countably_generated` is better; * add `infi.is_countably_generated` and `pi.is_countably_generated`; * prove `prod.fist_countable_topology` (from the sphere eversion project) and `pi.first_countable_topology`; * generalize `pi.second_countable_topology` from `encodable` to `countable` so that it automatically applies to finite types. --- src/order/filter/bases.lean | 51 ++++++++++++++----- src/order/filter/pi.lean | 6 ++- .../algebra/module/locally_convex.lean | 2 +- src/topology/bases.lean | 24 ++++++--- 4 files changed, 61 insertions(+), 22 deletions(-) diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index cc646bcbb2572..6feb44cc4f7bf 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -415,7 +415,7 @@ lemma has_basis.inf {ι ι' : Type*} {p : ι → Prop} {s : ι → set α} {p' : (hl.inf' hl').to_has_basis (λ i hi, ⟨⟨i.1, i.2⟩, hi, subset.rfl⟩) (λ i hi, ⟨⟨i.1, i.2⟩, hi, subset.rfl⟩) -lemma has_basis_infi {ι : Type*} {ι' : ι → Type*} {l : ι → filter α} +lemma has_basis_infi' {ι : Type*} {ι' : ι → Type*} {l : ι → filter α} {p : Π i, ι' i → Prop} {s : Π i, ι' i → set α} (hl : ∀ i, (l i).has_basis (p i) (s i)) : (⨅ i, l i).has_basis (λ If : set ι × Π i, ι' i, If.1.finite ∧ ∀ i ∈ If.1, p i (If.2 i)) (λ If : set ι × Π i, ι' i, ⋂ i ∈ If.1, s i (If.2 i)) := @@ -423,16 +423,29 @@ lemma has_basis_infi {ι : Type*} {ι' : ι → Type*} {l : ι → filter α} intro t, split, { simp only [mem_infi', (hl _).mem_iff], - rintros ⟨I, hI, V, hV, -, hVt, -⟩, + rintros ⟨I, hI, V, hV, -, rfl, -⟩, choose u hu using hV, - refine ⟨⟨I, u⟩, ⟨hI, λ i _, (hu i).1⟩, _⟩, - rw hVt, - exact Inter_mono (λ i, Inter_mono $ λ hi, (hu i).2) }, + exact ⟨⟨I, u⟩, ⟨hI, λ i _, (hu i).1⟩, Inter_mono (λ i, Inter_mono $ λ hi, (hu i).2)⟩ }, { rintros ⟨⟨I, f⟩, ⟨hI₁, hI₂⟩, hsub⟩, refine mem_of_superset _ hsub, exact (bInter_mem hI₁).mpr (λ i hi, mem_infi_of_mem i $ (hl i).mem_of_mem $ hI₂ _ hi) } end⟩ +lemma has_basis_infi {ι : Type*} {ι' : ι → Type*} {l : ι → filter α} + {p : Π i, ι' i → Prop} {s : Π i, ι' i → set α} (hl : ∀ i, (l i).has_basis (p i) (s i)) : + (⨅ i, l i).has_basis (λ If : Σ I : set ι, Π i : I, ι' i, If.1.finite ∧ ∀ i : If.1, p i (If.2 i)) + (λ If, ⋂ i : If.1, s i (If.2 i)) := +begin + refine ⟨λ t, ⟨λ ht, _, _⟩⟩, + { rcases (has_basis_infi' hl).mem_iff.mp ht with ⟨⟨I, f⟩, ⟨hI, hf⟩, hsub⟩, + exact ⟨⟨I, λ i, f i⟩, ⟨hI, subtype.forall.mpr hf⟩, + trans_rel_right _ (Inter_subtype _ _) hsub⟩ }, + { rintro ⟨⟨I, f⟩, ⟨hI, hf⟩, hsub⟩, + refine mem_of_superset _ hsub, + casesI hI.nonempty_fintype, + exact Inter_mem.2 (λ i, mem_infi_of_mem i $ (hl i).mem_of_mem $ hf _) } +end + lemma has_basis_infi_of_directed' {ι : Type*} {ι' : ι → Sort*} [nonempty ι] {l : ι → filter α} (s : Π i, (ι' i) → set α) (p : Π i, (ι' i) → Prop) @@ -932,9 +945,17 @@ begin ⟨hs.to_has_basis.sup ht.to_has_basis, set.to_countable _⟩ end +instance prod.is_countably_generated (la : filter α) (lb : filter β) [is_countably_generated la] + [is_countably_generated lb] : is_countably_generated (la ×ᶠ lb) := +filter.inf.is_countably_generated _ _ + +instance coprod.is_countably_generated (la : filter α) (lb : filter β) [is_countably_generated la] + [is_countably_generated lb] : is_countably_generated (la.coprod lb) := +filter.sup.is_countably_generated _ _ + end is_countably_generated -@[instance] lemma is_countably_generated_seq [encodable β] (x : β → set α) : +lemma is_countably_generated_seq [encodable β] (x : β → set α) : is_countably_generated (⨅ i, 𝓟 $ x i) := begin use [range x, countable_range x], @@ -971,14 +992,18 @@ by { rw ← principal_singleton, exact is_countably_generated_principal _, } @[instance] lemma is_countably_generated_top : is_countably_generated (⊤ : filter α) := @principal_univ α ▸ is_countably_generated_principal _ -instance is_countably_generated.prod {f : filter α} {g : filter β} - [hf : f.is_countably_generated] [hg : g.is_countably_generated] : - is_countably_generated (f ×ᶠ g) := +instance infi.is_countably_generated {ι : Sort*} [countable ι] (f : ι → filter α) + [∀ i, is_countably_generated (f i)] : is_countably_generated (⨅ i, f i) := begin - simp_rw is_countably_generated_iff_exists_antitone_basis at hf hg ⊢, - rcases hf with ⟨s, hs⟩, - rcases hg with ⟨t, ht⟩, - refine ⟨_, hs.prod ht⟩, + choose s hs using λ i, exists_antitone_basis (f i), + rw [← plift.down_surjective.infi_comp], + refine has_countable_basis.is_countably_generated + ⟨has_basis_infi (λ n, (hs _).to_has_basis), _⟩, + haveI := encodable.of_countable (plift ι), + refine (countable_range $ sigma.map (coe : finset (plift ι) → set (plift ι)) (λ _, id)).mono _, + rintro ⟨I, f⟩ ⟨hI, -⟩, + lift I to finset (plift ι) using hI, + exact ⟨⟨I, f⟩, rfl⟩ end end filter diff --git a/src/order/filter/pi.lean b/src/order/filter/pi.lean index 1557a4eeef974..f1c6e3a7ee21f 100644 --- a/src/order/filter/pi.lean +++ b/src/order/filter/pi.lean @@ -31,6 +31,10 @@ section pi /-- The product of an indexed family of filters. -/ def pi (f : Π i, filter (α i)) : filter (Π i, α i) := ⨅ i, comap (eval i) (f i) +instance pi.is_countably_generated [countable ι] [∀ i, is_countably_generated (f i)] : + is_countably_generated (pi f) := +infi.is_countably_generated _ + lemma tendsto_eval_pi (f : Π i, filter (α i)) (i : ι) : tendsto (eval i) (pi f) (f i) := tendsto_infi' i tendsto_comap @@ -91,7 +95,7 @@ lemma has_basis_pi {ι' : ι → Type} {s : Π i, ι' i → set (α i)} {p : Π (pi f).has_basis (λ If : set ι × Π i, ι' i, If.1.finite ∧ ∀ i ∈ If.1, p i (If.2 i)) (λ If : set ι × Π i, ι' i, If.1.pi (λ i, s i $ If.2 i)) := begin - have : (pi f).has_basis _ _ := has_basis_infi (λ i, (h i).comap (eval i : (Π j, α j) → α i)), + have : (pi f).has_basis _ _ := has_basis_infi' (λ i, (h i).comap (eval i : (Π j, α j) → α i)), convert this, ext, simp diff --git a/src/topology/algebra/module/locally_convex.lean b/src/topology/algebra/module/locally_convex.lean index 31b178cdec7db..3e8443f25d3e6 100644 --- a/src/topology/algebra/module/locally_convex.lean +++ b/src/topology/algebra/module/locally_convex.lean @@ -109,7 +109,7 @@ begin ((If.2 i) ∈ @nhds _ ↑i x ∧ convex 𝕜 (If.2 i))) (λ x, _) (λ x If hif, convex_Inter $ λ i, convex_Inter $ λ hi, (hif.2 i hi).2), rw [nhds_Inf, ← infi_subtype''], - exact has_basis_infi (λ i : ts, (@locally_convex_space_iff 𝕜 E _ _ _ ↑i).mp (h ↑i i.2) x), + exact has_basis_infi' (λ i : ts, (@locally_convex_space_iff 𝕜 E _ _ _ ↑i).mp (h ↑i i.2) x), end lemma locally_convex_space_infi {ts' : ι → topological_space E} diff --git a/src/topology/bases.lean b/src/topology/bases.lean index c57ca30354f10..44c214414a832 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -530,6 +530,20 @@ end first_countable_topology variables {α} +instance {β} [topological_space β] [first_countable_topology α] [first_countable_topology β] : + first_countable_topology (α × β) := +⟨λ ⟨x, y⟩, by { rw nhds_prod_eq, apply_instance }⟩ + +section pi + +omit t + +instance {ι : Type*} {π : ι → Type*} [countable ι] [Π i, topological_space (π i)] + [∀ i, first_countable_topology (π i)] : first_countable_topology (Π i, π i) := +⟨λ f, by { rw nhds_pi, apply_instance }⟩ + +end pi + instance is_countably_generated_nhds_within (x : α) [is_countably_generated (𝓝 x)] (s : set α) : is_countably_generated (𝓝[s] x) := inf.is_countably_generated _ _ @@ -622,10 +636,11 @@ instance {β : Type*} [topological_space β] ((is_basis_countable_basis α).prod (is_basis_countable_basis β)).second_countable_topology $ (countable_countable_basis α).image2 (countable_countable_basis β) _ -instance second_countable_topology_encodable {ι : Type*} {π : ι → Type*} - [encodable ι] [t : ∀a, topological_space (π a)] [∀a, second_countable_topology (π a)] : +instance {ι : Type*} {π : ι → Type*} + [countable ι] [t : ∀a, topological_space (π a)] [∀a, second_countable_topology (π a)] : second_countable_topology (∀a, π a) := begin + haveI := encodable.of_countable ι, have : t = (λa, generate_from (countable_basis (π a))), from funext (assume a, (is_basis_countable_basis (π a)).eq_generate_from), rw [this, pi_generate_from_eq], @@ -645,11 +660,6 @@ begin exact ⟨s, I, λ i hi, hs ⟨i, hi⟩, set.ext $ λ f, subtype.forall⟩ } end -instance second_countable_topology_fintype {ι : Type*} {π : ι → Type*} - [fintype ι] [t : ∀a, topological_space (π a)] [∀a, second_countable_topology (π a)] : - second_countable_topology (∀a, π a) := -by { letI := fintype.to_encodable ι, exact topological_space.second_countable_topology_encodable } - @[priority 100] -- see Note [lower instance priority] instance second_countable_topology.to_separable_space [second_countable_topology α] : separable_space α := From 0d79d1143dbd3d5178741f5baa47b2fc73e2fc97 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:01 +0000 Subject: [PATCH 058/173] feat(topology/metric_space/lipschitz): add several lemmas (#15634) Some of these lemmas come from the sphere eversion project, with slightly different names. --- src/topology/metric_space/lipschitz.lean | 25 ++++++++++++++++++++---- 1 file changed, 21 insertions(+), 4 deletions(-) diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index fd3303832e9cd..b0279dedeb874 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -84,6 +84,15 @@ lemma lipschitz_on_with_iff_restrict [pseudo_emetric_space α] [pseudo_emetric_s {f : α → β} {s : set α} : lipschitz_on_with K f s ↔ lipschitz_with K (s.restrict f) := by simp only [lipschitz_on_with, lipschitz_with, set_coe.forall', restrict, subtype.edist_eq] +alias lipschitz_on_with_iff_restrict ↔ lipschitz_on_with.to_restrict _ + +lemma maps_to.lipschitz_on_with_iff_restrict [pseudo_emetric_space α] [pseudo_emetric_space β] + {K : ℝ≥0} {f : α → β} {s : set α} {t : set β} (h : maps_to f s t) : + lipschitz_on_with K f s ↔ lipschitz_with K (h.restrict f s t) := +lipschitz_on_with_iff_restrict + +alias maps_to.lipschitz_on_with_iff_restrict ↔ lipschitz_on_with.to_restrict_maps_to _ + namespace lipschitz_with section emetric @@ -194,10 +203,7 @@ calc edist (f (g x)) (f (g y)) ≤ Kf * edist (g x) (g y) : hf _ _ lemma comp_lipschitz_on_with {Kf Kg : ℝ≥0} {f : β → γ} {g : α → β} {s : set α} (hf : lipschitz_with Kf f) (hg : lipschitz_on_with Kg g s) : lipschitz_on_with (Kf * Kg) (f ∘ g) s := -assume x hx y hy, -calc edist (f (g x)) (f (g y)) ≤ Kf * edist (g x) (g y) : hf _ _ -... ≤ Kf * (Kg * edist x y) : ennreal.mul_left_mono (hg hx hy) -... = (Kf * Kg : ℝ≥0) * edist x y : by rw [← mul_assoc, ennreal.coe_mul] +lipschitz_on_with_iff_restrict.mpr $ hf.comp hg.to_restrict protected lemma prod_fst : lipschitz_with 1 (@prod.fst α β) := lipschitz_with.of_edist_le $ assume x y, le_max_left _ _ @@ -214,6 +220,12 @@ begin exact max_le_max (hf x y) (hg x y) end +protected lemma prod_mk_left (a : α) : lipschitz_with 1 (prod.mk a : β → α × β) := +by simpa only [max_eq_right zero_le_one] using (lipschitz_with.const a).prod lipschitz_with.id + +protected lemma prod_mk_right (b : α) : lipschitz_with 1 (λ a : α, (a, b)) := +by simpa only [max_eq_left zero_le_one] using lipschitz_with.id.prod (lipschitz_with.const b) + protected lemma uncurry {f : α → β → γ} {Kα Kβ : ℝ≥0} (hα : ∀ b, lipschitz_with Kα (λ a, f a b)) (hβ : ∀ a, lipschitz_with Kβ (f a)) : lipschitz_with (Kα + Kβ) (function.uncurry f) := @@ -444,6 +456,11 @@ lemma edist_lt_of_edist_lt_div (hf : lipschitz_on_with K f s) {x y : α} (hx : x (lipschitz_on_with_iff_restrict.mp hf).edist_lt_of_edist_lt_div $ show edist (⟨x, hx⟩ : s) ⟨y, hy⟩ < d / K, from hd +protected lemma comp {g : β → γ} {t : set β} {Kg : ℝ≥0} (hg : lipschitz_on_with Kg g t) + (hf : lipschitz_on_with K f s) (hmaps : maps_to f s t) : + lipschitz_on_with (Kg * K) (g ∘ f) s := +lipschitz_on_with_iff_restrict.mpr $ hg.to_restrict.comp (hf.to_restrict_maps_to hmaps) + end emetric section metric From e1a7ea95967dd5879697a219319a46a8c12e6c6d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:02 +0000 Subject: [PATCH 059/173] feat(topology/separation): add `disjoint_nhds_nhds` (#15635) Prove that a topological space is a Hausdorff space if and only iff the neighborhood filters of distinct points are disjoint. Use existing API about filters to golf some proofs. --- src/order/filter/bases.lean | 2 +- src/topology/separation.lean | 84 +++++++++++------------------------- 2 files changed, 25 insertions(+), 61 deletions(-) diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 6feb44cc4f7bf..f3c0e164f2352 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -555,7 +555,7 @@ lemma has_basis.inf_principal_ne_bot_iff (hl : l.has_basis p s) {t : set α} : ne_bot (l ⊓ 𝓟 t) ↔ ∀ ⦃i⦄ (hi : p i), (s i ∩ t).nonempty := (hl.inf_principal t).ne_bot_iff -lemma has_basis.disjoint_basis_iff (hl : l.has_basis p s) (hl' : l'.has_basis p' s') : +lemma has_basis.disjoint_iff (hl : l.has_basis p s) (hl' : l'.has_basis p' s') : disjoint l l' ↔ ∃ i (hi : p i) i' (hi' : p' i'), disjoint (s i) (s' i') := not_iff_not.mp $ by simp only [disjoint_iff, ← ne.def, ← ne_bot_iff, hl.inf_basis_ne_bot_iff hl', not_exists, bot_eq_empty, ne_empty_iff_nonempty, inf_eq_inter] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 6d6a0f96087f0..28f14db63faef 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -733,6 +733,16 @@ lemma t2_separation [t2_space α] {x y : α} (h : x ≠ y) : ∃ u v : set α, is_open u ∧ is_open v ∧ x ∈ u ∧ y ∈ v ∧ disjoint u v := t2_space.t2 x y h +lemma t2_space_iff_disjoint_nhds : t2_space α ↔ ∀ x y : α, x ≠ y → disjoint (𝓝 x) (𝓝 y) := +begin + refine (t2_space_iff α).trans (forall₃_congr $ λ x y hne, _), + simp only [(nhds_basis_opens x).disjoint_iff (nhds_basis_opens y), exists_prop, + ← exists_and_distrib_left, and.assoc, and_comm, and.left_comm] +end + +@[simp] lemma disjoint_nhds_nhds [t2_space α] {x y : α} : disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := +⟨λ hd he, by simpa [he, nhds_ne_bot.ne] using hd, t2_space_iff_disjoint_nhds.mp ‹_› x y⟩ + /-- A finite set can be separated by open sets. -/ lemma t2_separation_finset [t2_space α] (s : finset α) : ∃ f : α → set α, set.pairwise_disjoint ↑s f ∧ ∀ x ∈ s, x ∈ f x ∧ is_open (f x) := @@ -766,37 +776,18 @@ end @[priority 100] -- see Note [lower instance priority] instance t2_space.t1_space [t2_space α] : t1_space α := -⟨λ x, is_open_compl_iff.1 $ is_open_iff_forall_mem_open.2 $ λ y hxy, -let ⟨u, v, hu, hv, hyu, hxv, huv⟩ := t2_separation (mt mem_singleton_of_eq hxy) in -⟨u, λ z hz1, huv.ne_of_mem hz1 hxv, hu, hyu⟩⟩ - -lemma eq_of_nhds_ne_bot [ht : t2_space α] {x y : α} (h : ne_bot (𝓝 x ⊓ 𝓝 y)) : x = y := -classical.by_contradiction $ assume : x ≠ y, -let ⟨u, v, hu, hv, hx, hy, huv⟩ := t2_space.t2 x y this in -(inf_ne_bot_iff.1 h (is_open.mem_nhds hu hx) (is_open.mem_nhds hv hy)).not_disjoint huv +t1_space_iff_disjoint_pure_nhds.mpr $ λ x y hne, (disjoint_nhds_nhds.2 hne).mono_left $ + pure_le_nhds _ /-- A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -/ lemma t2_iff_nhds : t2_space α ↔ ∀ {x y : α}, ne_bot (𝓝 x ⊓ 𝓝 y) → x = y := -⟨assume h, by exactI λ x y, eq_of_nhds_ne_bot, - assume h, ⟨assume x y xy, - have 𝓝 x ⊓ 𝓝 y = ⊥ := not_ne_bot.1 $ mt h xy, - let ⟨u', hu', v', hv', u'v'⟩ := empty_mem_iff_bot.mpr this, - ⟨u, uu', uo, hu⟩ := mem_nhds_iff.mp hu', - ⟨v, vv', vo, hv⟩ := mem_nhds_iff.mp hv' in - ⟨u, v, uo, vo, hu, hv, (disjoint_iff_inter_eq_empty.2 u'v'.symm).mono uu' vv'⟩⟩⟩ +by simp only [t2_space_iff_disjoint_nhds, disjoint_iff, ne_bot_iff, ne.def, not_imp_comm] + +lemma eq_of_nhds_ne_bot [t2_space α] {x y : α} (h : ne_bot (𝓝 x ⊓ 𝓝 y)) : x = y := +t2_iff_nhds.mp ‹_› h lemma t2_space_iff_nhds : t2_space α ↔ ∀ {x y : α}, x ≠ y → ∃ (U ∈ 𝓝 x) (V ∈ 𝓝 y), disjoint U V := -begin - split, - { rintro ⟨h⟩ x y hxy, - rcases h x y hxy with ⟨u, v, u_op, v_op, hx, hy, H⟩, - exact ⟨u, u_op.mem_nhds hx, v, v_op.mem_nhds hy, H⟩ }, - { refine λ h, ⟨λ x y hxy, _⟩, - rcases h hxy with ⟨u, u_in, v, v_in, H⟩, - rcases mem_nhds_iff.mp u_in with ⟨U, hUu, U_op, hxU⟩, - rcases mem_nhds_iff.mp v_in with ⟨V, hVv, V_op, hyV⟩, - exact ⟨U, V, U_op, V_op, hxU, hyV, H.mono hUu hVv⟩ } -end +by simp only [t2_space_iff_disjoint_nhds, filter.disjoint_iff] lemma t2_separation_nhds [t2_space α] {x y : α} (h : x ≠ y) : ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ disjoint u v := @@ -805,46 +796,19 @@ let ⟨u, v, open_u, open_v, x_in, y_in, huv⟩ := t2_separation h in lemma t2_separation_compact_nhds [locally_compact_space α] [t2_space α] {x y : α} (h : x ≠ y) : ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ is_compact u ∧ is_compact v ∧ disjoint u v := -begin - obtain ⟨u₀, v₀, u₀_in, v₀_in, hu₀v₀⟩ := t2_separation_nhds h, - obtain ⟨K₀, K₀_in, K₀_u₀, hK₀⟩ := local_compact_nhds u₀_in, - obtain ⟨L₀, L₀_in, L₀_v₀, hL₀⟩ := local_compact_nhds v₀_in, - exact ⟨K₀, L₀, K₀_in, L₀_in, hK₀, hL₀, hu₀v₀.mono K₀_u₀ L₀_v₀⟩, -end +by simpa only [exists_prop, ← exists_and_distrib_left, and_comm, and.assoc, and.left_comm] + using ((compact_basis_nhds x).disjoint_iff (compact_basis_nhds y)).1 (disjoint_nhds_nhds.2 h) lemma t2_iff_ultrafilter : t2_space α ↔ ∀ {x y : α} (f : ultrafilter α), ↑f ≤ 𝓝 x → ↑f ≤ 𝓝 y → x = y := t2_iff_nhds.trans $ by simp only [←exists_ultrafilter_iff, and_imp, le_inf_iff, exists_imp_distrib] -lemma is_closed_diagonal [t2_space α] : is_closed (diagonal α) := -begin - refine is_closed_iff_cluster_pt.mpr _, - rintro ⟨a₁, a₂⟩ h, - refine eq_of_nhds_ne_bot ⟨λ this : 𝓝 a₁ ⊓ 𝓝 a₂ = ⊥, h.ne _⟩, - obtain ⟨t₁, (ht₁ : t₁ ∈ 𝓝 a₁), t₂, (ht₂ : t₂ ∈ 𝓝 a₂), (h' : t₁ ∩ t₂ = ∅)⟩ := - inf_eq_bot_iff.1 this, - rw [inf_principal_eq_bot, nhds_prod_eq], - apply mem_of_superset (prod_mem_prod ht₁ ht₂), - rintro ⟨x, y⟩ ⟨x_in, y_in⟩ (heq : x = y), - rw ← heq at *, - have : x ∈ t₁ ∩ t₂ := ⟨x_in, y_in⟩, - rwa h' at this -end - lemma t2_iff_is_closed_diagonal : t2_space α ↔ is_closed (diagonal α) := -begin - split, - { introI h, - exact is_closed_diagonal }, - { intro h, - constructor, - intros x y hxy, - have : (x, y) ∈ (diagonal α)ᶜ, by rwa [mem_compl_iff], - obtain ⟨t, t_sub, t_op, xyt⟩ : ∃ t ⊆ (diagonal α)ᶜ, is_open t ∧ (x, y) ∈ t := - is_open_iff_forall_mem_open.mp h.is_open_compl _ this, - rcases is_open_prod_iff.mp t_op x y xyt with ⟨U, V, U_op, V_op, xU, yV, H⟩, - exact ⟨U, V, U_op, V_op, xU, yV, prod_subset_compl_diagonal_iff_disjoint.1 (H.trans t_sub)⟩ } -end +by simp only [t2_space_iff_disjoint_nhds, ← is_open_compl_iff, is_open_iff_mem_nhds, prod.forall, + nhds_prod_eq, compl_diagonal_mem_prod, mem_compl_iff, mem_diagonal_iff] + +lemma is_closed_diagonal [t2_space α] : is_closed (diagonal α) := +t2_iff_is_closed_diagonal.mp ‹_› section separated From f1c31ef8a592db5bb4a034404b20769e9f3b50fd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:03 +0000 Subject: [PATCH 060/173] chore(topology/locally_finite): move from `topology.basic` (#15640) Create a new file about `locally_finite`, move the definition and some lemmas from `topology.basic`. --- src/topology/basic.lean | 150 ---------------------------- src/topology/constructions.lean | 1 + src/topology/locally_finite.lean | 164 +++++++++++++++++++++++++++++++ 3 files changed, 165 insertions(+), 150 deletions(-) create mode 100644 src/topology/locally_finite.lean diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 5539c950a687b..aeddd6be515f6 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1175,151 +1175,6 @@ le_nhds_Lim h end lim -/-! -### Locally finite families --/ - -/- locally finite family [General Topology (Bourbaki, 1995)] -/ -section locally_finite - -/-- A family of sets in `set α` is locally finite if at every point `x:α`, - there is a neighborhood of `x` which meets only finitely many sets in the family -/ -def locally_finite (f : β → set α) := -∀x:α, ∃t ∈ 𝓝 x, {i | (f i ∩ t).nonempty}.finite - -lemma locally_finite.point_finite {f : β → set α} (hf : locally_finite f) (x : α) : - {b | x ∈ f b}.finite := -let ⟨t, hxt, ht⟩ := hf x in ht.subset $ λ b hb, ⟨x, hb, mem_of_mem_nhds hxt⟩ - -lemma locally_finite_of_finite [finite β] (f : β → set α) : locally_finite f := -assume x, ⟨univ, univ_mem, to_finite _⟩ - -lemma locally_finite.subset - {f₁ f₂ : β → set α} (hf₂ : locally_finite f₂) (hf : ∀b, f₁ b ⊆ f₂ b) : locally_finite f₁ := -assume a, -let ⟨t, ht₁, ht₂⟩ := hf₂ a in -⟨t, ht₁, ht₂.subset $ assume i hi, hi.mono $ inter_subset_inter (hf i) $ subset.refl _⟩ - -lemma locally_finite.comp_inj_on {ι} {f : β → set α} {g : ι → β} (hf : locally_finite f) - (hg : inj_on g {i | (f (g i)).nonempty}) : locally_finite (f ∘ g) := -λ x, let ⟨t, htx, htf⟩ := hf x in ⟨t, htx, htf.preimage $ hg.mono $ λ i hi, - hi.out.mono $ inter_subset_left _ _⟩ - -lemma locally_finite.comp_injective {ι} {f : β → set α} {g : ι → β} (hf : locally_finite f) - (hg : function.injective g) : locally_finite (f ∘ g) := -hf.comp_inj_on (hg.inj_on _) - -lemma locally_finite.eventually_finite {f : β → set α} (hf : locally_finite f) (x : α) : - ∀ᶠ s in (𝓝 x).small_sets, {i | (f i ∩ s).nonempty}.finite := -eventually_small_sets.2 $ let ⟨s, hsx, hs⟩ := hf x in - ⟨s, hsx, λ t hts, hs.subset $ λ i hi, hi.out.mono $ inter_subset_inter_right _ hts⟩ - -lemma locally_finite.exists_mem_basis {f : β → set α} (hf : locally_finite f) {p : ι → Prop} - {s : ι → set α} {x : α} (hb : (𝓝 x).has_basis p s) : - ∃ i (hi : p i), {j | (f j ∩ s i).nonempty}.finite := -let ⟨i, hpi, hi⟩ := hb.small_sets.eventually_iff.mp (hf.eventually_finite x) -in ⟨i, hpi, hi subset.rfl⟩ - -lemma locally_finite.sum_elim {γ} {f : β → set α} {g : γ → set α} (hf : locally_finite f) - (hg : locally_finite g) : locally_finite (sum.elim f g) := -begin - intro x, - obtain ⟨s, hsx, hsf, hsg⟩ : - ∃ s, s ∈ 𝓝 x ∧ {i | (f i ∩ s).nonempty}.finite ∧ {j | (g j ∩ s).nonempty}.finite, - from ((𝓝 x).frequently_small_sets_mem.and_eventually - ((hf.eventually_finite x).and (hg.eventually_finite x))).exists, - refine ⟨s, hsx, _⟩, - convert (hsf.image sum.inl).union (hsg.image sum.inr) using 1, - ext (i|j); simp -end - -lemma locally_finite.closure {f : β → set α} (hf : locally_finite f) : - locally_finite (λ i, closure (f i)) := -begin - intro x, - rcases hf x with ⟨s, hsx, hsf⟩, - refine ⟨interior s, interior_mem_nhds.2 hsx, hsf.subset $ λ i hi, _⟩, - exact (hi.mono (closure_inter_open' is_open_interior)).of_closure.mono - (inter_subset_inter_right _ interior_subset) -end - -lemma locally_finite.is_closed_Union {f : β → set α} - (h₁ : locally_finite f) (h₂ : ∀i, is_closed (f i)) : is_closed (⋃i, 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, (h₂ i).is_open_compl.mem_nhds (ha i), - rcases h₁ 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_eq, mem_Inter], - rintros b ⟨hbt, hn⟩ i hfb, - exact hn i ⟨b, hfb, hbt⟩ hfb, -end - -lemma locally_finite.closure_Union {f : β → set α} (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 _ _) - -/-- 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`. -/ -lemma locally_finite.Inter_compl_mem_nhds {f : β → set α} (hf : locally_finite f) - (hc : ∀ i, is_closed (f i)) (x : α) : (⋂ i (hi : x ∉ f i), (f i)ᶜ) ∈ 𝓝 x := -begin - refine is_open.mem_nhds _ (mem_Inter₂.2 $ λ i, id), - suffices : is_closed (⋃ i : {i // x ∉ f i}, f i), - by rwa [← is_open_compl_iff, compl_Union, Inter_subtype] at this, - exact (hf.comp_injective subtype.coe_injective).is_closed_Union (λ i, hc _) -end - -/-- Let `f : ℕ → Π a, β a` be a sequence of (dependent) functions on a topological space. Suppose -that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a -function `F : Π a, β a` such that for any `x`, we have `f n x = F x` on the product of an infinite -interval `[N, +∞)` and a neighbourhood of `x`. - -We formulate the conclusion in terms of the product of filter `filter.at_top` and `𝓝 x`. -/ -lemma locally_finite.exists_forall_eventually_eq_prod {β : α → Sort*} {f : ℕ → Π a : α, β a} - (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : - ∃ F : Π a : α, β a, ∀ x, ∀ᶠ p : ℕ × α in at_top ×ᶠ 𝓝 x, f p.1 p.2 = F p.2 := -begin - choose U hUx hU using hf, - choose N hN using λ x, (hU x).bdd_above, - replace hN : ∀ x (n > N x) (y ∈ U x), f (n + 1) y = f n y, - from λ x n hn y hy, by_contra (λ hne, hn.lt.not_le $ hN x ⟨y, hne, hy⟩), - replace hN : ∀ x (n ≥ N x + 1) (y ∈ U x), f n y = f (N x + 1) y, - from λ x n hn y hy, nat.le_induction rfl (λ k hle, (hN x _ hle _ hy).trans) n hn, - refine ⟨λ x, f (N x + 1) x, λ x, _⟩, - filter_upwards [filter.prod_mem_prod (eventually_gt_at_top (N x)) (hUx x)], - rintro ⟨n, y⟩ ⟨hn : N x < n, hy : y ∈ U x⟩, - calc f n y = f (N x + 1) y : hN _ _ hn _ hy - ... = f (max (N x + 1) (N y + 1)) y : (hN _ _ (le_max_left _ _) _ hy).symm - ... = f (N y + 1) y : hN _ _ (le_max_right _ _) _ (mem_of_mem_nhds $ hUx y) -end - -/-- Let `f : ℕ → Π a, β a` be a sequence of (dependent) functions on a topological space. Suppose -that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a -function `F : Π a, β a` such that for any `x`, for sufficiently large values of `n`, we have -`f n y = F y` in a neighbourhood of `x`. -/ -lemma locally_finite.exists_forall_eventually_at_top_eventually_eq' {β : α → Sort*} - {f : ℕ → Π a : α, β a} (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : - ∃ F : Π a : α, β a, ∀ x, ∀ᶠ n : ℕ in at_top, ∀ᶠ y : α in 𝓝 x, f n y = F y := -hf.exists_forall_eventually_eq_prod.imp $ λ F hF x, (hF x).curry - -/-- Let `f : ℕ → α → β` be a sequence of functions on a topological space. Suppose -that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a -function `F : α → β` such that for any `x`, for sufficiently large values of `n`, we have -`f n =ᶠ[𝓝 x] F`. -/ -lemma locally_finite.exists_forall_eventually_at_top_eventually_eq - {f : ℕ → α → β} (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : - ∃ F : α → β, ∀ x, ∀ᶠ n : ℕ in at_top, f n =ᶠ[𝓝 x] F := -hf.exists_forall_eventually_at_top_eventually_eq' - -end locally_finite - end topological_space /-! @@ -1446,11 +1301,6 @@ nat.rec_on n continuous_at_id $ λ n ihn, show continuous_at (f^[n] ∘ f) x, from continuous_at.comp (hx.symm ▸ ihn) hf -lemma locally_finite.preimage_continuous {ι} {f : ι → set α} {g : β → α} (hf : locally_finite f) - (hg : continuous g) : locally_finite (λ i, g ⁻¹' (f i)) := -λ x, let ⟨s, hsx, hs⟩ := hf (g x) - in ⟨g ⁻¹' s, hg.continuous_at hsx, hs.subset $ λ i ⟨y, hy⟩, ⟨g y, hy⟩⟩ - lemma continuous_iff_is_closed {f : α → β} : continuous f ↔ (∀s, is_closed s → is_closed (f ⁻¹' s)) := ⟨assume hf s hs, by simpa using (continuous_def.1 hf sᶜ hs.is_open_compl).is_closed_compl, diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 7164c8bcc540a..8292f143f06a4 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import topology.maps +import topology.locally_finite import order.filter.pi import data.fin.tuple diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean new file mode 100644 index 0000000000000..02bbec2d1364b --- /dev/null +++ b/src/topology/locally_finite.lean @@ -0,0 +1,164 @@ +/- +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 + +/-! +### Locally finite families of sets + +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. + +In this file we give the definition and prove basic properties of locally finite families of sets. +-/ + +/- locally finite family [General Topology (Bourbaki, 1995)] -/ + +open set function filter +open_locale topological_space filter + +variables {ι ι' α X Y : Type*} [topological_space X] [topological_space Y] + {f g : ι → set X} + +/-- A family of sets in `set X` is locally finite if at every point `x : X`, +there is a neighborhood of `x` which meets only finitely many sets in the family. -/ +def locally_finite (f : ι → set X) := +∀ x : X, ∃t ∈ 𝓝 x, {i | (f i ∩ t).nonempty}.finite + +lemma locally_finite_of_finite [finite ι] (f : ι → set X) : locally_finite f := +assume x, ⟨univ, univ_mem, to_finite _⟩ + +namespace locally_finite + +lemma point_finite (hf : locally_finite f) (x : X) : {b | x ∈ f b}.finite := +let ⟨t, hxt, ht⟩ := hf x in ht.subset $ λ b hb, ⟨x, hb, mem_of_mem_nhds hxt⟩ + +protected lemma subset (hf : locally_finite f) (hg : ∀ i, g i ⊆ f i) : locally_finite g := +assume a, +let ⟨t, ht₁, ht₂⟩ := hf a in +⟨t, ht₁, ht₂.subset $ assume i hi, hi.mono $ inter_subset_inter (hg i) subset.rfl⟩ + +lemma comp_inj_on {g : ι' → ι} (hf : locally_finite f) + (hg : inj_on g {i | (f (g i)).nonempty}) : locally_finite (f ∘ g) := +λ x, let ⟨t, htx, htf⟩ := hf x in ⟨t, htx, htf.preimage $ hg.mono $ λ i hi, + hi.out.mono $ inter_subset_left _ _⟩ + +lemma comp_injective {g : ι' → ι} (hf : locally_finite f) + (hg : function.injective g) : locally_finite (f ∘ g) := +hf.comp_inj_on (hg.inj_on _) + +lemma eventually_finite (hf : locally_finite f) (x : X) : + ∀ᶠ s in (𝓝 x).small_sets, {i | (f i ∩ s).nonempty}.finite := +eventually_small_sets.2 $ let ⟨s, hsx, hs⟩ := hf x in + ⟨s, hsx, λ t hts, hs.subset $ λ i hi, hi.out.mono $ inter_subset_inter_right _ hts⟩ + +lemma exists_mem_basis {ι' : Sort*} (hf : locally_finite f) {p : ι' → Prop} + {s : ι' → set X} {x : X} (hb : (𝓝 x).has_basis p s) : + ∃ i (hi : p i), {j | (f j ∩ s i).nonempty}.finite := +let ⟨i, hpi, hi⟩ := hb.small_sets.eventually_iff.mp (hf.eventually_finite x) +in ⟨i, hpi, hi subset.rfl⟩ + +lemma sum_elim {g : ι' → set X} (hf : locally_finite f) (hg : locally_finite g) : + locally_finite (sum.elim f g) := +begin + intro x, + obtain ⟨s, hsx, hsf, hsg⟩ : + ∃ s, s ∈ 𝓝 x ∧ {i | (f i ∩ s).nonempty}.finite ∧ {j | (g j ∩ s).nonempty}.finite, + from ((𝓝 x).frequently_small_sets_mem.and_eventually + ((hf.eventually_finite x).and (hg.eventually_finite x))).exists, + refine ⟨s, hsx, _⟩, + convert (hsf.image sum.inl).union (hsg.image sum.inr) using 1, + ext (i|j); simp +end + +protected lemma closure (hf : locally_finite f) : locally_finite (λ i, closure (f i)) := +begin + intro x, + rcases hf x with ⟨s, hsx, hsf⟩, + refine ⟨interior s, interior_mem_nhds.2 hsx, hsf.subset $ λ i hi, _⟩, + exact (hi.mono (closure_inter_open' is_open_interior)).of_closure.mono + (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) := +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_eq, mem_Inter], + rintros b ⟨hbt, hn⟩ i hfb, + exact hn i ⟨b, hfb, hbt⟩ hfb, +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 _ _) + +/-- 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`. -/ +lemma Inter_compl_mem_nhds (hf : locally_finite f) (hc : ∀ i, is_closed (f i)) (x : X) : + (⋂ i (hi : x ∉ f i), (f i)ᶜ) ∈ 𝓝 x := +begin + refine is_open.mem_nhds _ (mem_Inter₂.2 $ λ i, id), + suffices : is_closed (⋃ i : {i // x ∉ f i}, f i), + by rwa [← is_open_compl_iff, compl_Union, Inter_subtype] at this, + exact (hf.comp_injective subtype.coe_injective).is_closed_Union (λ i, hc _) +end + +/-- Let `f : ℕ → Π a, β a` be a sequence of (dependent) functions on a topological space. Suppose +that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a +function `F : Π a, β a` such that for any `x`, we have `f n x = F x` on the product of an infinite +interval `[N, +∞)` and a neighbourhood of `x`. + +We formulate the conclusion in terms of the product of filter `filter.at_top` and `𝓝 x`. -/ +lemma exists_forall_eventually_eq_prod {π : X → Sort*} {f : ℕ → Π x : X, π x} + (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : + ∃ F : Π x : X, π x, ∀ x, ∀ᶠ p : ℕ × X in at_top ×ᶠ 𝓝 x, f p.1 p.2 = F p.2 := +begin + choose U hUx hU using hf, + choose N hN using λ x, (hU x).bdd_above, + replace hN : ∀ x (n > N x) (y ∈ U x), f (n + 1) y = f n y, + from λ x n hn y hy, by_contra (λ hne, hn.lt.not_le $ hN x ⟨y, hne, hy⟩), + replace hN : ∀ x (n ≥ N x + 1) (y ∈ U x), f n y = f (N x + 1) y, + from λ x n hn y hy, nat.le_induction rfl (λ k hle, (hN x _ hle _ hy).trans) n hn, + refine ⟨λ x, f (N x + 1) x, λ x, _⟩, + filter_upwards [filter.prod_mem_prod (eventually_gt_at_top (N x)) (hUx x)], + rintro ⟨n, y⟩ ⟨hn : N x < n, hy : y ∈ U x⟩, + calc f n y = f (N x + 1) y : hN _ _ hn _ hy + ... = f (max (N x + 1) (N y + 1)) y : (hN _ _ (le_max_left _ _) _ hy).symm + ... = f (N y + 1) y : hN _ _ (le_max_right _ _) _ (mem_of_mem_nhds $ hUx y) +end + +/-- Let `f : ℕ → Π a, β a` be a sequence of (dependent) functions on a topological space. Suppose +that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a +function `F : Π a, β a` such that for any `x`, for sufficiently large values of `n`, we have +`f n y = F y` in a neighbourhood of `x`. -/ +lemma exists_forall_eventually_at_top_eventually_eq' {π : X → Sort*} + {f : ℕ → Π x : X, π x} (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : + ∃ F : Π x : X, π x, ∀ x, ∀ᶠ n : ℕ in at_top, ∀ᶠ y : X in 𝓝 x, f n y = F y := +hf.exists_forall_eventually_eq_prod.imp $ λ F hF x, (hF x).curry + +/-- Let `f : ℕ → α → β` be a sequence of functions on a topological space. Suppose +that the family of sets `s n = {x | f (n + 1) x ≠ f n x}` is locally finite. Then there exists a +function `F : α → β` such that for any `x`, for sufficiently large values of `n`, we have +`f n =ᶠ[𝓝 x] F`. -/ +lemma exists_forall_eventually_at_top_eventually_eq {f : ℕ → X → α} + (hf : locally_finite (λ n, {x | f (n + 1) x ≠ f n x})) : + ∃ F : X → α, ∀ x, ∀ᶠ n : ℕ in at_top, f n =ᶠ[𝓝 x] F := +hf.exists_forall_eventually_at_top_eventually_eq' + +lemma preimage_continuous {g : Y → X} (hf : locally_finite f) (hg : continuous g) : + locally_finite (λ i, g ⁻¹' (f i)) := +λ x, let ⟨s, hsx, hs⟩ := hf (g x) + in ⟨g ⁻¹' s, hg.continuous_at hsx, hs.subset $ λ i ⟨y, hy⟩, ⟨g y, hy⟩⟩ + +end locally_finite From 1690a4eed435bfd937dda1b95d9eaaebdb81e96f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:04 +0000 Subject: [PATCH 061/173] feat(tooplogy/metric_space/hausdorff_distance): add lemmas about `thickening` (#15641) * rename `emetric.exists_pos_forall_le_edist` -> `emetric.exists_pos_forall_lt_edist`; - don't assume that the compact set is nonempty; - provide an `nnreal` number instead of an `ennreal`; - claim strict inequality; * add `metric.thickening_mem_nhds_set` and `metric.cthickening_mem_nhds_set`; * move `is_compact.exists_thickening_subset_open` below the `cthickening` version, make it clear from the proof that the latter implies the former; * add `metric.has_basis_nhds_set_thickening` and `metric.has_basis_nhds_set_cthickening`; * add `continuous.tendsto_nhds_set` --- .../metric_space/hausdorff_distance.lean | 60 ++++++++++++------- src/topology/nhds_set.lean | 10 +++- 2 files changed, 46 insertions(+), 24 deletions(-) diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index b0bc756d86604..1e59651bbb2e9 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -174,15 +174,16 @@ begin exact ⟨y, ys, le_antisymm (inf_edist_le_edist_of_mem ys) (by rwa le_inf_edist)⟩ end -lemma exists_pos_forall_le_edist (hs : is_compact s) (hs' : s.nonempty) (ht : is_closed t) - (hst : disjoint s t) : - ∃ r, 0 < r ∧ ∀ (x ∈ s) (y ∈ t), r ≤ edist x y := +lemma exists_pos_forall_lt_edist (hs : is_compact s) (ht : is_closed t) (hst : disjoint s t) : + ∃ r : ℝ≥0, 0 < r ∧ ∀ (x ∈ s) (y ∈ t), (r : ℝ≥0∞) < edist x y := begin + rcases s.eq_empty_or_nonempty with rfl|hne, { use 1, simp }, obtain ⟨x, hx, h⟩ : ∃ x ∈ s, ∀ y ∈ s, inf_edist x t ≤ inf_edist y t := - hs.exists_forall_le hs' continuous_inf_edist.continuous_on, - refine ⟨inf_edist x t, pos_iff_ne_zero.2 $ λ H, hst ⟨hx, _⟩, λ y hy, le_inf_edist.1 $ h y hy⟩, - rw ←ht.closure_eq, - exact mem_closure_iff_inf_edist_zero.2 H, + hs.exists_forall_le hne continuous_inf_edist.continuous_on, + have : 0 < inf_edist x t, + from pos_iff_ne_zero.2 (λ H, hst ⟨hx, (mem_iff_inf_edist_zero_of_closed ht).mpr H⟩), + rcases ennreal.lt_iff_exists_nnreal_btwn.1 this with ⟨r, h₀, hr⟩, + exact ⟨r, ennreal.coe_pos.mp h₀, λ y hy z hz, hr.trans_le $ le_inf_edist.1 (h y hy) z hz⟩ end end inf_edist --section @@ -1016,6 +1017,12 @@ lemma self_subset_cthickening {δ : ℝ} (E : set α) : E ⊆ cthickening δ E := subset_closure.trans (closure_subset_cthickening δ E) +lemma thickening_mem_nhds_set (E : set α) {δ : ℝ} (hδ : 0 < δ) : thickening δ E ∈ 𝓝ˢ E := +is_open_thickening.mem_nhds_set.2 $ self_subset_thickening hδ E + +lemma cthickening_mem_nhds_set (E : set α) {δ : ℝ} (hδ : 0 < δ) : cthickening δ E ∈ 𝓝ˢ E := +mem_of_superset (thickening_mem_nhds_set E hδ) (thickening_subset_cthickening _ _) + @[simp] lemma thickening_union (δ : ℝ) (s t : set α) : thickening δ (s ∪ t) = thickening δ s ∪ thickening δ t := by simp_rw [thickening, inf_edist_union, inf_eq_min, min_lt_iff, set_of_or] @@ -1040,20 +1047,17 @@ lemma _root_.disjoint.exists_thickenings (hst : disjoint s t) (hs : is_compact s (ht : is_closed t) : ∃ δ, 0 < δ ∧ disjoint (thickening δ s) (thickening δ t) := begin - obtain rfl | hs' := s.eq_empty_or_nonempty, - { simp_rw thickening_empty, - exact ⟨1, zero_lt_one, empty_disjoint _⟩ }, - obtain ⟨r, hr, h⟩ := exists_pos_forall_le_edist hs hs' ht hst, - refine ⟨(min 1 (r/2)).to_real, to_real_pos (lt_min ennreal.zero_lt_one $ half_pos hr.ne').ne' - (min_lt_of_left_lt one_lt_top).ne, _⟩, + obtain ⟨r, hr, h⟩ := exists_pos_forall_lt_edist hs ht hst, + refine ⟨r / 2, half_pos (nnreal.coe_pos.2 hr), _⟩, rintro z ⟨hzs, hzt⟩, rw mem_thickening_iff_exists_edist_lt at hzs hzt, + rw [← nnreal.coe_two, ← nnreal.coe_div, ennreal.of_real_coe_nnreal] at hzs hzt, obtain ⟨x, hx, hzx⟩ := hzs, obtain ⟨y, hy, hzy⟩ := hzt, - refine (((h _ hx _ hy).trans $ edist_triangle_left _ _ _).trans_lt $ - ennreal.add_lt_add hzx hzy).not_le _, - rw ←two_mul, - exact ennreal.mul_le_of_le_div' (of_real_to_real_le.trans $ min_le_right _ _), + 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] end lemma _root_.disjoint.exists_cthickenings (hst : disjoint s t) (hs : is_compact s) @@ -1065,18 +1069,28 @@ begin exact (cthickening_subset_thickening' hδ (half_lt_self hδ) _), end -lemma _root_.is_compact.exists_thickening_subset_open (hs : is_compact s) (ht : is_open t) - (hst : s ⊆ t) : - ∃ δ, 0 < δ ∧ thickening δ s ⊆ t := -(hst.disjoint_compl_right.exists_thickenings hs ht.is_closed_compl).imp $ λ δ h, - ⟨h.1, disjoint_compl_right_iff_subset.1 $ h.2.mono_right $ self_subset_thickening h.1 _⟩ - lemma _root_.is_compact.exists_cthickening_subset_open (hs : is_compact s) (ht : is_open t) (hst : s ⊆ t) : ∃ δ, 0 < δ ∧ cthickening δ s ⊆ t := (hst.disjoint_compl_right.exists_cthickenings hs ht.is_closed_compl).imp $ λ δ h, ⟨h.1, disjoint_compl_right_iff_subset.1 $ h.2.mono_right $ self_subset_cthickening _⟩ +lemma _root_.is_compact.exists_thickening_subset_open (hs : is_compact s) (ht : is_open t) + (hst : s ⊆ t) : + ∃ δ, 0 < δ ∧ thickening δ s ⊆ t := +let ⟨δ, h₀, hδ⟩ := hs.exists_cthickening_subset_open ht hst + in ⟨δ, h₀, (thickening_subset_cthickening _ _).trans hδ⟩ + +lemma has_basis_nhds_set_thickening {K : set α} (hK : is_compact K) : + (𝓝ˢ K).has_basis (λ δ : ℝ, 0 < δ) (λ δ, thickening δ K) := +(has_basis_nhds_set K).to_has_basis' (λ U hU, hK.exists_thickening_subset_open hU.1 hU.2) $ + λ _, thickening_mem_nhds_set K + +lemma has_basis_nhds_set_cthickening {K : set α} (hK : is_compact K) : + (𝓝ˢ K).has_basis (λ δ : ℝ, 0 < δ) (λ δ, cthickening δ K) := +(has_basis_nhds_set K).to_has_basis' (λ U hU, hK.exists_cthickening_subset_open hU.1 hU.2) $ + λ _, cthickening_mem_nhds_set K + lemma cthickening_eq_Inter_cthickening' {δ : ℝ} (s : set ℝ) (hsδ : s ⊆ Ioi δ) (hs : ∀ ε, δ < ε → (s ∩ (Ioc δ ε)).nonempty) (E : set α) : cthickening δ E = ⋂ ε ∈ s, cthickening ε E := diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index 9a215090a215b..23a632815070e 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -26,7 +26,8 @@ Furthermore, we have the following results: open set filter open_locale topological_space -variables {α : Type*} [topological_space α] {s t s₁ s₂ t₁ t₂ : set α} {x : α} +variables {α β : Type*} [topological_space α] [topological_space β] + {s t s₁ s₂ t₁ t₂ : set α} {x : α} /-- The filter of neighborhoods of a set in a topological space. -/ def nhds_set (s : set α) : filter α := @@ -73,3 +74,10 @@ by simp only [nhds_set, image_union, Sup_union] lemma union_mem_nhds_set (h₁ : s₁ ∈ 𝓝ˢ t₁) (h₂ : s₂ ∈ 𝓝ˢ t₂) : s₁ ∪ s₂ ∈ 𝓝ˢ (t₁ ∪ t₂) := by { rw nhds_set_union, exact union_mem_sup h₁ h₂ } + +/-- Preimage of a set neighborhood of `t` under a continuous map `f` is a set neighborhood of `s` +provided that `f` maps `s` to `t`. -/ +lemma continuous.tendsto_nhds_set {f : α → β} {t : set β} (hf : continuous f) + (hst : maps_to f s t) : tendsto f (𝓝ˢ s) (𝓝ˢ t) := +((has_basis_nhds_set s).tendsto_iff (has_basis_nhds_set t)).mpr $ λ U hU, + ⟨f ⁻¹' U, ⟨hU.1.preimage hf, hst.mono subset.rfl hU.2⟩, λ x, id⟩ From 81ab6207183d490c190adac37ee96dd0779e6474 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:05 +0000 Subject: [PATCH 062/173] feat(topology/basic): add lemmas about `filter.lift' _ closure` (#15653) Use these lemmas to golf some proofs about `uniformity`. --- src/topology/basic.lean | 14 ++++++++++++++ src/topology/uniform_space/basic.lean | 27 ++++++++------------------- 2 files changed, 22 insertions(+), 19 deletions(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index aeddd6be515f6..1ac0b45759c1b 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -449,6 +449,20 @@ theorem mem_closure_iff {s : set α} {a : α} : λ H c ⟨h₁, h₂⟩, classical.by_contradiction $ λ nc, let ⟨x, hc, hs⟩ := (H _ h₁.is_open_compl nc) in hc (h₂ hs)⟩ +lemma filter.le_lift'_closure (l : filter α) : l ≤ l.lift' closure := +le_infi₂ $ λ s hs, le_principal_iff.2 $ mem_of_superset hs subset_closure + +lemma filter.has_basis.lift'_closure {l : filter α} {p : ι → Prop} {s : ι → set α} + (h : l.has_basis p s) : + (l.lift' closure).has_basis p (λ i, closure (s i)) := +h.lift' (monotone_closure α) + +lemma filter.has_basis.lift'_closure_eq_self {l : filter α} {p : ι → Prop} {s : ι → set α} + (h : l.has_basis p s) (hc : ∀ i, p i → is_closed (s i)) : + l.lift' closure = l := +le_antisymm (h.ge_iff.2 $ λ i hi, (hc i hi).closure_eq ▸ mem_lift' (h.mem_of_mem hi)) + l.le_lift'_closure + /-- A set is dense in a topological space if every point belongs to its closure. -/ def dense (s : set α) : Prop := ∀ x, x ∈ closure s diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index ae28ce297ba17..c8cd86da18aa6 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -786,20 +786,16 @@ begin exact ⟨w_in, w_symm⟩ end +lemma uniformity_eq_uniformity_closure : 𝓤 α = (𝓤 α).lift' closure := +eq.symm $ uniformity_has_basis_closed.lift'_closure_eq_self $ λ _, and.right + +lemma filter.has_basis.uniformity_closure {p : ι → Prop} {U : ι → set (α × α)} + (h : (𝓤 α).has_basis p U) : (𝓤 α).has_basis p (λ i, closure (U i)) := +(@uniformity_eq_uniformity_closure α _).symm ▸ h.lift'_closure + /-- Closed entourages form a basis of the uniformity filter. -/ lemma uniformity_has_basis_closure : has_basis (𝓤 α) (λ V : set (α × α), V ∈ 𝓤 α) closure := -⟨begin - intro t, - rw uniformity_has_basis_closed.mem_iff, - split, - { rintros ⟨r, ⟨r_in, r_closed⟩, r_sub⟩, - use [r, r_in], - convert r_sub, - rw r_closed.closure_eq, - refl }, - { rintros ⟨r, r_in, r_sub⟩, - exact ⟨closure r, ⟨mem_of_superset r_in subset_closure, is_closed_closure⟩, r_sub⟩ } -end⟩ +(𝓤 α).basis_sets.uniformity_closure lemma closure_eq_inter_uniformity {t : set (α×α)} : closure t = (⋂ d ∈ 𝓤 α, d ○ (t ○ d)) := @@ -829,13 +825,6 @@ calc (a, b) ∈ closure t ↔ (𝓝 (a, b) ⊓ 𝓟 t ≠ ⊥) : mem_closure_iff assume ⟨x, hx, y, hxyt, hy⟩, ⟨⟨x, y⟩, ⟨⟨hx, hy⟩, hxyt⟩⟩⟩ ... ↔ _ : by simp -lemma uniformity_eq_uniformity_closure : 𝓤 α = (𝓤 α).lift' closure := -le_antisymm - (le_infi $ assume s, le_infi $ assume hs, by simp; filter_upwards [hs] using subset_closure) - (calc (𝓤 α).lift' closure ≤ (𝓤 α).lift' (λd, d ○ (d ○ d)) : - lift'_mono' (by intros s hs; rw [closure_eq_inter_uniformity]; exact bInter_subset_of_mem hs) - ... ≤ (𝓤 α) : comp_le_uniformity3) - lemma uniformity_eq_uniformity_interior : 𝓤 α = (𝓤 α).lift' interior := le_antisymm (le_infi $ assume d, le_infi $ assume hd, From 7238f9380d5c546fa9f84eadf847358a1eb5d2bb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 11:31:06 +0000 Subject: [PATCH 063/173] feat(order/filter/basic): add `filter.has_basis.bInter_mem` (#15661) Use it to golf a few proofs. --- src/order/filter/bases.lean | 18 ++++++---- src/topology/uniform_space/basic.lean | 50 +++++++-------------------- 2 files changed, 25 insertions(+), 43 deletions(-) diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index f3c0e164f2352..48f550d9bfa7c 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -698,12 +698,18 @@ begin exact ⟨λ h i hi, h (s i) i hi subset.rfl, λ h t i hi ht, ht (h i hi)⟩ end -lemma has_basis.sInter_sets (h : has_basis l p s) : - ⋂₀ l.sets = ⋂ i (hi : p i), s i := -begin - ext x, - simp only [mem_Inter, mem_sInter, filter.mem_sets, h.forall_mem_mem], -end +protected lemma has_basis.binfi_mem [complete_lattice β] {f : set α → β} (h : has_basis l p s) + (hf : monotone f) : + (⨅ t ∈ l, f t) = ⨅ i (hi : p i), f (s i) := +le_antisymm (le_infi₂ $ λ i hi, infi₂_le (s i) (h.mem_of_mem hi)) $ + le_infi₂ $ λ t ht, let ⟨i, hpi, hi⟩ := h.mem_iff.1 ht in infi₂_le_of_le i hpi (hf hi) + +protected lemma has_basis.bInter_mem {f : set α → set β} (h : has_basis l p s) (hf : monotone f) : + (⋂ t ∈ l, f t) = ⋂ i (hi : p i), f (s i) := +h.binfi_mem hf + +lemma has_basis.sInter_sets (h : has_basis l p s) : ⋂₀ l.sets = ⋂ i (hi : p i), s i := +by { rw [sInter_eq_bInter], exact h.bInter_mem monotone_id } variables {ι'' : Type*} [preorder ι''] (l) (s'' : ι'' → set α) diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index c8cd86da18aa6..878b9a8bf7d08 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -428,6 +428,12 @@ begin exact image_mem_map h, end +/-- Symmetric entourages form a basis of `𝓤 α` -/ +lemma uniform_space.has_basis_symmetric : + (𝓤 α).has_basis (λ s : set (α × α), s ∈ 𝓤 α ∧ symmetric_rel s) id := +has_basis_self.2 $ λ t t_in, ⟨symmetrize_rel t, symmetrize_mem_uniformity t_in, + symmetric_symmetrize_rel t, symmetrize_rel_subset_self t⟩ + theorem uniformity_lift_le_swap {g : set (α×α) → filter β} {f : filter β} (hg : monotone g) (h : (𝓤 α).lift (λs, g (preimage prod.swap s)) ≤ f) : (𝓤 α).lift g ≤ f := calc (𝓤 α).lift g ≤ (filter.map (@prod.swap α α) $ 𝓤 α).lift g : @@ -766,12 +772,9 @@ lemma closure_eq_uniformity (s : set $ α × α) : closure s = ⋂ V ∈ {V | V ∈ 𝓤 α ∧ symmetric_rel V}, V ○ s ○ V := begin ext ⟨x, y⟩, - simp_rw [mem_closure_iff_nhds_basis (uniform_space.has_basis_nhds_prod x y), - mem_Inter, mem_set_of_eq], - refine forall₂_congr (λ V, _), - rintros ⟨V_in, V_symm⟩, - simp_rw [mem_comp_comp V_symm, inter_comm, exists_prop], - exact iff.rfl, + simp only [mem_closure_iff_nhds_basis (uniform_space.has_basis_nhds_prod x y), mem_Inter, + mem_set_of_eq, and_imp, mem_comp_comp, exists_prop, ← mem_inter_eq, inter_comm, set.nonempty] + { contextual := tt } end lemma uniformity_has_basis_closed : has_basis (𝓤 α) (λ V : set (α × α), V ∈ 𝓤 α ∧ is_closed V) id := @@ -799,31 +802,10 @@ lemma uniformity_has_basis_closure : has_basis (𝓤 α) (λ V : set (α × α), lemma closure_eq_inter_uniformity {t : set (α×α)} : closure t = (⋂ d ∈ 𝓤 α, d ○ (t ○ d)) := -set.ext $ assume ⟨a, b⟩, -calc (a, b) ∈ closure t ↔ (𝓝 (a, b) ⊓ 𝓟 t ≠ ⊥) : mem_closure_iff_nhds_ne_bot - ... ↔ (((@prod.swap α α) <$> 𝓤 α).lift' - (λ (s : set (α × α)), {x : α | (x, a) ∈ s} ×ˢ {y : α | (b, y) ∈ s}) ⊓ 𝓟 t ≠ ⊥) : - by rw [←uniformity_eq_symm, nhds_eq_uniformity_prod] - ... ↔ ((map (@prod.swap α α) (𝓤 α)).lift' - (λ (s : set (α × α)), {x : α | (x, a) ∈ s} ×ˢ {y : α | (b, y) ∈ s}) ⊓ 𝓟 t ≠ ⊥) : - by refl - ... ↔ ((𝓤 α).lift' - (λ (s : set (α × α)), {y : α | (a, y) ∈ s} ×ˢ {x : α | (x, b) ∈ s}) ⊓ 𝓟 t ≠ ⊥) : - begin - rw [map_lift'_eq2], - simp [image_swap_eq_preimage_swap, function.comp], - exact monotone_preimage.set_prod monotone_preimage - end - ... ↔ (∀s ∈ 𝓤 α, ({y : α | (a, y) ∈ s} ×ˢ {x : α | (x, b) ∈ s} ∩ t).nonempty) : - begin - rw [lift'_inf_principal_eq, ← ne_bot_iff, lift'_ne_bot_iff], - exact (monotone_preimage.set_prod monotone_preimage).inter monotone_const - end - ... ↔ (∀ s ∈ 𝓤 α, (a, b) ∈ s ○ (t ○ s)) : - forall₂_congr $ λ s hs, - ⟨assume ⟨⟨x, y⟩, ⟨⟨hx, hy⟩, hxyt⟩⟩, ⟨x, hx, y, hxyt, hy⟩, - assume ⟨x, hx, y, hxyt, hy⟩, ⟨⟨x, y⟩, ⟨⟨hx, hy⟩, hxyt⟩⟩⟩ - ... ↔ _ : by simp +calc closure t = ⋂ V (hV : V ∈ 𝓤 α ∧ symmetric_rel V), V ○ t ○ V : closure_eq_uniformity t +... = ⋂ V ∈ 𝓤 α, V ○ t ○ V : eq.symm $ uniform_space.has_basis_symmetric.bInter_mem $ + λ V₁ V₂ hV, comp_rel_mono (comp_rel_mono hV subset.rfl) hV +... = ⋂ V ∈ 𝓤 α, V ○ (t ○ V) : by simp only [comp_rel_assoc] lemma uniformity_eq_uniformity_interior : 𝓤 α = (𝓤 α).lift' interior := le_antisymm @@ -884,12 +866,6 @@ lemma filter.has_basis.mem_uniformity_iff {p : β → Prop} {s : β → set (α t ∈ 𝓤 α ↔ ∃ i (hi : p i), ∀ a b, (a, b) ∈ s i → (a, b) ∈ t := h.mem_iff.trans $ by simp only [prod.forall, subset_def] -/-- Symmetric entourages form a basis of `𝓤 α` -/ -lemma uniform_space.has_basis_symmetric : - (𝓤 α).has_basis (λ s : set (α × α), s ∈ 𝓤 α ∧ symmetric_rel s) id := -has_basis_self.2 $ λ t t_in, ⟨symmetrize_rel t, symmetrize_mem_uniformity t_in, - symmetric_symmetrize_rel t, symmetrize_rel_subset_self t⟩ - /-- Open elements `s : set (α × α)` of `𝓤 α` such that `(x, y) ∈ s ↔ (y, x) ∈ s` form a basis of `𝓤 α`. -/ lemma uniformity_has_basis_open_symmetric : From 80862a6d66f950dc7530aea3e26c4bbd87765761 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Tue, 26 Jul 2022 11:31:08 +0000 Subject: [PATCH 064/173] chore(analysis/inner_product_space): golf two proofs (#15679) Golf two proofs and move the lemmas into `inner_product_space/basic` since they now only depend on elementary facts about inner product spaces. --- src/analysis/inner_product_space/basic.lean | 14 +++++++++++++- src/analysis/inner_product_space/dual.lean | 19 ------------------- 2 files changed, 13 insertions(+), 20 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 1157808b31d9a..1b6bff179d103 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -212,7 +212,7 @@ 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] +by rw [←inner_conj_sym, abs_conj] lemma inner_neg_left {x y : F} : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } @@ -596,6 +596,18 @@ begin ring, end +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)] + +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)] + +omit 𝕜 +variable {𝕜} + /-- Parallelogram law -/ lemma parallelogram_law {x y : E} : ⟪x + y, x + y⟫ + ⟪x - y, x - y⟫ = 2 * (⟪x, x⟫ + ⟪y, y⟫) := diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index ee9a7b56a03b0..a9bb1f002feae 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -65,25 +65,6 @@ lemma innerSL_norm [nontrivial E] : ∥(innerSL : E →L⋆[𝕜] E →L[𝕜] show ∥(to_dual_map 𝕜 E).to_continuous_linear_map∥ = 1, from linear_isometry.norm_to_continuous_linear_map _ -variable (𝕜) -include 𝕜 -lemma ext_inner_left {x y : E} (h : ∀ v, ⟪v, x⟫ = ⟪v, y⟫) : x = y := -begin - apply (to_dual_map 𝕜 E).map_eq_iff.mp, - ext v, - rw [to_dual_map_apply, to_dual_map_apply, ←inner_conj_sym], - nth_rewrite_rhs 0 [←inner_conj_sym], - exact congr_arg conj (h v) -end - -lemma ext_inner_right {x y : E} (h : ∀ v, ⟪x, v⟫ = ⟪y, v⟫) : x = y := -begin - refine ext_inner_left 𝕜 (λ v, _), - rw [←inner_conj_sym], - nth_rewrite_rhs 0 [←inner_conj_sym], - exact congr_arg conj (h v) -end -omit 𝕜 variable {𝕜} lemma ext_inner_left_basis {ι : Type*} {x y : E} (b : basis ι 𝕜 E) From 600812216ac0adb0181a469b33802ffb5064e258 Mon Sep 17 00:00:00 2001 From: prakol16 Date: Tue, 26 Jul 2022 12:57:20 +0000 Subject: [PATCH 065/173] feat(algebraic_topology/fundamental_groupoid): Define simply connected spaces (#12788) Proves contractible spaces are simply connected, and that simply connected spaces are characterized by the property that any two paths between the same endpoints are homotopic. --- .../simply_connected.lean | 88 +++++++++++++++++++ 1 file changed, 88 insertions(+) create mode 100644 src/algebraic_topology/fundamental_groupoid/simply_connected.lean diff --git a/src/algebraic_topology/fundamental_groupoid/simply_connected.lean b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean new file mode 100644 index 0000000000000..e04bbe10c63b6 --- /dev/null +++ b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2022 Praneeth Kolichala. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Praneeth Kolichala +-/ +import algebraic_topology.fundamental_groupoid.induced_maps +import topology.homotopy.contractible +import category_theory.punit +import algebraic_topology.fundamental_groupoid.punit + +/-! +# Simply connected spaces +This file defines simply connected spaces. +A topological space is simply connected if its fundamental groupoid is equivalent to `unit`. + +## Main theorems + - `simply_connected_iff_unique_homotopic` - A space is simply connected if and only if it is + nonempty and there is a unique path up to homotopy between any two points + + - `simply_connected_space.of_contractible` - A contractible space is simply connected +-/ +noncomputable theory + +open category_theory +open continuous_map +open_locale continuous_map + +/-- A simply connected space is one whose fundamental groupoid is equivalent to `discrete unit` -/ +class simply_connected_space (X : Type*) [topological_space X] : Prop := +(equiv_unit [] : nonempty (fundamental_groupoid X ≌ discrete unit)) + +lemma simply_connected_def (X : Type*) [topological_space X] : + simply_connected_space X ↔ nonempty (fundamental_groupoid X ≌ discrete unit) := +⟨λ h, @simply_connected_space.equiv_unit X _ h, λ h, ⟨h⟩⟩ + +lemma simply_connected_iff_unique_homotopic (X : Type*) [topological_space X] : + simply_connected_space X ↔ (nonempty X) ∧ + ∀ (x y : X), nonempty (unique (path.homotopic.quotient x y)) := +by { rw [simply_connected_def, equiv_punit_iff_unique], refl, } + +namespace simply_connected_space +variables {X : Type*} [topological_space X] [simply_connected_space X] + +instance (x y : X) : subsingleton (path.homotopic.quotient x y) := +@unique.subsingleton _ (nonempty.some (by { rw simply_connected_iff_unique_homotopic at *, tauto })) + +local attribute [instance] path.homotopic.setoid + +@[priority 100] +instance : path_connected_space X := +let unique_homotopic := (simply_connected_iff_unique_homotopic X).mp infer_instance in +{ nonempty := unique_homotopic.1, joined := λ x y, ⟨(unique_homotopic.2 x y).some.default.out⟩, } + +/-- In a simply connected space, any two paths are homotopic -/ +lemma paths_homotopic {x y : X} (p₁ p₂ : path x y) : path.homotopic p₁ p₂ := +by simpa using @subsingleton.elim (path.homotopic.quotient x y) _ ⟦p₁⟧ ⟦p₂⟧ + +@[priority 100] +instance of_contractible (Y : Type*) [topological_space Y] [contractible_space Y] : + simply_connected_space Y := +{ equiv_unit := + let H : Top.of Y ≃ₕ Top.of unit := (contractible_space.hequiv_unit Y).some in + ⟨(fundamental_groupoid_functor.equiv_of_homotopy_equiv H).trans + fundamental_groupoid.punit_equiv_discrete_punit⟩, } + +end simply_connected_space + +local attribute [instance] path.homotopic.setoid + +/-- A space is simply connected iff it is path connected, and there is at most one path + up to homotopy between any two points. -/ +lemma simply_connected_iff_paths_homotopic {Y : Type*} [topological_space Y] : + simply_connected_space Y ↔ (path_connected_space Y) ∧ + (∀ x y : Y, subsingleton (path.homotopic.quotient x y)) := +⟨by { introI, split; apply_instance, }, +λ h, begin + casesI h, rw simply_connected_iff_unique_homotopic, + exact ⟨infer_instance, λ x y, ⟨unique_of_subsingleton ⟦path_connected_space.some_path x y⟧⟩⟩, +end⟩ + +/-- Another version of `simply_connected_iff_paths_homotopic` -/ +lemma simply_connected_iff_paths_homotopic' {Y : Type*} [topological_space Y] : + simply_connected_space Y ↔ (path_connected_space Y) ∧ + (∀ {x y : Y} (p₁ p₂ : path x y), path.homotopic p₁ p₂) := +begin + convert simply_connected_iff_paths_homotopic, + simp [path.homotopic.quotient, setoid.eq_top_iff], refl, +end From eed43447a72070662c8cbfe3ea256faafca2a5ed Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Tue, 26 Jul 2022 12:57:21 +0000 Subject: [PATCH 066/173] feat(category_theory/limits): preserves biproducts if comparison is iso (#14419) --- .../limits/preserves/shapes/biproducts.lean | 161 +++++++++++++++++- .../limits/shapes/biproducts.lean | 67 +++++++- .../preadditive/additive_functor.lean | 1 + 3 files changed, 217 insertions(+), 12 deletions(-) diff --git a/src/category_theory/limits/preserves/shapes/biproducts.lean b/src/category_theory/limits/preserves/shapes/biproducts.lean index 2092a52bfd359..9d6601853f409 100644 --- a/src/category_theory/limits/preserves/shapes/biproducts.lean +++ b/src/category_theory/limits/preserves/shapes/biproducts.lean @@ -5,6 +5,8 @@ Authors: Markus Himmel -/ import category_theory.limits.shapes.biproducts import category_theory.limits.preserves.shapes.zero +import category_theory.limits.preserves.shapes.binary_products +import category_theory.limits.preserves.shapes.products /-! # Preservation of biproducts @@ -13,7 +15,10 @@ We define the image of a (binary) bicone under a functor that preserves zero mor classes `preserves_biproduct` and `preserves_binary_biproduct`. We then * show that a functor that preserves biproducts of a two-element type preserves binary biproducts, -* give the canonical isomorphism between the image of a biproduct and the biproduct of the images, +* construct the comparison morphisms between the image of a biproduct and the biproduct of the + images and show that the biproduct is preserved if one of them is an isomorphism, +* give the canonical isomorphism between the image of a biproduct and the biproduct of the images + in case that the biproduct is preserved, * show that in a preadditive category, a functor preserves a biproduct if and only if it preserves the corresponding product if and only if it preserves the corresponding coproduct. @@ -179,8 +184,53 @@ open category_theory.limits namespace functor section bicone -variables {J : Type w₁} (F : C ⥤ D) [preserves_zero_morphisms F] (f : J → C) - [has_biproduct f] [preserves_biproduct f F] +variables {J : Type w₁} (F : C ⥤ D) (f : J → C) + [has_biproduct f] + +section +variables [has_biproduct (F.obj ∘ f)] + +/-- As for products, any functor between categories with biproducts gives rise to a morphism + `F.obj (⨁ f) ⟶ ⨁ (F.obj ∘ f)`. -/ +def biproduct_comparison : F.obj (⨁ f) ⟶ ⨁ (F.obj ∘ f) := +biproduct.lift (λ j, F.map (biproduct.π f j)) + +@[simp, reassoc] lemma biproduct_comparison_π (j : J) : + biproduct_comparison F f ≫ biproduct.π _ j = F.map (biproduct.π f j) := +biproduct.lift_π _ _ + +/-- As for coproducts, any functor between categories with biproducts gives rise to a morphism + `⨁ (F.obj ∘ f) ⟶ F.obj (⨁ f)` -/ +def biproduct_comparison' : ⨁ (F.obj ∘ f) ⟶ F.obj (⨁ f) := +biproduct.desc (λ j, F.map (biproduct.ι f j)) + +@[simp, reassoc] lemma ι_biproduct_comparison' (j : J) : + biproduct.ι _ j ≫ biproduct_comparison' F f = F.map (biproduct.ι f j) := +biproduct.ι_desc _ _ + +variables [preserves_zero_morphisms F] + +/-- The composition in the opposite direction is equal to the identity if and only if `F` preserves + the biproduct, see `preserves_biproduct_of_mono_biproduct_comparison`. -/ +@[simp, reassoc] lemma biproduct_comparison'_comp_biproduct_comparison : + biproduct_comparison' F f ≫ biproduct_comparison F f = 𝟙 (⨁ (F.obj ∘ f)) := +by { classical, ext, simp [biproduct.ι_π, ← functor.map_comp, eq_to_hom_map] } + +instance : split_epi (biproduct_comparison F f) := +⟨biproduct_comparison' F f⟩ + +@[simp] lemma section_biproduct_comparison : + section_ (biproduct_comparison F f) = biproduct_comparison' F f := rfl + +instance : split_mono (biproduct_comparison' F f) := +⟨biproduct_comparison F f⟩ + +@[simp] lemma retraction_biproduct_comparison' : + retraction (biproduct_comparison' F f) = biproduct_comparison F f := rfl + +end + +variables [preserves_zero_morphisms F] [preserves_biproduct f F] instance has_biproduct_of_preserves : has_biproduct (F.obj ∘ f) := has_biproduct.mk @@ -201,8 +251,60 @@ rfl end bicone -variables (F : C ⥤ D) [preserves_zero_morphisms F] (X Y : C) [has_binary_biproduct X Y] - [preserves_binary_biproduct X Y F] +variables (F : C ⥤ D) (X Y : C) [has_binary_biproduct X Y] + +section +variables [has_binary_biproduct (F.obj X) (F.obj Y)] + +/-- As for products, any functor between categories with binary biproducts gives rise to a + morphism `F.obj (X ⊞ Y) ⟶ F.obj X ⊞ F.obj Y`. -/ +def biprod_comparison : F.obj (X ⊞ Y) ⟶ F.obj X ⊞ F.obj Y := +biprod.lift (F.map biprod.fst) (F.map biprod.snd) + +@[simp, reassoc] lemma biprod_comparison_fst : + biprod_comparison F X Y ≫ biprod.fst = F.map biprod.fst := +biprod.lift_fst _ _ + +@[simp, reassoc] lemma biprod_comparison_snd : + biprod_comparison F X Y ≫ biprod.snd = F.map biprod.snd := +biprod.lift_snd _ _ + +/-- As for coproducts, any functor between categories with binary biproducts gives rise to a + morphism `F.obj X ⊞ F.obj Y ⟶ F.obj (X ⊞ Y)`. -/ +def biprod_comparison' : F.obj X ⊞ F.obj Y ⟶ F.obj (X ⊞ Y) := +biprod.desc (F.map biprod.inl) (F.map biprod.inr) + +@[simp, reassoc] lemma inl_biprod_comparison' : + biprod.inl ≫ biprod_comparison' F X Y = F.map biprod.inl := +biprod.inl_desc _ _ + +@[simp, reassoc] lemma inr_biprod_comparison' : + biprod.inr ≫ biprod_comparison' F X Y = F.map biprod.inr := +biprod.inr_desc _ _ + +variables [preserves_zero_morphisms F] + +/-- The composition in the opposite direction is equal to the identity if and only if `F` preserves + the biproduct, see `preserves_binary_biproduct_of_mono_biprod_comparison`. -/ +@[simp, reassoc] lemma biprod_comparison'_comp_biprod_comparison : + biprod_comparison' F X Y ≫ biprod_comparison F X Y = 𝟙 (F.obj X ⊞ F.obj Y) := +by { ext; simp [← functor.map_comp] } + +instance : split_epi (biprod_comparison F X Y) := +⟨biprod_comparison' F X Y⟩ + +@[simp] lemma section_biprod_comparison : + section_ (biprod_comparison F X Y) = biprod_comparison' F X Y := rfl + +instance : split_mono (biprod_comparison' F X Y) := +⟨biprod_comparison F X Y⟩ + +@[simp] lemma retraction_biprod_comparison' : + retraction (biprod_comparison' F X Y) = biprod_comparison F X Y := rfl + +end + +variables [preserves_zero_morphisms F] [preserves_binary_biproduct X Y F] instance has_binary_biproduct_of_preserves : has_binary_biproduct (F.obj X) (F.obj Y) := has_binary_biproduct.mk @@ -311,6 +413,30 @@ def preserves_biproduct_of_preserves_product {f : J → C} [preserves_limit (dis (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ cones.ext (iso.refl _) (by tidy) } +/-- If the (product-like) biproduct comparison for `F` and `f` is a monomorphism, then `F` + preserves the biproduct of `f`. For the converse, see `map_biproduct`. -/ +def preserves_biproduct_of_mono_biproduct_comparison {f : J → C} [has_biproduct f] + [has_biproduct (F.obj ∘ f)] [mono (biproduct_comparison F f)] : preserves_biproduct f F := +begin + have : pi_comparison F f = (F.map_iso (biproduct.iso_product f)).inv ≫ + biproduct_comparison F f ≫ (biproduct.iso_product _).hom, + { ext, convert pi_comparison_comp_π F f j.as; simp [← functor.map_comp] }, + haveI : is_iso (biproduct_comparison F f) := is_iso_of_mono_of_split_epi _, + haveI : is_iso (pi_comparison F f) := by { rw this, apply_instance }, + haveI := preserves_product.of_iso_comparison F f, + apply preserves_biproduct_of_preserves_product +end + +/-- If the (coproduct-like) biproduct comparison for `F` and `f` is an epimorphism, then `F` + preserves the biproduct of `F` and `f`. For the converse, see `map_biproduct`. -/ +def preserves_biproduct_of_epi_biproduct_comparison' {f : J → C} [has_biproduct f] + [has_biproduct (F.obj ∘ f)] [epi (biproduct_comparison' F f)] : preserves_biproduct f F := +begin + haveI : epi (section_ (biproduct_comparison F f)) := by simpa, + haveI : is_iso (biproduct_comparison F f) := is_iso.of_epi_section, + apply preserves_biproduct_of_mono_biproduct_comparison +end + /-- A functor between preadditive categories that preserves (zero morphisms and) finite products preserves finite biproducts. -/ def preserves_biproducts_of_shape_of_preserves_products_of_shape @@ -385,6 +511,31 @@ def preserves_binary_biproduct_of_preserves_binary_product {X Y : C} (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ cones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } +/-- If the (product-like) biproduct comparison for `F`, `X` and `Y` is a monomorphism, then + `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ +def preserves_binary_biproduct_of_mono_biprod_comparison {X Y : C} [has_binary_biproduct X Y] + [has_binary_biproduct (F.obj X) (F.obj Y)] [mono (biprod_comparison F X Y)] : + preserves_binary_biproduct X Y F := +begin + have : prod_comparison F X Y = (F.map_iso (biprod.iso_prod X Y)).inv ≫ + biprod_comparison F X Y ≫ (biprod.iso_prod _ _).hom := by { ext; simp [← functor.map_comp] }, + haveI : is_iso (biprod_comparison F X Y) := is_iso_of_mono_of_split_epi _, + haveI : is_iso (prod_comparison F X Y) := by { rw this, apply_instance }, + haveI := preserves_limit_pair.of_iso_prod_comparison F X Y, + apply preserves_binary_biproduct_of_preserves_binary_product +end + +/-- If the (coproduct-like) biproduct comparison for `F`, `X` and `Y` is an epimorphism, then + `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ +def preserves_binary_biproduct_of_epi_biprod_comparison' {X Y : C} [has_binary_biproduct X Y] + [has_binary_biproduct (F.obj X) (F.obj Y)] [epi (biprod_comparison' F X Y)] : + preserves_binary_biproduct X Y F := +begin + haveI : epi (section_ (biprod_comparison F X Y)) := by simpa, + haveI : is_iso (biprod_comparison F X Y) := is_iso.of_epi_section, + apply preserves_binary_biproduct_of_mono_biprod_comparison +end + /-- A functor between preadditive categories that preserves (zero morphisms and) binary products preserves binary biproducts. -/ def preserves_binary_biproducts_of_preserves_binary_products diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index 18cb405f2951e..6cec24a59572b 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -102,7 +102,9 @@ def to_cone (B : bicone F) : cone (discrete.functor F) := @[simp] lemma to_cone_X (B : bicone F) : B.to_cone.X = B.X := rfl -@[simp] lemma to_cone_π_app (B : bicone F) (j : J) : B.to_cone.π.app ⟨j⟩ = B.π j := rfl +@[simp] lemma to_cone_π_app (B : bicone F) (j : discrete J) : B.to_cone.π.app j = B.π j.as := rfl + +lemma to_cone_π_app_mk (B : bicone F) (j : J) : B.to_cone.π.app ⟨j⟩ = B.π j := rfl /-- Extract the cocone from a bicone. -/ def to_cocone (B : bicone F) : cocone (discrete.functor F) := @@ -111,7 +113,10 @@ def to_cocone (B : bicone F) : cocone (discrete.functor F) := @[simp] lemma to_cocone_X (B : bicone F) : B.to_cocone.X = B.X := rfl -@[simp] lemma to_cocone_ι_app (B : bicone F) (j : J) : B.to_cocone.ι.app ⟨j⟩ = B.ι j := rfl +@[simp] lemma to_cocone_ι_app (B : bicone F) (j : discrete J) : B.to_cocone.ι.app j = B.ι j.as := +rfl + +lemma to_cocone_ι_app_mk (B : bicone F) (j : J) : B.to_cocone.ι.app ⟨j⟩ = B.ι j := rfl /-- We can turn any limit cone over a discrete collection of objects into a bicone. -/ @[simps] @@ -384,12 +389,36 @@ is_colimit.map (biproduct.is_colimit f) (biproduct.bicone g).to_cocone (w : ∀ j, biproduct.ι f j ≫ g = biproduct.ι f j ≫ h) : g = h := (biproduct.is_colimit f).hom_ext (λ j, w j.as) +/-- The canonical isomorphism between the chosen biproduct and the chosen product. -/ +def biproduct.iso_product (f : J → C) [has_biproduct f] : ⨁ f ≅ ∏ f := +is_limit.cone_point_unique_up_to_iso (biproduct.is_limit f) (limit.is_limit _) + +@[simp] lemma biproduct.iso_product_hom {f : J → C} [has_biproduct f] : + (biproduct.iso_product f).hom = pi.lift (biproduct.π f) := +limit.hom_ext $ λ j, by simp [biproduct.iso_product] + +@[simp] lemma biproduct.iso_product_inv {f : J → C} [has_biproduct f] : + (biproduct.iso_product f).inv = biproduct.lift (pi.π f) := +biproduct.hom_ext _ _ $ λ j, by simp [iso.inv_comp_eq] + +/-- The canonical isomorphism between the chosen biproduct and the chosen coproduct. -/ +def biproduct.iso_coproduct (f : J → C) [has_biproduct f] : ⨁ f ≅ ∐ f := +is_colimit.cocone_point_unique_up_to_iso (biproduct.is_colimit f) (colimit.is_colimit _) + +@[simp] lemma biproduct.iso_coproduct_inv {f : J → C} [has_biproduct f] : + (biproduct.iso_coproduct f).inv = sigma.desc (biproduct.ι f) := +colimit.hom_ext $ λ j, by simp [biproduct.iso_coproduct] + +@[simp] lemma biproduct.iso_coproduct_hom {f : J → C} [has_biproduct f] : + (biproduct.iso_coproduct f).hom = biproduct.desc (sigma.ι f) := +biproduct.hom_ext' _ _ $ λ j, by simp [← iso.eq_comp_inv] + lemma biproduct.map_eq_map' {f g : J → C} [has_biproduct f] [has_biproduct g] (p : Π b, f b ⟶ g b) : biproduct.map p = biproduct.map' p := begin ext j j', simp only [discrete.nat_trans_app, limits.is_colimit.ι_map, limits.is_limit.map_π, category.assoc, - ←bicone.to_cone_π_app, ←biproduct.bicone_π, ←bicone.to_cocone_ι_app, ←biproduct.bicone_ι], + ←bicone.to_cone_π_app_mk, ←biproduct.bicone_π, ←bicone.to_cocone_ι_app_mk, ←biproduct.bicone_ι], simp only [biproduct.bicone_ι, biproduct.bicone_π, bicone.to_cocone_ι_app, bicone.to_cone_π_app], dsimp, rw [biproduct.ι_π_assoc, biproduct.ι_π], @@ -1068,6 +1097,30 @@ binary_fan.is_limit.hom_ext (binary_biproduct.is_limit X Y) h₀ h₁ (h₀ : biprod.inl ≫ f = biprod.inl ≫ g) (h₁ : biprod.inr ≫ f = biprod.inr ≫ g) : f = g := binary_cofan.is_colimit.hom_ext (binary_biproduct.is_colimit X Y) h₀ h₁ +/-- The canonical isomorphism between the chosen biproduct and the chosen product. -/ +def biprod.iso_prod (X Y : C) [has_binary_biproduct X Y] : X ⊞ Y ≅ X ⨯ Y := +is_limit.cone_point_unique_up_to_iso (binary_biproduct.is_limit X Y) (limit.is_limit _) + +@[simp] lemma biprod.iso_prod_hom {X Y : C} [has_binary_biproduct X Y] : + (biprod.iso_prod X Y).hom = prod.lift biprod.fst biprod.snd := +by ext; simp [biprod.iso_prod] + +@[simp] lemma biprod.iso_prod_inv {X Y : C} [has_binary_biproduct X Y] : + (biprod.iso_prod X Y).inv = biprod.lift prod.fst prod.snd := +by apply biprod.hom_ext; simp [iso.inv_comp_eq] + +/-- The canonical isomorphism between the chosen biproduct and the chosen coproduct. -/ +def biprod.iso_coprod (X Y : C) [has_binary_biproduct X Y] : X ⊞ Y ≅ X ⨿ Y := +is_colimit.cocone_point_unique_up_to_iso (binary_biproduct.is_colimit X Y) (colimit.is_colimit _) + +@[simp] lemma biprod.iso_coprod_inv {X Y : C} [has_binary_biproduct X Y] : + (biprod.iso_coprod X Y).inv = coprod.desc biprod.inl biprod.inr := +by ext; simp [biprod.iso_coprod]; refl + +@[simp] lemma biprod_iso_coprod_hom {X Y : C} [has_binary_biproduct X Y] : + (biprod.iso_coprod X Y).hom = biprod.desc coprod.inl coprod.inr := +by apply biprod.hom_ext'; simp [← iso.eq_comp_inv] + lemma biprod.map_eq_map' {W X Y Z : C} [has_binary_biproduct W X] [has_binary_biproduct Y Z] (f : W ⟶ Y) (g : X ⟶ Z) : biprod.map f g = biprod.map' f g := begin @@ -1094,19 +1147,19 @@ end instance biprod.inl_mono {X Y : C} [has_binary_biproduct X Y] : split_mono (biprod.inl : X ⟶ X ⊞ Y) := -{ retraction := biprod.desc (𝟙 X) (biprod.inr ≫ biprod.fst) } +{ retraction := biprod.fst } instance biprod.inr_mono {X Y : C} [has_binary_biproduct X Y] : split_mono (biprod.inr : Y ⟶ X ⊞ Y) := -{ retraction := biprod.desc (biprod.inl ≫ biprod.snd) (𝟙 Y)} +{ retraction := biprod.snd } instance biprod.fst_epi {X Y : C} [has_binary_biproduct X Y] : split_epi (biprod.fst : X ⊞ Y ⟶ X) := -{ section_ := biprod.lift (𝟙 X) (biprod.inl ≫ biprod.snd) } +{ section_ := biprod.inl } instance biprod.snd_epi {X Y : C} [has_binary_biproduct X Y] : split_epi (biprod.snd : X ⊞ Y ⟶ Y) := -{ section_ := biprod.lift (biprod.inr ≫ biprod.fst) (𝟙 Y) } +{ section_ := biprod.inr } @[simp,reassoc] lemma biprod.map_fst {W X Y Z : C} [has_binary_biproduct W X] [has_binary_biproduct Y Z] diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index 8bfead2488c20..78590a36a8d41 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -118,6 +118,7 @@ instance preserves_finite_biproducts_of_additive [additive F] : preserves_finite simp_rw [← F.map_id], refine congr_arg _ (hb.is_limit.hom_ext (λ j, hb.is_colimit.hom_ext (λ j', _))), cases j, cases j', + dsimp only [limits.bicone.to_cone_π_app], simp [sum_comp, comp_sum, bicone.ι_π, comp_dite, dite_comp], end } } } From 313a4436f85e3d6f0c2273ca537385845477208d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 26 Jul 2022 12:57:22 +0000 Subject: [PATCH 067/173] feat(probability/moments): mgf/cgf of a sum of independent random variables (#15140) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémy Degenne --- src/probability/independence.lean | 79 +++++++++++++++++++++++++++++++ src/probability/moments.lean | 69 +++++++++++++++++++++++++-- 2 files changed, 143 insertions(+), 5 deletions(-) diff --git a/src/probability/independence.lean b/src/probability/independence.lean index 220694ebf6dad..205b2eee2a0f1 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -630,6 +630,85 @@ begin rw [h_sets_s'_univ hi, set.univ_inter], }, end +lemma Indep_fun.indep_fun_prod [is_probability_measure μ] + {ι : Type*} {β : ι → Type*} {m : Π i, measurable_space (β i)} + {f : Π i, α → β i} (hf_Indep : Indep_fun m f μ) (hf_meas : ∀ i, measurable (f i)) + (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : + indep_fun (λ a, (f i a, f j a)) (f k) μ := +begin + classical, + have h_right : f k = (λ p : (Π j : ({k} : finset ι), β j), p ⟨k, finset.mem_singleton_self k⟩) + ∘ (λ a (j : ({k} : finset ι)), f j a) := rfl, + have h_meas_right : measurable + (λ p : (Π j : ({k} : finset ι), β j), p ⟨k, finset.mem_singleton_self k⟩), + from measurable_pi_apply ⟨k, finset.mem_singleton_self k⟩, + let s : finset ι := {i, j}, + have h_left : (λ ω, (f i ω, f j ω)) + = (λ p : (Π l : s, β l), (p ⟨i, finset.mem_insert_self i _⟩, + p ⟨j, finset.mem_insert_of_mem (finset.mem_singleton_self _)⟩)) + ∘ (λ a (j : s), f j a), + { ext1 a, + simp only [prod.mk.inj_iff], + split; refl, }, + have h_meas_left : measurable (λ p : (Π l : s, β l), (p ⟨i, finset.mem_insert_self i _⟩, + p ⟨j, finset.mem_insert_of_mem (finset.mem_singleton_self _)⟩)), + from measurable.prod (measurable_pi_apply ⟨i, finset.mem_insert_self i {j}⟩) + (measurable_pi_apply ⟨j, finset.mem_insert_of_mem (finset.mem_singleton_self j)⟩), + rw [h_left, h_right], + refine (hf_Indep.indep_fun_finset s {k} _ hf_meas).comp h_meas_left h_meas_right, + intros x hx, + simp only [finset.inf_eq_inter, finset.mem_inter, finset.mem_insert, finset.mem_singleton] at hx, + simp only [finset.bot_eq_empty, finset.not_mem_empty], + cases hx.1 with hx_eq hx_eq; rw hx_eq at hx, + { exact hik hx.2, }, + { exact hjk hx.2, }, +end + +@[to_additive] +lemma Indep_fun.mul [is_probability_measure μ] + {ι : Type*} {β : Type*} {m : measurable_space β} [has_mul β] [has_measurable_mul₂ β] + {f : ι → α → β} (hf_Indep : Indep_fun (λ _, m) f μ) (hf_meas : ∀ i, measurable (f i)) + (i j k : ι) (hik : i ≠ k) (hjk : j ≠ k) : + indep_fun (f i * f j) (f k) μ := +begin + have : indep_fun (λ ω, (f i ω, f j ω)) (f k) μ := hf_Indep.indep_fun_prod hf_meas i j k hik hjk, + change indep_fun ((λ p : β × β, p.fst * p.snd) ∘ (λ ω, (f i ω, f j ω))) (id ∘ (f k)) μ, + exact indep_fun.comp this (measurable_fst.mul measurable_snd) measurable_id, +end + +@[to_additive] +lemma Indep_fun.indep_fun_finset_prod_of_not_mem [is_probability_measure μ] + {ι : Type*} {β : Type*} {m : measurable_space β} [comm_monoid β] [has_measurable_mul₂ β] + {f : ι → α → β} (hf_Indep : Indep_fun (λ _, m) f μ) (hf_meas : ∀ i, measurable (f i)) + {s : finset ι} {i : ι} (hi : i ∉ s) : + indep_fun (∏ j in s, f j) (f i) μ := +begin + classical, + have h_right : f i = (λ p : (Π j : ({i} : finset ι), β), p ⟨i, finset.mem_singleton_self i⟩) + ∘ (λ a (j : ({i} : finset ι)), f j a) := rfl, + have h_meas_right : measurable + (λ p : (Π j : ({i} : finset ι), β), p ⟨i, finset.mem_singleton_self i⟩), + from measurable_pi_apply ⟨i, finset.mem_singleton_self i⟩, + have h_left : (∏ j in s, f j) = (λ p : (Π j : s, β), ∏ j, p j) ∘ (λ a (j : s), f j a), + { ext1 a, + simp only [function.comp_app], + have : (∏ (j : ↥s), f ↑j a) = (∏ (j : ↥s), f ↑j) a, by rw finset.prod_apply, + rw [this, finset.prod_coe_sort], }, + have h_meas_left : measurable (λ p : (Π j : s, β), ∏ j, p j), + from finset.univ.measurable_prod (λ (j : ↥s) (H : j ∈ finset.univ), measurable_pi_apply j), + rw [h_left, h_right], + exact (hf_Indep.indep_fun_finset s {i} (finset.disjoint_singleton_left.mpr hi).symm hf_meas).comp + h_meas_left h_meas_right, +end + +@[to_additive] +lemma Indep_fun.indep_fun_prod_range_succ [is_probability_measure μ] + {β : Type*} {m : measurable_space β} [comm_monoid β] [has_measurable_mul₂ β] + {f : ℕ → α → β} (hf_Indep : Indep_fun (λ _, m) f μ) (hf_meas : ∀ i, measurable (f i)) + (n : ℕ) : + indep_fun (∏ j in finset.range n, f j) (f n) μ := +hf_Indep.indep_fun_finset_prod_of_not_mem hf_meas finset.not_mem_range_self + end indep_fun end probability_theory diff --git a/src/probability/moments.lean b/src/probability/moments.lean index f397185559717..46efcf3c96c5f 100644 --- a/src/probability/moments.lean +++ b/src/probability/moments.lean @@ -43,7 +43,7 @@ open_locale big_operators measure_theory probability_theory ennreal nnreal namespace probability_theory -variables {Ω : Type*} {m : measurable_space Ω} {X : Ω → ℝ} {p : ℕ} {μ : measure Ω} +variables {Ω ι : Type*} {m : measurable_space Ω} {X : Ω → ℝ} {p : ℕ} {μ : measure Ω} include m @@ -182,16 +182,22 @@ by simp_rw [mgf, pi.neg_apply, mul_neg, neg_mul] lemma cgf_neg : cgf (-X) μ t = cgf X μ (-t) := by simp_rw [cgf, mgf_neg] +/-- This is a trivial application of `indep_fun.comp` but it will come up frequently. -/ +lemma indep_fun.exp_mul {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) (s t : ℝ) : + indep_fun (λ ω, exp (s * X ω)) (λ ω, exp (t * Y ω)) μ := +begin + have h_meas : ∀ t, measurable (λ x, exp (t * x)) := λ t, (measurable_id'.const_mul t).exp, + change indep_fun ((λ x, exp (s * x)) ∘ X) ((λ x, exp (t * x)) ∘ Y) μ, + exact indep_fun.comp h_indep (h_meas s) (h_meas t), +end + lemma indep_fun.mgf_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) (h_int_X : integrable (λ ω, exp (t * X ω)) μ) (h_int_Y : integrable (λ ω, exp (t * Y ω)) μ) : mgf (X + Y) μ t = mgf X μ t * mgf Y μ t := begin simp_rw [mgf, pi.add_apply, mul_add, exp_add], - refine indep_fun.integral_mul_of_integrable' _ h_int_X h_int_Y, - have h_meas : measurable (λ x, exp (t * x)) := (measurable_id'.const_mul t).exp, - change indep_fun ((λ x, exp (t * x)) ∘ X) ((λ x, exp (t * x)) ∘ Y) μ, - exact indep_fun.comp h_indep h_meas h_meas, + exact (h_indep.exp_mul t t).integral_mul_of_integrable' h_int_X h_int_Y, end lemma indep_fun.cgf_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) @@ -205,6 +211,59 @@ begin exact log_mul (mgf_pos' hμ h_int_X).ne' (mgf_pos' hμ h_int_Y).ne', end +lemma indep_fun.integrable_exp_mul_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) + (h_int_X : integrable (λ ω, exp (t * X ω)) μ) + (h_int_Y : integrable (λ ω, exp (t * Y ω)) μ) : + integrable (λ ω, exp (t * (X + Y) ω)) μ := +begin + simp_rw [pi.add_apply, mul_add, exp_add], + exact (h_indep.exp_mul t t).integrable_mul h_int_X h_int_Y, +end + +lemma Indep_fun.integrable_exp_mul_sum [is_probability_measure μ] + {X : ι → Ω → ℝ} (h_indep : Indep_fun (λ i, infer_instance) X μ) (h_meas : ∀ i, measurable (X i)) + {s : finset ι} (h_int : ∀ i ∈ s, integrable (λ ω, exp (t * X i ω)) μ) : + integrable (λ ω, exp (t * (∑ i in s, X i) ω)) μ := +begin + classical, + induction s using finset.induction_on with i s hi_notin_s h_rec h_int, + { simp only [pi.zero_apply, sum_apply, sum_empty, mul_zero, exp_zero], + exact integrable_const _, }, + { have : ∀ (i : ι), i ∈ s → integrable (λ (ω : Ω), exp (t * X i ω)) μ, + from λ i hi, h_int i (mem_insert_of_mem hi), + specialize h_rec this, + rw sum_insert hi_notin_s, + refine indep_fun.integrable_exp_mul_add _ (h_int i (mem_insert_self _ _)) h_rec, + exact (h_indep.indep_fun_finset_sum_of_not_mem h_meas hi_notin_s).symm, }, +end + +lemma Indep_fun.mgf_sum [is_probability_measure μ] + {X : ι → Ω → ℝ} (h_indep : Indep_fun (λ i, infer_instance) X μ) (h_meas : ∀ i, measurable (X i)) + {s : finset ι} (h_int : ∀ i ∈ s, integrable (λ ω, exp (t * X i ω)) μ) : + mgf (∑ i in s, X i) μ t = ∏ i in s, mgf (X i) μ t := +begin + classical, + induction s using finset.induction_on with i s hi_notin_s h_rec h_int, + { simp only [sum_empty, mgf_zero_fun, measure_univ, ennreal.one_to_real, prod_empty], }, + { have h_int' : ∀ (i : ι), i ∈ s → integrable (λ (ω : Ω), exp (t * X i ω)) μ, + from λ i hi, h_int i (mem_insert_of_mem hi), + rw [sum_insert hi_notin_s, indep_fun.mgf_add + (h_indep.indep_fun_finset_sum_of_not_mem h_meas hi_notin_s).symm + (h_int i (mem_insert_self _ _)) (h_indep.integrable_exp_mul_sum h_meas h_int'), + h_rec h_int', prod_insert hi_notin_s], }, +end + +lemma Indep_fun.cgf_sum [is_probability_measure μ] + {X : ι → Ω → ℝ} (h_indep : Indep_fun (λ i, infer_instance) X μ) (h_meas : ∀ i, measurable (X i)) + {s : finset ι} (h_int : ∀ i ∈ s, integrable (λ ω, exp (t * X i ω)) μ) : + cgf (∑ i in s, X i) μ t = ∑ i in s, cgf (X i) μ t := +begin + simp_rw cgf, + rw ← log_prod _ _ (λ j hj, _), + { rw h_indep.mgf_sum h_meas h_int, }, + { exact (mgf_pos (h_int j hj)).ne', }, +end + /-- **Chernoff bound** on the upper tail of a real random variable. -/ lemma measure_ge_le_exp_mul_mgf [is_finite_measure μ] (ε : ℝ) (ht : 0 ≤ t) (h_int : integrable (λ ω, exp (t * X ω)) μ) : From 4918ed169d2264ea0d8cf3701448385db15c2bfd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 26 Jul 2022 12:57:23 +0000 Subject: [PATCH 068/173] chore(number_theory/padics/padic_integers): golf the comm_ring instance (#15590) This results in nicer definitional equalities that don't involve the application of a recursor. This also renames `padic_int.coe_coe` to `padic_int.coe_nat_cast` and `padic_int.coe_coe_int` to `padic_int.coe_int_cast` to match other similar lemmas in mathlib. Finally, this fixes the TODO comment ```lean -- TODO: define nat_cast/int_cast so that coe_coe and coe_coe_int are rfl ``` --- src/number_theory/padics/padic_integers.lean | 143 ++++++------------- src/number_theory/padics/ring_homs.lean | 9 +- 2 files changed, 47 insertions(+), 105 deletions(-) diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index da7c603769e7f..7543b9e3de9c3 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -61,31 +61,37 @@ variables {p : ℕ} [fact p.prime] instance : has_coe ℤ_[p] ℚ_[p] := ⟨subtype.val⟩ -lemma ext {x y : ℤ_[p]} : (x : ℚ_[p]) = y → x = y := subtype.ext_iff_val.2 +lemma ext {x y : ℤ_[p]} : (x : ℚ_[p]) = y → x = y := subtype.ext + +variables (p) + +/-- The padic integers as a subring of the padics -/ +def subring : subring (ℚ_[p]) := +{ carrier := {x : ℚ_[p] | ∥x∥ ≤ 1}, + zero_mem' := by norm_num, + one_mem' := by norm_num, + add_mem' := λ x y hx hy, (padic_norm_e.nonarchimedean _ _).trans $ max_le_iff.2 ⟨hx, hy⟩, + mul_mem' := λ x y hx hy, (padic_norm_e.mul _ _).trans_le $ mul_le_one hx (norm_nonneg _) hy, + neg_mem' := λ x hx, (norm_neg _).trans_le hx } + +@[simp] lemma mem_subring_iff {x : ℚ_[p]} : x ∈ subring p ↔ ∥x∥ ≤ 1 := iff.rfl + +variables {p} /-- Addition on ℤ_p is inherited from ℚ_p. -/ -instance : has_add ℤ_[p] := -⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x+y, - le_trans (padic_norm_e.nonarchimedean _ _) (max_le_iff.2 ⟨hx,hy⟩)⟩⟩ +instance : has_add ℤ_[p] := (by apply_instance : has_add (subring p)) /-- Multiplication on ℤ_p is inherited from ℚ_p. -/ -instance : has_mul ℤ_[p] := -⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x*y, - begin rw padic_norm_e.mul, apply mul_le_one; {assumption <|> apply norm_nonneg} end⟩⟩ +instance : has_mul ℤ_[p] := (by apply_instance : has_mul (subring p)) /-- Negation on ℤ_p is inherited from ℚ_p. -/ -instance : has_neg ℤ_[p] := -⟨λ ⟨x, hx⟩, ⟨-x, by simpa⟩⟩ +instance : has_neg ℤ_[p] := (by apply_instance : has_neg (subring p)) /-- Subtraction on ℤ_p is inherited from ℚ_p. -/ -instance : has_sub ℤ_[p] := -⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x - y, - by { rw sub_eq_add_neg, rw ← norm_neg at hy, - exact le_trans (padic_norm_e.nonarchimedean _ _) (max_le_iff.2 ⟨hx, hy⟩) }⟩⟩ +instance : has_sub ℤ_[p] := (by apply_instance : has_sub (subring p)) /-- Zero on ℤ_p is inherited from ℚ_p. -/ -instance : has_zero ℤ_[p] := -⟨⟨0, by norm_num⟩⟩ +instance : has_zero ℤ_[p] := (by apply_instance : has_zero (subring p)) instance : inhabited ℤ_[p] := ⟨0⟩ @@ -97,69 +103,28 @@ instance : has_one ℤ_[p] := @[simp] lemma val_eq_coe (z : ℤ_[p]) : z.val = z := rfl -@[simp, norm_cast] lemma coe_add : ∀ (z1 z2 : ℤ_[p]), ((z1 + z2 : ℤ_[p]) : ℚ_[p]) = z1 + z2 -| ⟨_, _⟩ ⟨_, _⟩ := rfl - -@[simp, norm_cast] lemma coe_mul : ∀ (z1 z2 : ℤ_[p]), ((z1 * z2 : ℤ_[p]) : ℚ_[p]) = z1 * z2 -| ⟨_, _⟩ ⟨_, _⟩ := rfl - -@[simp, norm_cast] lemma coe_neg : ∀ (z1 : ℤ_[p]), ((-z1 : ℤ_[p]) : ℚ_[p]) = -z1 -| ⟨_, _⟩ := rfl - -@[simp, norm_cast] lemma coe_sub : ∀ (z1 z2 : ℤ_[p]), ((z1 - z2 : ℤ_[p]) : ℚ_[p]) = z1 - z2 -| ⟨_, _⟩ ⟨_, _⟩ := rfl - +@[simp, norm_cast] lemma coe_add (z1 z2 : ℤ_[p]) : ((z1 + z2 : ℤ_[p]) : ℚ_[p]) = z1 + z2 := rfl +@[simp, norm_cast] lemma coe_mul (z1 z2 : ℤ_[p]) : ((z1 * z2 : ℤ_[p]) : ℚ_[p]) = z1 * z2 := rfl +@[simp, norm_cast] lemma coe_neg (z1 : ℤ_[p]) : ((-z1 : ℤ_[p]) : ℚ_[p]) = -z1 := rfl +@[simp, norm_cast] lemma coe_sub (z1 z2 : ℤ_[p]) : ((z1 - z2 : ℤ_[p]) : ℚ_[p]) = z1 - z2 := rfl @[simp, norm_cast] lemma coe_one : ((1 : ℤ_[p]) : ℚ_[p]) = 1 := rfl - @[simp, norm_cast] lemma coe_zero : ((0 : ℤ_[p]) : ℚ_[p]) = 0 := rfl instance : add_comm_group ℤ_[p] := -by refine_struct -{ add := (+), - neg := has_neg.neg, - zero := (0 : ℤ_[p]), - sub := has_sub.sub, - nsmul := @nsmul_rec _ ⟨(0 : ℤ_[p])⟩ ⟨(+)⟩, - zsmul := @zsmul_rec _ ⟨(0 : ℤ_[p])⟩ ⟨(+)⟩ ⟨has_neg.neg⟩ }; -intros; try { refl }; ext; simp; ring - -instance : add_group_with_one ℤ_[p] := --- TODO: define nat_cast/int_cast so that coe_coe and coe_coe_int are rfl -{ one := 1, .. padic_int.add_comm_group } - -instance : ring ℤ_[p] := -by refine_struct -{ add := (+), - mul := (*), - neg := has_neg.neg, - zero := (0 : ℤ_[p]), - one := 1, - sub := has_sub.sub, - npow := @npow_rec _ ⟨(1 : ℤ_[p])⟩ ⟨(*)⟩, - .. padic_int.add_group_with_one }; -intros; try { refl }; ext; simp; ring - -@[simp, norm_cast] lemma coe_coe : ∀ n : ℕ, ((n : ℤ_[p]) : ℚ_[p]) = n -| 0 := rfl -| (k+1) := by simp [coe_coe] - -@[simp, norm_cast] lemma coe_coe_int : ∀ (z : ℤ), ((z : ℤ_[p]) : ℚ_[p]) = z -| (int.of_nat n) := by simp -| -[1+n] := by simp +(by apply_instance : add_comm_group (subring p)) + +instance : comm_ring ℤ_[p] := +(by apply_instance : comm_ring (subring p)) + +@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : ((n : ℤ_[p]) : ℚ_[p]) = n := rfl +@[simp, norm_cast] lemma coe_int_cast (z : ℤ) : ((z : ℤ_[p]) : ℚ_[p]) = z := rfl /-- The coercion from ℤ[p] to ℚ[p] as a ring homomorphism. -/ -def coe.ring_hom : ℤ_[p] →+* ℚ_[p] := -{ to_fun := (coe : ℤ_[p] → ℚ_[p]), - map_zero' := rfl, - map_one' := rfl, - map_mul' := coe_mul, - map_add' := coe_add } +def coe.ring_hom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype -@[simp, norm_cast] lemma coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x^n) : ℚ_[p]) = (↑x : ℚ_[p])^n := -(coe.ring_hom : ℤ_[p] →+* ℚ_[p]).map_pow x n +@[simp, norm_cast] lemma coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x^n) : ℚ_[p]) = (↑x : ℚ_[p])^n := rfl -@[simp] lemma mk_coe : ∀ (k : ℤ_[p]), (⟨k, k.2⟩ : ℤ_[p]) = k -| ⟨_, _⟩ := rfl +@[simp] lemma mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := subtype.coe_eta _ _ /-- The inverse of a p-adic integer with norm equal to 1 is also a p-adic integer. Otherwise, the inverse is defined to be 0. -/ @@ -212,20 +177,6 @@ instance : has_norm ℤ_[p] := ⟨λ z, ∥(z : ℚ_[p])∥⟩ variables {p} -protected lemma mul_comm : ∀ z1 z2 : ℤ_[p], z1*z2 = z2*z1 -| ⟨q1, h1⟩ ⟨q2, h2⟩ := show (⟨q1*q2, _⟩ : ℤ_[p]) = ⟨q2*q1, _⟩, by simp [_root_.mul_comm] - -protected lemma zero_ne_one : (0 : ℤ_[p]) ≠ 1 := -show (⟨(0 : ℚ_[p]), _⟩ : ℤ_[p]) ≠ ⟨(1 : ℚ_[p]), _⟩, from mt subtype.ext_iff_val.1 zero_ne_one - -protected lemma eq_zero_or_eq_zero_of_mul_eq_zero : - ∀ (a b : ℤ_[p]), a * b = 0 → a = 0 ∨ b = 0 -| ⟨a, ha⟩ ⟨b, hb⟩ := λ h : (⟨a * b, _⟩ : ℤ_[p]) = ⟨0, _⟩, -have a * b = 0, from subtype.ext_iff_val.1 h, -(mul_eq_zero.1 this).elim - (λ h1, or.inl (by simp [h1]; refl)) - (λ h2, or.inr (by simp [h2]; refl)) - lemma norm_def {z : ℤ_[p]} : ∥z∥ = ∥(z : ℚ_[p])∥ := rfl variables (p) @@ -233,8 +184,7 @@ variables (p) instance : normed_comm_ring ℤ_[p] := { dist_eq := λ ⟨_, _⟩ ⟨_, _⟩, rfl, norm_mul := by simp [norm_def], - mul_comm := padic_int.mul_comm, - norm := norm, .. padic_int.ring, .. padic_int.metric_space p } + norm := norm, .. padic_int.comm_ring, .. padic_int.metric_space p } instance : norm_one_class ℤ_[p] := ⟨norm_def.trans norm_one⟩ @@ -247,9 +197,7 @@ instance is_absolute_value : is_absolute_value (λ z : ℤ_[p], ∥z∥) := variables {p} instance : is_domain ℤ_[p] := -{ eq_zero_or_eq_zero_of_mul_eq_zero := λ x y, padic_int.eq_zero_or_eq_zero_of_mul_eq_zero x y, - exists_pair_ne := ⟨0, 1, padic_int.zero_ne_one⟩, - .. padic_int.normed_comm_ring p } +function.injective.is_domain (subring p).subtype subtype.coe_injective end padic_int @@ -257,8 +205,7 @@ namespace padic_int /-! ### Norm -/ variables {p : ℕ} [fact p.prime] -lemma norm_le_one : ∀ z : ℤ_[p], ∥z∥ ≤ 1 -| ⟨_, h⟩ := h +lemma norm_le_one (z : ℤ_[p]) : ∥z∥ ≤ 1 := z.2 @[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ := by simp [norm_def] @@ -267,11 +214,10 @@ by simp [norm_def] | 0 := by simp | (k+1) := by { rw [pow_succ, pow_succ, norm_mul], congr, apply norm_pow } -theorem nonarchimedean : ∀ (q r : ℤ_[p]), ∥q + r∥ ≤ max (∥q∥) (∥r∥) -| ⟨_, _⟩ ⟨_, _⟩ := padic_norm_e.nonarchimedean _ _ +theorem nonarchimedean (q r : ℤ_[p]) : ∥q + r∥ ≤ max (∥q∥) (∥r∥) := padic_norm_e.nonarchimedean _ _ -theorem norm_add_eq_max_of_ne : ∀ {q r : ℤ_[p]}, ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥) -| ⟨_, _⟩ ⟨_, _⟩ := padic_norm_e.add_eq_max_of_ne +theorem norm_add_eq_max_of_ne {q r : ℤ_[p]} : ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥) := +padic_norm_e.add_eq_max_of_ne lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]} (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := @@ -292,12 +238,9 @@ by simp [norm_def] @[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ∥q∥ ≤ 1) : @norm ℤ_[p] _ ⟨q, hq⟩ = ∥q∥ := rfl -@[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ := -show ∥((p : ℤ_[p]) : ℚ_[p])∥ = p⁻¹, by exact_mod_cast padic_norm_e.norm_p +@[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ := padic_norm_e.norm_p -@[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) := -show ∥((p^n : ℤ_[p]) : ℚ_[p])∥ = p^(-n:ℤ), -by { convert padic_norm_e.norm_p_pow n, simp, } +@[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) := padic_norm_e.norm_p_pow n private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) : cau_seq ℚ_[p] (λ a, ∥a∥) := diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index 248ec0b169c20..2e90b9c350fea 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -81,7 +81,7 @@ lemma is_unit_denom (r : ℚ) (h : ∥(r : ℚ_[p])∥ ≤ 1) : is_unit (r.denom begin rw is_unit_iff, apply le_antisymm (r.denom : ℤ_[p]).2, - rw [← not_lt, val_eq_coe, coe_coe], + rw [← not_lt, val_eq_coe, coe_nat_cast], intro norm_denom_lt, have hr : ∥(r * r.denom : ℚ_[p])∥ = ∥(r.num : ℚ_[p])∥, { rw_mod_cast @rat.mul_denom_eq_num r, refl, }, @@ -126,7 +126,7 @@ begin simp only [sub_mul, int.cast_coe_nat, ring_hom.eq_int_cast, int.cast_mul, sub_left_inj, int.cast_sub], apply subtype.coe_injective, - simp only [coe_mul, subtype.coe_mk, coe_coe], + simp only [coe_mul, subtype.coe_mk, coe_nat_cast], rw_mod_cast @rat.mul_denom_eq_num r, refl }, exact norm_sub_mod_part_aux r h end @@ -182,7 +182,7 @@ begin lift n to ℕ using hzn, use n, split, {exact_mod_cast hnp}, - simp only [norm_def, coe_sub, subtype.coe_mk, coe_coe] at hn ⊢, + simp only [norm_def, coe_sub, subtype.coe_mk, coe_nat_cast] at hn ⊢, rw show (x - n : ℚ_[p]) = (x - r) + (r - n), by ring, apply lt_of_le_of_lt (padic_norm_e.nonarchimedean _ _), apply max_lt hr, @@ -587,8 +587,7 @@ begin apply lt_trans _ hε', change ↑(padic_norm_e _) < _, norm_cast, - convert hN _ hn, - simp [nth_hom, lim_nth_hom, nth_hom_seq, of_int_seq], + exact hN _ hn, end lemma lim_nth_hom_zero : lim_nth_hom f_compat 0 = 0 := From 4a579f0e63829ab4bbb4e3b2d6f3f001863346bf Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 12:57:24 +0000 Subject: [PATCH 069/173] feat(order/filter/bases): lemmas about bases of product filters (#15630) * Move `has_basis.prod_self` and `mem_prod_self_iff` down to golf the proof. * Rename `has_basis.prod''` to `has_basis.prod_pprod`. * Rename `has_basis.prod'` to `has_basis.prod_same_index`. * Add `has_basis.prod_same_index_mono ` and `has_basis.prod_same_index_anti`. --- src/order/filter/bases.lean | 61 +++++++++++++-------------- src/topology/uniform_space/basic.lean | 2 +- 2 files changed, 31 insertions(+), 32 deletions(-) diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 48f550d9bfa7c..a5a149a05c99b 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -675,22 +675,6 @@ lemma comap_has_basis (f : α → β) (l : filter β) : has_basis (comap f l) (λ s : set β, s ∈ l) (λ s, f ⁻¹' s) := ⟨λ t, mem_comap⟩ -lemma has_basis.prod_self (hl : l.has_basis p s) : - (l ×ᶠ l).has_basis p (λ i, s i ×ˢ s i) := -⟨begin - intro t, - apply mem_prod_iff.trans, - split, - { rintros ⟨t₁, ht₁, t₂, ht₂, H⟩, - rcases hl.mem_iff.1 (inter_mem ht₁ ht₂) with ⟨i, hi, ht⟩, - exact ⟨i, hi, λ p ⟨hp₁, hp₂⟩, H ⟨(ht hp₁).1, (ht hp₂).2⟩⟩ }, - { rintros ⟨i, hi, H⟩, - exact ⟨s i, hl.mem_of_mem hi, s i, hl.mem_of_mem hi, H⟩ } -end⟩ - -lemma mem_prod_self_iff {s} : s ∈ l ×ᶠ l ↔ ∃ t ∈ l, t ×ˢ t ⊆ s := -l.basis_sets.prod_self.mem_iff - lemma has_basis.forall_mem_mem (h : has_basis l p s) {x : α} : (∀ t ∈ l, x ∈ t) ↔ ∀ i, p i → x ∈ s i := begin @@ -755,7 +739,7 @@ lemma tendsto.basis_both (H : tendsto f la lb) (hla : la.has_basis pa sa) ∀ ib (hib : pb ib), ∃ ia (hia : pa ia), ∀ x ∈ sa ia, f x ∈ sb ib := (hla.tendsto_iff hlb).1 H -lemma has_basis.prod'' (hla : la.has_basis pa sa) (hlb : lb.has_basis pb sb) : +lemma has_basis.prod_pprod (hla : la.has_basis pa sa) (hlb : lb.has_basis pb sb) : (la ×ᶠ lb).has_basis (λ i : pprod ι ι', pa i.1 ∧ pb i.2) (λ i, sa i.1 ×ˢ sb i.2) := (hla.comap prod.fst).inf' (hlb.comap prod.snd) @@ -764,13 +748,12 @@ lemma has_basis.prod {ι ι' : Type*} {pa : ι → Prop} {sa : ι → set α} {p (la ×ᶠ lb).has_basis (λ i : ι × ι', pa i.1 ∧ pb i.2) (λ i, sa i.1 ×ˢ sb i.2) := (hla.comap prod.fst).inf (hlb.comap prod.snd) -lemma has_basis.prod' {la : filter α} {lb : filter β} {ι : Type*} {p : ι → Prop} - {sa : ι → set α} {sb : ι → set β} +lemma has_basis.prod_same_index {p : ι → Prop} {sb : ι → set β} (hla : la.has_basis p sa) (hlb : lb.has_basis p sb) (h_dir : ∀ {i j}, p i → p j → ∃ k, p k ∧ sa k ⊆ sa i ∧ sb k ⊆ sb j) : (la ×ᶠ lb).has_basis p (λ i, sa i ×ˢ sb i) := begin - simp only [has_basis_iff, (hla.prod hlb).mem_iff], + simp only [has_basis_iff, (hla.prod_pprod hlb).mem_iff], refine λ t, ⟨_, _⟩, { rintros ⟨⟨i, j⟩, ⟨hi, hj⟩, hsub : sa i ×ˢ sb j ⊆ t⟩, rcases h_dir hi hj with ⟨k, hk, ki, kj⟩, @@ -779,18 +762,34 @@ begin exact ⟨⟨i, i⟩, ⟨hi, hi⟩, h⟩ }, end -lemma has_antitone_basis.prod {f : filter α} {g : filter β} - {s : ℕ → set α} {t : ℕ → set β} (hf : has_antitone_basis f s) (hg : has_antitone_basis g t) : +lemma has_basis.prod_same_index_mono {ι : Type*} [linear_order ι] + {p : ι → Prop} {sa : ι → set α} {sb : ι → set β} + (hla : la.has_basis p sa) (hlb : lb.has_basis p sb) + (hsa : monotone_on sa {i | p i}) (hsb : monotone_on sb {i | p i}) : + (la ×ᶠ lb).has_basis p (λ i, sa i ×ˢ sb i) := +hla.prod_same_index hlb $ λ i j hi hj, + have p (min i j), from min_rec' _ hi hj, + ⟨min i j, this, hsa this hi $ min_le_left _ _, hsb this hj $ min_le_right _ _⟩ + +lemma has_basis.prod_same_index_anti {ι : Type*} [linear_order ι] + {p : ι → Prop} {sa : ι → set α} {sb : ι → set β} + (hla : la.has_basis p sa) (hlb : lb.has_basis p sb) + (hsa : antitone_on sa {i | p i}) (hsb : antitone_on sb {i | p i}) : + (la ×ᶠ lb).has_basis p (λ i, sa i ×ˢ sb i) := +@has_basis.prod_same_index_mono _ _ _ _ ιᵒᵈ _ _ _ _ hla hlb hsa.dual_left hsb.dual_left + +lemma has_basis.prod_self (hl : la.has_basis pa sa) : + (la ×ᶠ la).has_basis pa (λ i, sa i ×ˢ sa i) := +hl.prod_same_index hl $ λ i j hi hj, by simpa only [exists_prop, subset_inter_iff] + using hl.mem_iff.1 (inter_mem (hl.mem_of_mem hi) (hl.mem_of_mem hj)) + +lemma mem_prod_self_iff {s} : s ∈ la ×ᶠ la ↔ ∃ t ∈ la, t ×ˢ t ⊆ s := +la.basis_sets.prod_self.mem_iff + +lemma has_antitone_basis.prod {ι : Type*} [linear_order ι] {f : filter α} {g : filter β} + {s : ι → set α} {t : ι → set β} (hf : has_antitone_basis f s) (hg : has_antitone_basis g t) : has_antitone_basis (f ×ᶠ g) (λ n, s n ×ˢ t n) := -begin - have h : has_basis (f ×ᶠ g) _ _ := has_basis.prod' hf.to_has_basis hg.to_has_basis _, - swap, - { intros i j, - simp only [true_and, forall_true_left], - exact ⟨max i j, hf.antitone (le_max_left _ _), hg.antitone (le_max_right _ _)⟩, }, - refine ⟨h, λ n m hn_le_m, set.prod_mono _ _⟩, - exacts [hf.antitone hn_le_m, hg.antitone hn_le_m] -end +⟨hf.1.prod_same_index_anti hg.1 (hf.2.antitone_on _) (hg.2.antitone_on _), hf.2.set_prod hg.2⟩ lemma has_basis.coprod {ι ι' : Type*} {pa : ι → Prop} {sa : ι → set α} {pb : ι' → Prop} {sb : ι' → set β} (hla : la.has_basis pa sa) (hlb : lb.has_basis pb sb) : diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 878b9a8bf7d08..a9b7e29604558 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -661,7 +661,7 @@ lemma uniform_space.has_basis_nhds_prod (x y : α) : has_basis (𝓝 (x, y)) (λ s, s ∈ 𝓤 α ∧ symmetric_rel s) $ λ s, ball x s ×ˢ ball y s := begin rw nhds_prod_eq, - apply (has_basis_nhds x).prod' (has_basis_nhds y), + apply (has_basis_nhds x).prod_same_index (has_basis_nhds y), rintro U V ⟨U_in, U_symm⟩ ⟨V_in, V_symm⟩, exact ⟨U ∩ V, ⟨(𝓤 α).inter_sets U_in V_in, U_symm.inter V_symm⟩, ball_inter_left x U V, ball_inter_right y U V⟩, From 7a83f118850e0ffafa8183e22ba4a9d05b819a39 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 26 Jul 2022 14:35:09 +0000 Subject: [PATCH 070/173] feat(tactic/positivity): a tactic for proving positivity/nonnegativity (#15618) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is a new tactic, `positivity`, which solves goals of the form `0 ≤ x` and `0 < x`. The tactic works recursively according to the syntax of the expression `x`. Co-authored-by: Mario Carneiro --- src/algebra/order/field.lean | 4 + src/analysis/normed/group/basic.lean | 11 + src/data/real/sqrt.lean | 20 +- src/tactic/positivity.lean | 366 +++++++++++++++++++++++++++ src/topology/metric_space/basic.lean | 12 + test/positivity.lean | 105 ++++++++ 6 files changed, 517 insertions(+), 1 deletion(-) create mode 100644 src/tactic/positivity.lean create mode 100644 test/positivity.lean diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index a96b6797c29d2..b08aa87da169c 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -101,9 +101,13 @@ suffices ∀ a : α, 0 < a → 0 < a⁻¹, from ⟨λ h, inv_inv a ▸ this _ h, this a⟩, assume a ha, flip lt_of_mul_lt_mul_left ha.le $ by simp [ne_of_gt ha, zero_lt_one] +alias inv_pos ↔ _ inv_pos_of_pos + @[simp] lemma inv_nonneg : 0 ≤ a⁻¹ ↔ 0 ≤ a := by simp only [le_iff_eq_or_lt, inv_pos, zero_eq_inv] +alias inv_nonneg ↔ _ inv_nonneg_of_nonneg + @[simp] lemma inv_lt_zero : a⁻¹ < 0 ↔ a < 0 := by simp only [← not_le, inv_nonneg] diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 50a3481fe0ff7..cdab8241e5096 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -216,6 +216,17 @@ by simpa only [dist_add_left, dist_add_right, dist_comm h₂] @[simp] lemma norm_nonneg (g : E) : 0 ≤ ∥g∥ := by { rw[←dist_zero_right], exact dist_nonneg } +section +open tactic tactic.positivity + +/-- Extension for the `positivity` tactic: norms are nonnegative. -/ +@[positivity] +meta def _root_.tactic.positivity_norm : expr → tactic strictness +| `(∥%%a∥) := nonnegative <$> mk_app ``norm_nonneg [a] +| _ := failed + +end + @[simp] lemma norm_zero : ∥(0 : E)∥ = 0 := by rw [← dist_zero_right, dist_self] lemma ne_zero_of_norm_ne_zero {g : E} : ∥g∥ ≠ 0 → g ≠ 0 := mt $ by { rintro rfl, exact norm_zero } diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index 3be9eb443f665..256c82f0571c2 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Floris van Doorn, Yury Kudryashov -/ import topology.algebra.order.monotone_continuity import topology.instances.nnreal -import tactic.norm_cast +import tactic.positivity /-! # Square root of a real number @@ -272,6 +272,24 @@ by rw [← not_le, not_iff_not, sqrt_eq_zero'] lt_iff_lt_of_le_iff_le (iff.trans (by simp [le_antisymm_iff, sqrt_nonneg]) sqrt_eq_zero') +alias sqrt_pos ↔ _ sqrt_pos_of_pos + +section +open tactic tactic.positivity + +/-- Extension for the `positivity` tactic: a square root is nonnegative, and is strictly positive if +its input is. -/ +@[positivity] +meta def _root_.tactic.positivity_sqrt : expr → tactic strictness +| `(real.sqrt %%a) := do + (do -- if can prove `0 < a`, report positivity + positive pa ← core a, + positive <$> mk_app ``sqrt_pos_of_pos [pa]) <|> + nonnegative <$> mk_app ``sqrt_nonneg [a] -- else report nonnegativity +| _ := failed + +end + @[simp] theorem sqrt_mul (hx : 0 ≤ x) (y : ℝ) : sqrt (x * y) = sqrt x * sqrt y := by simp_rw [sqrt, ← nnreal.coe_mul, nnreal.coe_eq, real.to_nnreal_mul hx, nnreal.sqrt_mul] diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean new file mode 100644 index 0000000000000..b9c36aa5a13cb --- /dev/null +++ b/src/tactic/positivity.lean @@ -0,0 +1,366 @@ +/- +Copyright (c) 2022 Mario Carneiro, Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Heather Macbeth +-/ +import tactic.norm_num + +/-! # `positivity` tactic + +The `positivity` tactic in this file solves goals of the form `0 ≤ x` and `0 < x`. The tactic works +recursively according to the syntax of the expression `x`. For example, a goal of the form +`0 ≤ 3 * a ^ 2 + b * c` can be solved either +* by a hypothesis such as `5 ≤ 3 * a ^ 2 + b * c` which directly implies the nonegativity of + `3 * a ^ 2 + b * c`; or, +* by the application of the lemma `add_nonneg` and the success of the `positivity` tactic on the two + sub-expressions `3 * a ^ 2` and `b * c`. + +For each supported operation, one must write a small tactic, tagged with the attribute +`@[positivity]`, which operates only on goals whose leading function application is that operation. +Typically, this small tactic will run the full `positivity` tactic on one or more of the function's +arguments (which is where the recursion comes in), and if successful will combine this with an +appropriate lemma to give positivity of the full expression. + +This file contains the core `positivity` logic and the small tactics handling the basic operations: +`min`, `max`, `+`, `*`, `/`, `⁻¹`, raising to natural powers, and taking absolute values. Further +extensions, e.g. to handle `real.sqrt` and norms, can be found in the files of the library which +introduce these operations. + +## Main declarations + +* `tactic.norm_num.positivity` tries to prove positivity of an expression by running `norm_num` on + it. This is one of the base cases of the recursion. +* `tactic.positivity.compare_hyp` tries to prove positivity of an expression by comparing with a + provided hypothesis. If the hypothesis is of the form `a ≤ b` or similar, with `b` matching the + expression whose proof of positivity is desired, then it will check whether `a` can be proved + positive via `tactic.norm_num.positivity` and if so apply a transitivity lemma. This is the other + base case of the recursion. +* `tactic.positivity.attr` creates the `positivity` user attribute for tagging the extension + tactics handling specific operations, and specifies the behaviour for a single step of the + recursion +* `tactic.positivity.core` collects the list of tactics with the `@[positivity]` attribute and + calls the first recursion step as specified in `tactic.positivity.attr`. Its input is `e : expr` + and its output (if it succeeds) is a term of a custom inductive type + `tactic.positivity.strictness`, containing an `expr` which is a proof of the + strict-positivity/nonnegativity of `e` as well as an indication of whether what could be proved + was strict-positivity or nonnegativity +* `tactic.interactive.positivity` is the user-facing tactic. It parses the goal and, if it is of + one of the forms `0 ≤ e`, `0 < e`, `e > 0`, `e ≥ 0`, it sends `e` to `tactic.positivity.core`. + +## TODO + +Implement extensions for other operations (raising to non-numeral powers, `exp`, `log`, coercions +from `ℕ` and `ℝ≥0`). + +-/ + +namespace tactic + +/-- Inductive type recording either `positive` and an expression (typically a proof of a fact +`0 < x`) or `nonnegative` and an expression (typically a proof of a fact `0 ≤ x`). -/ +@[derive [decidable_eq]] +meta inductive positivity.strictness : Type +| positive : expr → positivity.strictness +| nonnegative : expr → positivity.strictness + +export positivity.strictness (positive nonnegative) + +private lemma lt_of_lt_of_eq'' {α} [preorder α] {b b' a : α} : b = b' → a < b' → a < b := +λ h1 h2, lt_of_lt_of_eq h2 h1.symm + +/-- First base case of the `positivity` tactic. We try `norm_num` to prove directly that an +expression `e` is positive or nonnegative. -/ +meta def norm_num.positivity (e : expr) : tactic strictness := do + (e', p) ← norm_num.derive e <|> refl_conv e, + e'' ← e'.to_rat, + typ ← infer_type e', + ic ← mk_instance_cache typ, + if e'' > 0 then do + (ic, p₁) ← norm_num.prove_pos ic e', + p ← mk_app ``lt_of_lt_of_eq'' [p, p₁], + pure (positive p) + else if e'' = 0 then do + p' ← mk_app ``ge_of_eq [p], + pure (nonnegative p') + else failed + +namespace positivity + +/-- Given two tactics whose result is `strictness`, report a `strictness`: +- if at least one gives `positive`, report `positive` and one of the expressions giving a proof of + positivity +- if neither gives `pos` but at least one gives `nonnegative`, report `nonnegative` and one of the + expressions giving a proof of nonnegativity +- if both fail, fail -/ +meta def orelse' (tac1 tac2 : tactic strictness) : tactic strictness := do + res ← try_core tac1, + match res with + | none := tac2 + | some res@(nonnegative e) := tac2 <|> pure res + | some res@(positive _) := pure res + end + +/-! ### Core logic of the `positivity` tactic -/ + +/-- Second base case of the `positivity` tactic. Prove an expression `e` is positive/nonnegative by +finding a hypothesis of the form `a < e` or `a ≤ e` in which `a` can be proved positive/nonnegative +by `norm_num`. -/ +meta def compare_hyp (e p₂ : expr) : tactic strictness := do + p_typ ← infer_type p₂, + (lo, hi, strict₂) ← match p_typ with -- TODO also handle equality hypotheses + | `(%%lo ≤ %%hi) := pure (lo, hi, ff) + | `(%%hi ≥ %%lo) := pure (lo, hi, ff) + | `(%%lo < %%hi) := pure (lo, hi, tt) + | `(%%hi > %%lo) := pure (lo, hi, tt) + | _ := failed + end, + is_def_eq e hi, + strictness₁ ← norm_num.positivity lo, + match strictness₁, strict₂ with + | (positive p₁), tt := positive <$> mk_app ``lt_trans [p₁, p₂] + | (positive p₁), ff := positive <$> mk_app `lt_of_lt_of_le [p₁, p₂] + | (nonnegative p₁), tt := positive <$> mk_app `lt_of_le_of_lt [p₁, p₂] + | (nonnegative p₁), ff := nonnegative <$> mk_app `le_trans [p₁, p₂] + end + +/-- Attribute allowing a user to tag a tactic as an extension for `tactic.positivity`. The main +(recursive) step of this tactic is to try successively all the extensions tagged with this attribute +on the expression at hand, and also to try the two "base case" tactics `tactic.norm_num.positivity`, +`tactic.positivity.compare_hyp` on the expression at hand. -/ +@[user_attribute] +meta def attr : user_attribute (expr → tactic strictness) unit := +{ name := `positivity, + descr := "extensions handling particular operations for the `positivity` tactic", + cache_cfg := + { mk_cache := λ ns, do + { t ← ns.mfoldl + (λ (t : expr → tactic strictness) n, do + t' ← eval_expr (expr → tactic strictness) (expr.const n []), + pure (λ e, orelse' (t' e) (t e))) + (λ _, failed), + pure $ λ e, orelse' + (t e) $ orelse' -- run all the extensions on `e` + (norm_num.positivity e) $ -- directly try `norm_num` on `e` + -- loop over hypotheses and try to compare with `e` + local_context >>= list.foldl (λ tac h, orelse' tac (compare_hyp e h)) failed }, + dependencies := [] } } + +/-- Look for a proof of positivity/nonnegativity of an expression `e`; if found, return the proof +together with a `strictness` stating whether the proof found was for strict positivity +(`positive p`) or only for nonnegativity (`nonnegative p`). -/ +meta def core (e : expr) : tactic strictness := do + f ← attr.get_cache, + f e <|> fail "failed to prove positivity/nonnegativity" + +end positivity + +open positivity + +namespace interactive + +setup_tactic_parser + +/-- Tactic solving goals of the form `0 ≤ x` and `0 < x`. The tactic works recursively according to +the syntax of the expression `x`, if the atoms composing the expression all have numeric lower +bounds which can be proved positive/nonnegative by `norm_num`. This tactic either closes the goal +or fails. + +Examples: +``` +example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 3 + a := by positivity + +example {a : ℤ} (ha : 1 < a) : 0 < |(3:ℤ) + a| := by positivity + +example {b : ℤ} : 0 ≤ max (-3) (b ^ 2) := by positivity +``` +-/ +meta def positivity : tactic unit := focus1 $ do + t ← target, + (rel_desired, a) ← match t with + | `(0 ≤ %%e₂) := pure (ff, e₂) + | `(%%e₂ ≥ 0) := pure (ff, e₂) + | `(0 < %%e₂) := pure (tt, e₂) + | `(%%e₂ > 0) := pure (tt, e₂) + | _ := fail "not a positivity/nonnegativity goal" + end, + strictness_proved ← tactic.positivity.core a, + match rel_desired, strictness_proved with + | tt, (positive p) := pure p + | tt, (nonnegative _) := fail ("failed to prove strict positivity, but it would be possible to " + ++ "prove nonnegativity if desired") + | ff, (positive p) := mk_app ``le_of_lt [p] + | ff, (nonnegative p) := pure p + end >>= tactic.exact + +add_tactic_doc +{ name := "positivity", + category := doc_category.tactic, + decl_names := [`tactic.interactive.positivity], + tags := ["arithmetic", "monotonicity", "finishing"] } + +end interactive + +variables {R : Type*} + +/-! ### `positivity` extensions for particular arithmetic operations -/ + +private lemma le_min_of_lt_of_le [linear_order R] (a b c : R) (ha : a < b) (hb : a ≤ c) : + a ≤ min b c := +le_min ha.le hb + +private lemma le_min_of_le_of_lt [linear_order R] (a b c : R) (ha : a ≤ b) (hb : a < c) : + a ≤ min b c := +le_min ha hb.le + +/-- Extension for the `positivity` tactic: the `min` of two numbers is nonnegative if both are +nonnegative, and strictly positive if both are. -/ +@[positivity] +meta def positivity_min : expr → tactic strictness +| `(min %%a %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | (positive pa), (positive pb) := positive <$> mk_app ``lt_min [pa, pb] + | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``le_min_of_lt_of_le [pa, pb] + | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``le_min_of_le_of_lt [pa, pb] + | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``le_min [pa, pb] + end +| _ := failed + +/-- Extension for the `positivity` tactic: the `max` of two numbers is nonnegative if at least one +is nonnegative, and strictly positive if at least one is positive. -/ +@[positivity] +meta def positivity_max : expr → tactic strictness +| `(max %%a %%b) := tactic.positivity.orelse' (do + strictness_a ← core a, + match strictness_a with + | (positive pa) := positive <$> mk_mapp ``lt_max_of_lt_left [none, none, none, a, b, pa] + | (nonnegative pa) := + nonnegative <$> mk_mapp ``le_max_of_le_left [none, none, none, a, b, pa] + end) + (do + strictness_b ← core b, + match strictness_b with + | (positive pb) := positive <$> mk_mapp ``lt_max_of_lt_right [none, none, none, a, b, pb] + | (nonnegative pb) := + nonnegative <$> mk_mapp ``le_max_of_le_right [none, none, none, a, b, pb] + end) +| _ := failed + +/-- Extension for the `positivity` tactic: addition is nonnegative if both summands are nonnegative, +and strictly positive if at least one summand is. -/ +@[positivity] +meta def positivity_add : expr → tactic strictness +| `(%%a + %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | (positive pa), (positive pb) := positive <$> mk_app ``add_pos [pa, pb] + | (positive pa), (nonnegative pb) := positive <$> mk_app ``lt_add_of_pos_of_le [pa, pb] + | (nonnegative pa), (positive pb) := positive <$> mk_app ``lt_add_of_le_of_pos [pa, pb] + | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``add_nonneg [pa, pb] + end +| _ := failed + +private lemma mul_nonneg_of_pos_of_nonneg [linear_ordered_semiring R] (a b : R) (ha : 0 < a) + (hb : 0 ≤ b) : + 0 ≤ a * b := +mul_nonneg ha.le hb + +private lemma mul_nonneg_of_nonneg_of_pos [linear_ordered_semiring R] (a b : R) (ha : 0 ≤ a) + (hb : 0 < b) : + 0 ≤ a * b := +mul_nonneg ha hb.le + +/-- Extension for the `positivity` tactic: multiplication is nonnegative if both multiplicands are +nonnegative, and strictly positive if both multiplicands are. -/ +@[positivity] +meta def positivity_mul : expr → tactic strictness +| `(%%a * %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | (positive pa), (positive pb) := positive <$> mk_app ``mul_pos [pa, pb] + | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``mul_nonneg_of_pos_of_nonneg [pa, pb] + | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``mul_nonneg_of_nonneg_of_pos [pa, pb] + | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``mul_nonneg [pa, pb] + end +| _ := failed + +private lemma div_nonneg_of_pos_of_nonneg [linear_ordered_field R] {a b : R} (ha : 0 < a) + (hb : 0 ≤ b) : + 0 ≤ a / b := +div_nonneg ha.le hb + +private lemma div_nonneg_of_nonneg_of_pos [linear_ordered_field R] {a b : R} (ha : 0 ≤ a) + (hb : 0 < b) : + 0 ≤ a / b := +div_nonneg ha hb.le + +/-- Extension for the `positivity` tactic: division is nonnegative if both numerator and denominator +are nonnegative, and strictly positive if both numerator and denominator are. -/ +@[positivity] +meta def positivity_div : expr → tactic strictness +| `(%%a / %%b) := do -- TODO handle eg `int.div_nonneg` + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | (positive pa), (positive pb) := positive <$> mk_app ``div_pos [pa, pb] + | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``div_nonneg_of_pos_of_nonneg [pa, pb] + | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``div_nonneg_of_nonneg_of_pos [pa, pb] + | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``div_nonneg [pa, pb] + end +| _ := failed + +/-- Extension for the `positivity` tactic: an inverse of a positive number is positive, an inverse +of a nonnegative number is nonnegative. -/ +@[positivity] +meta def positivity_inv : expr → tactic strictness +| `((%%a)⁻¹) := do + strictness_a ← core a, + match strictness_a with + | (positive pa) := positive <$> mk_app ``inv_pos_of_pos [pa] + | (nonnegative pa) := nonnegative <$> mk_app ``inv_nonneg_of_nonneg [pa] + end +| _ := failed + +private lemma pow_zero_pos [ordered_semiring R] [nontrivial R] (a : R) : 0 < a ^ 0 := +zero_lt_one.trans_le (pow_zero a).ge + +/-- Extension for the `positivity` tactic: raising a number `a` to a natural number power `n` is +known to be positive if `n = 0` (since `a ^ 0 = 1`) or if `0 < a`, and is known to be nonnegative if +`n = 2` (squares are nonnegative) or if `0 ≤ a`. -/ +@[positivity] +meta def positivity_pow : expr → tactic strictness +| `(%%a ^ %%n) := do + n_typ ← infer_type n, + match n_typ with + | `(ℕ) := do + if n = `(0) then + positive <$> mk_app ``pow_zero_pos [a] + else tactic.positivity.orelse' + -- squares are nonnegative (TODO: similar for any `bit0` exponent?) + (nonnegative <$> mk_app ``sq_nonneg [a]) + -- moreover `a ^ n` is positive if `a` is and nonnegative if `a` is + (do + strictness_a ← core a, + match strictness_a with + | (positive pa) := positive <$> mk_app ``pow_pos [pa, n] + | (nonnegative pa) := nonnegative <$> mk_app ``pow_nonneg [pa, n] + end) + | _ := failed -- TODO handle integer powers, maybe even real powers + end +| _ := failed + +/-- Extension for the `positivity` tactic: an absolute value is nonnegative, and is strictly +positive if its input is. -/ +@[positivity] +meta def positivity_abs : expr → tactic strictness +| `(|%%a|) := do + (do -- if can prove `0 < a`, report positivity + positive pa ← core a, + positive <$> mk_app ``abs_pos_of_pos [pa]) <|> + nonnegative <$> mk_app ``abs_nonneg [a] -- else report nonnegativity +| _ := failed + +end tactic diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index e45791b7d41e0..abf1f2204afcc 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas -/ import data.int.interval +import tactic.positivity import topology.algebra.order.compact import topology.metric_space.emetric_space import topology.bornology.constructions @@ -309,6 +310,17 @@ abs_sub_le_iff.2 theorem dist_nonneg {x y : α} : 0 ≤ dist x y := pseudo_metric_space.dist_nonneg' dist dist_self dist_comm dist_triangle +section +open tactic tactic.positivity + +/-- Extension for the `positivity` tactic: distances are nonnegative. -/ +@[positivity] +meta def _root_.tactic.positivity_dist : expr → tactic strictness +| `(dist %%a %%b) := nonnegative <$> mk_app ``dist_nonneg [a, b] +| _ := failed + +end + @[simp] theorem abs_dist {a b : α} : |dist a b| = dist a b := abs_of_nonneg dist_nonneg diff --git a/test/positivity.lean b/test/positivity.lean new file mode 100644 index 0000000000000..0e7bd9fd059f5 --- /dev/null +++ b/test/positivity.lean @@ -0,0 +1,105 @@ +import analysis.normed.group.basic +import data.real.sqrt +import tactic.positivity + +/-! # Tests for the `positivity` tactic + +This tactic proves goals of the form `0 ≤ a` and `0 < a`. +-/ + +/- ## Numeric goals -/ + +example : 0 ≤ 0 := by positivity + +example : 0 ≤ 3 := by positivity + +example : 0 < 3 := by positivity + +/- ## Goals working directly from a hypothesis -/ + +example {a : ℤ} (ha : 0 ≤ a) : 0 ≤ a := by positivity + +example {a : ℤ} (ha : 0 < a) : 0 ≤ a := by positivity + +example {a : ℤ} (ha : 0 < a) : 0 < a := by positivity + +example {a : ℤ} (ha : 3 ≤ a) : 0 < a := by positivity + +example {a : ℤ} (ha : 3 < a) : 0 ≤ a := by positivity + +example {a : ℤ} (ha : 3 < a) : 0 < a := by positivity + +example {a b : ℤ} (h : 0 ≤ a + b) : 0 ≤ a + b := by positivity + +/- ## Tests of the @[positivity] plugin tactics (addition, multiplication, division) -/ + +example {a : ℤ} (ha : 3 < a) : 0 ≤ a + a := by positivity + +example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 ≤ 3 + a + b + b + 14 := by positivity + +example {H : Type*} [linear_ordered_add_comm_group H] {a b : H} (ha : 0 < a) (hb : 0 ≤ b) : + 0 ≤ a + a + b := +by positivity + +example {a : ℤ} (ha : 3 < a) : 0 < a + a := by positivity + +example {a b : ℚ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 := by positivity + +-- TODO: this fails because `div_nonneg` doesn't apply directly to `ℤ` -- it requires a linearly +-- ordered field +-- example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 := by positivity + +example {a : ℕ} : 0 < a ^ 0 := by positivity + +example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 2 + a := by positivity + +example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 3 + a := by positivity + +example {a : ℤ} (ha : 3 < a) : 0 < a ^ 2 + a := by positivity + +example {a b : ℤ} (ha : 3 < a) (hb : b ≥ 4) : 0 ≤ 3 * a ^ 2 * b + b * 7 + 14 := by positivity + +example {a b : ℤ} (ha : 3 < a) (hb : b ≥ 4) : 0 < 3 * a ^ 2 * b + b * 7 + 14 := by positivity + +example {x : ℚ} (hx : 0 ≤ x) : 0 ≤ x⁻¹ := by positivity + +example {a : ℤ} : 0 ≤ |a| := by positivity + +example {a : ℤ} : 0 < |a| + 3 := by positivity + +example {a : ℤ} (ha : 1 < a) : 0 < |(3:ℤ) + a| := by positivity + +example {a : ℝ} (ha : 0 ≤ a) : 0 ≤ real.sqrt a := by positivity + +example {a : ℝ} (ha : 0 ≤ a) : 0 < real.sqrt (a + 3) := by positivity + +example {a b : ℤ} (ha : 3 < a) : 0 ≤ min a (b ^ 2) := by positivity + +-- test that the tactic can ignore arithmetic operations whose associated extension tactic requires +-- more typeclass assumptions than are available +example {R : Type*} [has_zero R] [has_div R] [linear_order R] {a b c : R} (h1 : 0 < a) (h2 : 0 < b) + (h3 : 0 < c) : + 0 < max (a / b) c := +by positivity + +example : 0 ≤ max 3 4 := by positivity + +example {b : ℤ} : 0 ≤ max (-3) (b ^ 2) := by positivity + +example {b : ℤ} : 0 ≤ max (b ^ 2) 0 := by positivity + +example : 0 ≤ max (0:ℤ) (-3) := by positivity + +example : 0 ≤ max (-3 : ℤ) 5 := by positivity + +example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity + +example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity + +/- ## Tests that the tactic is agnostic on reversed inequalities -/ + +example {a : ℤ} (ha : a > 0) : 0 ≤ a := by positivity + +example {a : ℤ} (ha : 0 < a) : a ≥ 0 := by positivity + +example {a : ℤ} (ha : a > 0) : a ≥ 0 := by positivity From aa270c3d7b51c8b2c516c555bd04069e3a8c4fda Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 26 Jul 2022 16:40:46 +0000 Subject: [PATCH 071/173] feat(measure_theory/function/conditional_expectation): pull-out property of the conditional expectation (#15274) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove this result: ```lean lemma condexp_strongly_measurable_mul {f g : α → ℝ} (hf : strongly_measurable[m] f) (hfg : integrable (f * g) μ) (hg : integrable g μ) : μ[f * g | m] =ᵐ[μ] f * μ[g | m] := ``` This could be extended beyond multiplication, to any bounded bilinear map, but we leave this to a future PR. For now we only prove the real multiplication case, which is needed for #14909 and #14933. --- .../function/conditional_expectation.lean | 207 ++++++++++++++++++ src/measure_theory/function/l1_space.lean | 10 + .../function/strongly_measurable.lean | 124 ++++++++++- src/measure_theory/integral/set_integral.lean | 28 +++ 4 files changed, 358 insertions(+), 11 deletions(-) diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index 2fd940dcc4024..9c14580e7f849 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -2145,6 +2145,50 @@ begin ((condexp_L1_mono hf hg hfg).trans_eq (condexp_ae_eq_condexp_L1 hm _).symm), end +/-- **Lebesgue dominated convergence theorem**: sufficient conditions under which almost + everywhere convergence of a sequence of functions implies the convergence of their image by + `condexp_L1`. -/ +lemma tendsto_condexp_L1_of_dominated_convergence (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + {fs : ℕ → α → F'} {f : α → F'} (bound_fs : α → ℝ) + (hfs_meas : ∀ n, ae_strongly_measurable (fs n) μ) (h_int_bound_fs : integrable bound_fs μ) + (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ∥fs n x∥ ≤ bound_fs x) + (hfs : ∀ᵐ x ∂μ, tendsto (λ n, fs n x) at_top (𝓝 (f x))) : + tendsto (λ n, condexp_L1 hm μ (fs n)) at_top (𝓝 (condexp_L1 hm μ f)) := +tendsto_set_to_fun_of_dominated_convergence _ bound_fs hfs_meas h_int_bound_fs hfs_bound hfs + +/-- If two sequences of functions have a.e. equal conditional expectations at each step, converge +and verify dominated convergence hypotheses, then the conditional expectations of their limits are +a.e. equal. -/ +lemma tendsto_condexp_unique (fs gs : ℕ → α → F') (f g : α → F') + (hfs_int : ∀ n, integrable (fs n) μ) (hgs_int : ∀ n, integrable (gs n) μ) + (hfs : ∀ᵐ x ∂μ, tendsto (λ n, fs n x) at_top (𝓝 (f x))) + (hgs : ∀ᵐ x ∂μ, tendsto (λ n, gs n x) at_top (𝓝 (g x))) + (bound_fs : α → ℝ) (h_int_bound_fs : integrable bound_fs μ) + (bound_gs : α → ℝ) (h_int_bound_gs : integrable bound_gs μ) + (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ∥fs n x∥ ≤ bound_fs x) + (hgs_bound : ∀ n, ∀ᵐ x ∂μ, ∥gs n x∥ ≤ bound_gs x) + (hfg : ∀ n, μ[fs n | m] =ᵐ[μ] μ[gs n | m]) : + μ[f | m] =ᵐ[μ] μ[g | m] := +begin + by_cases hm : m ≤ m0, swap, { simp_rw condexp_of_not_le hm, }, + by_cases hμm : sigma_finite (μ.trim hm), swap, { simp_rw condexp_of_not_sigma_finite hm hμm, }, + haveI : sigma_finite (μ.trim hm) := hμm, + refine (condexp_ae_eq_condexp_L1 hm f).trans ((condexp_ae_eq_condexp_L1 hm g).trans _).symm, + rw ← Lp.ext_iff, + have hn_eq : ∀ n, condexp_L1 hm μ (gs n) = condexp_L1 hm μ (fs n), + { intros n, + ext1, + refine (condexp_ae_eq_condexp_L1 hm (gs n)).symm.trans ((hfg n).symm.trans _), + exact (condexp_ae_eq_condexp_L1 hm (fs n)), }, + have hcond_fs : tendsto (λ n, condexp_L1 hm μ (fs n)) at_top (𝓝 (condexp_L1 hm μ f)), + from tendsto_condexp_L1_of_dominated_convergence hm _ (λ n, (hfs_int n).1) h_int_bound_fs + hfs_bound hfs, + have hcond_gs : tendsto (λ n, condexp_L1 hm μ (gs n)) at_top (𝓝 (condexp_L1 hm μ g)), + from tendsto_condexp_L1_of_dominated_convergence hm _ (λ n, (hgs_int n).1) h_int_bound_gs + hgs_bound hgs, + exact tendsto_nhds_unique_of_eventually_eq hcond_gs hcond_fs (eventually_of_forall hn_eq), +end + section indicator lemma condexp_ae_eq_restrict_zero (hs : measurable_set[m] s) (hf : f =ᵐ[μ.restrict s] 0) : @@ -2313,6 +2357,169 @@ end end indicator +section pull_out +-- TODO: this section could be generalized beyond multiplication, to any bounded bilinear map. + +/-- Auxiliary lemma for `condexp_measurable_mul`. -/ +lemma condexp_strongly_measurable_simple_func_mul (hm : m ≤ m0) + (f : @simple_func α m ℝ) {g : α → ℝ} (hg : integrable g μ) : + μ[f * g | m] =ᵐ[μ] f * μ[g | m] := +begin + have : ∀ s c (f : α → ℝ), set.indicator s (function.const α c) * f = s.indicator (c • f), + { intros s c f, + ext1 x, + by_cases hx : x ∈ s, + { simp only [hx, pi.mul_apply, set.indicator_of_mem, pi.smul_apply, algebra.id.smul_eq_mul] }, + { simp only [hx, pi.mul_apply, set.indicator_of_not_mem, not_false_iff, zero_mul], }, }, + refine @simple_func.induction _ _ m _ _ (λ c s hs, _) (λ g₁ g₂ h_disj h_eq₁ h_eq₂, _) f, + { simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, + simple_func.coe_zero, set.piecewise_eq_indicator], + rw [this, this], + refine (condexp_indicator (hg.smul c) hs).trans _, + filter_upwards [@condexp_smul α ℝ ℝ _ _ _ _ _ m m0 μ c g] with x hx, + classical, + simp_rw [set.indicator_apply, hx], }, + { have h_add := @simple_func.coe_add _ _ m _ g₁ g₂, + calc μ[⇑(g₁ + g₂) * g|m] =ᵐ[μ] μ[(⇑g₁ + ⇑g₂) * g|m] : + by { refine condexp_congr_ae (eventually_eq.mul _ eventually_eq.rfl), rw h_add, } + ... =ᵐ[μ] μ[⇑g₁ * g|m] + μ[⇑g₂ * g|m] : + by { rw add_mul, exact condexp_add (hg.simple_func_mul' hm _) (hg.simple_func_mul' hm _), } + ... =ᵐ[μ] ⇑g₁ * μ[g|m] + ⇑g₂ * μ[g|m] : eventually_eq.add h_eq₁ h_eq₂ + ... =ᵐ[μ] ⇑(g₁ + g₂) * μ[g|m] : by rw [h_add, add_mul], }, +end + +lemma condexp_strongly_measurable_mul_of_bound (hm : m ≤ m0) [is_finite_measure μ] + {f g : α → ℝ} (hf : strongly_measurable[m] f) (hg : integrable g μ) (c : ℝ) + (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + μ[f * g | m] =ᵐ[μ] f * μ[g | m] := +begin + let fs := hf.approx_bounded c, + have hfs_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, fs n x) at_top (𝓝 (f x)), + from hf.tendsto_approx_bounded_ae hf_bound, + by_cases hμ : μ = 0, + { simp only [hμ, ae_zero], }, + haveI : μ.ae.ne_bot, by simp only [hμ, ae_ne_bot, ne.def, not_false_iff], + have hc : 0 ≤ c, + { have h_exists : ∃ x, ∥f x∥ ≤ c := eventually.exists hf_bound, + exact (norm_nonneg _).trans h_exists.some_spec, }, + have hfs_bound : ∀ n x, ∥fs n x∥ ≤ c := hf.norm_approx_bounded_le hc, + have hn_eq : ∀ n, μ[fs n * g | m] =ᵐ[μ] fs n * μ[g | m], + from λ n, condexp_strongly_measurable_simple_func_mul hm _ hg, + have : μ[f * μ[g|m]|m] = f * μ[g|m], + { refine condexp_of_strongly_measurable hm (hf.mul strongly_measurable_condexp) _, + exact integrable_condexp.bdd_mul' ((hf.mono hm).ae_strongly_measurable) hf_bound, }, + rw ← this, + refine tendsto_condexp_unique (λ n x, fs n x * g x) (λ n x, fs n x * μ[g|m] x) (f * g) + (f * μ[g|m]) _ _ _ _ (λ x, c * ∥g x∥) _ (λ x, c * ∥μ[g|m] x∥) _ _ _ _, + { exact λ n, hg.bdd_mul' + ((simple_func.strongly_measurable (fs n)).mono hm).ae_strongly_measurable + (eventually_of_forall (hfs_bound n)), }, + { exact λ n, integrable_condexp.bdd_mul' + ((simple_func.strongly_measurable (fs n)).mono hm).ae_strongly_measurable + (eventually_of_forall (hfs_bound n)), }, + { filter_upwards [hfs_tendsto] with x hx, + rw pi.mul_apply, + exact tendsto.mul hx tendsto_const_nhds, }, + { filter_upwards [hfs_tendsto] with x hx, + rw pi.mul_apply, + exact tendsto.mul hx tendsto_const_nhds, }, + { exact hg.norm.const_mul c, }, + { exact integrable_condexp.norm.const_mul c, }, + { refine λ n, eventually_of_forall (λ x, _), + exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)), }, + { refine λ n, eventually_of_forall (λ x, _), + exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)), }, + { intros n, + simp_rw ← pi.mul_apply, + refine (condexp_strongly_measurable_simple_func_mul hm _ hg).trans _, + rw condexp_of_strongly_measurable hm + ((simple_func.strongly_measurable _).mul strongly_measurable_condexp) _, + { apply_instance, }, + { apply_instance, }, + exact integrable_condexp.bdd_mul' + ((simple_func.strongly_measurable (fs n)).mono hm).ae_strongly_measurable + (eventually_of_forall (hfs_bound n)), }, +end + +lemma condexp_strongly_measurable_mul_of_bound₀ (hm : m ≤ m0) [is_finite_measure μ] + {f g : α → ℝ} (hf : ae_strongly_measurable' m f μ) (hg : integrable g μ) (c : ℝ) + (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + μ[f * g | m] =ᵐ[μ] f * μ[g | m] := +begin + have : μ[f * g | m] =ᵐ[μ] μ[hf.mk f * g | m], + from condexp_congr_ae (eventually_eq.mul hf.ae_eq_mk eventually_eq.rfl), + refine this.trans _, + have : f * μ[g | m] =ᵐ[μ] hf.mk f * μ[g | m] := eventually_eq.mul hf.ae_eq_mk eventually_eq.rfl, + refine eventually_eq.trans _ this.symm, + refine condexp_strongly_measurable_mul_of_bound hm hf.strongly_measurable_mk hg c _, + filter_upwards [hf_bound, hf.ae_eq_mk] with x hxc hx_eq, + rw ← hx_eq, + exact hxc, +end + +/-- Pull-out property of the conditional expectation. -/ +lemma condexp_strongly_measurable_mul {f g : α → ℝ} (hf : strongly_measurable[m] f) + (hfg : integrable (f * g) μ) (hg : integrable g μ) : + μ[f * g | m] =ᵐ[μ] f * μ[g | m] := +begin + by_cases hm : m ≤ m0, swap, { simp_rw condexp_of_not_le hm, rw mul_zero, }, + by_cases hμm : sigma_finite (μ.trim hm), + swap, { simp_rw condexp_of_not_sigma_finite hm hμm, rw mul_zero, }, + haveI : sigma_finite (μ.trim hm) := hμm, + obtain ⟨sets, sets_prop, h_univ⟩ := hf.exists_spanning_measurable_set_norm_le hm μ, + simp_rw forall_and_distrib at sets_prop, + obtain ⟨h_meas, h_finite, h_norm⟩ := sets_prop, + + suffices : ∀ n, ∀ᵐ x ∂μ, x ∈ sets n → μ[f * g|m] x = f x * μ[g|m] x, + { rw ← ae_all_iff at this, + filter_upwards [this] with x hx, + rw pi.mul_apply, + obtain ⟨i, hi⟩ : ∃ i, x ∈ sets i, + { have h_mem : x ∈ ⋃ i, sets i, + { rw h_univ, exact set.mem_univ _, }, + simpa using h_mem, }, + exact hx i hi, }, + refine λ n, ae_imp_of_ae_restrict _, + suffices : (μ.restrict (sets n))[f * g | m] + =ᵐ[μ.restrict (sets n)] f * (μ.restrict (sets n))[g | m], + { simp_rw ← pi.mul_apply, + refine (condexp_restrict_ae_eq_restrict hm (h_meas n) hfg).symm.trans _, + exact this.trans (eventually_eq.rfl.mul (condexp_restrict_ae_eq_restrict hm (h_meas n) hg)), }, + suffices : (μ.restrict (sets n))[((sets n).indicator f) * g | m] + =ᵐ[μ.restrict (sets n)] ((sets n).indicator f) * (μ.restrict (sets n))[g | m], + { refine eventually_eq.trans _ (this.trans _), + { exact condexp_congr_ae + ((indicator_ae_eq_restrict (hm _ (h_meas n))).symm.mul eventually_eq.rfl), }, + { exact (indicator_ae_eq_restrict (hm _ (h_meas n))).mul eventually_eq.rfl, }, }, + haveI : is_finite_measure (μ.restrict (sets n)), + { constructor, + rw measure.restrict_apply_univ, + exact h_finite n, }, + refine condexp_strongly_measurable_mul_of_bound hm (hf.indicator (h_meas n)) hg.integrable_on n _, + refine eventually_of_forall (λ x, _), + by_cases hxs : x ∈ sets n, + { simp only [hxs, set.indicator_of_mem], + exact h_norm n x hxs, }, + { simp only [hxs, set.indicator_of_not_mem, not_false_iff, _root_.norm_zero, nat.cast_nonneg], }, +end + +/-- Pull-out property of the conditional expectation. -/ +lemma condexp_strongly_measurable_mul₀ {f g : α → ℝ} (hf : ae_strongly_measurable' m f μ) + (hfg : integrable (f * g) μ) (hg : integrable g μ) : + μ[f * g | m] =ᵐ[μ] f * μ[g | m] := +begin + have : μ[f * g | m] =ᵐ[μ] μ[hf.mk f * g | m], + from condexp_congr_ae (eventually_eq.mul hf.ae_eq_mk eventually_eq.rfl), + refine this.trans _, + have : f * μ[g | m] =ᵐ[μ] hf.mk f * μ[g | m] := eventually_eq.mul hf.ae_eq_mk eventually_eq.rfl, + refine eventually_eq.trans _ this.symm, + refine condexp_strongly_measurable_mul hf.strongly_measurable_mk _ hg, + refine (integrable_congr _).mp hfg, + exact eventually_eq.mul hf.ae_eq_mk eventually_eq.rfl, +end + +end pull_out + section real lemma rn_deriv_ae_eq_condexp {hm : m ≤ m0} [hμm : sigma_finite (μ.trim hm)] {f : α → ℝ} diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index a9fce93c233dc..e3b03c584789b 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -903,6 +903,16 @@ 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 μ) + (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + integrable (λ x, f x * g x) μ := +begin + refine integrable.mono' (hg.norm.smul c) (hf.mul hg.1) _, + filter_upwards [hf_bound] with x hx, + rw [pi.smul_apply, smul_eq_mul], + exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right hx (norm_nonneg _)), +end + end normed_space section normed_space_over_complete_field diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index 05e68ddec5413..f7e62a731a501 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -181,6 +181,75 @@ protected lemma tendsto_approx {m : measurable_space α} (hf : strongly_measurab ∀ x, tendsto (λ n, hf.approx n x) at_top (𝓝 (f x)) := hf.some_spec +/-- Similar to `strongly_measurable.approx`, but enforces that the norm of every function in the +sequence is less than `c` everywhere. If `∥f x∥ ≤ c` this sequence of simple functions verifies +`tendsto (λ n, hf.approx_bounded n x) at_top (𝓝 (f x))`. -/ +noncomputable +def approx_bounded {m : measurable_space α} + [has_norm β] [has_smul ℝ β] (hf : strongly_measurable f) (c : ℝ) : + ℕ → simple_func α β := +λ n, (hf.approx n).map (λ x, (min 1 (c / ∥x∥)) • x) + +lemma tendsto_approx_bounded_of_norm_le {β} {f : α → β} [normed_add_comm_group β] [normed_space ℝ β] + {m : measurable_space α} (hf : strongly_measurable[m] f) {c : ℝ} {x : α} (hfx : ∥f x∥ ≤ c) : + tendsto (λ n, hf.approx_bounded c n x) at_top (𝓝 (f x)) := +begin + have h_tendsto := hf.tendsto_approx x, + simp only [strongly_measurable.approx_bounded, simple_func.coe_map, function.comp_app], + by_cases hfx0 : ∥f x∥ = 0, + { rw norm_eq_zero at hfx0, + rw hfx0 at h_tendsto ⊢, + have h_tendsto_norm : tendsto (λ n, ∥hf.approx n x∥) at_top (𝓝 0), + { convert h_tendsto.norm, + rw norm_zero, }, + refine squeeze_zero_norm (λ n, _) h_tendsto_norm, + calc ∥min 1 (c / ∥hf.approx n x∥) • hf.approx n x∥ + = ∥min 1 (c / ∥hf.approx n x∥)∥ * ∥hf.approx n x∥ : norm_smul _ _ + ... ≤ ∥(1 : ℝ)∥ * ∥hf.approx n x∥ : + begin + refine mul_le_mul_of_nonneg_right _ (norm_nonneg _), + rw [norm_one, real.norm_of_nonneg], + { exact min_le_left _ _, }, + { exact le_min zero_le_one + (div_nonneg ((norm_nonneg _).trans hfx) (norm_nonneg _)), }, + end + ... = ∥hf.approx n x∥ : by rw [norm_one, one_mul], }, + rw ← one_smul ℝ (f x), + refine tendsto.smul _ h_tendsto, + have : min 1 (c / ∥f x∥) = 1, + { rw [min_eq_left_iff, one_le_div (lt_of_le_of_ne (norm_nonneg _) (ne.symm hfx0))], + exact hfx, }, + nth_rewrite 0 this.symm, + refine tendsto.min tendsto_const_nhds _, + refine tendsto.div tendsto_const_nhds h_tendsto.norm hfx0, +end + +lemma tendsto_approx_bounded_ae {β} {f : α → β} [normed_add_comm_group β] [normed_space ℝ β] + {m m0 : measurable_space α} {μ : measure α} + (hf : strongly_measurable[m] f) {c : ℝ} + (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + ∀ᵐ x ∂μ, tendsto (λ n, hf.approx_bounded c n x) at_top (𝓝 (f x)) := +by filter_upwards [hf_bound] with x hfx using tendsto_approx_bounded_of_norm_le hf hfx + +lemma norm_approx_bounded_le {β} {f : α → β} [seminormed_add_comm_group β] [normed_space ℝ β] + {m : measurable_space α} {c : ℝ} (hf : strongly_measurable[m] f) (hc : 0 ≤ c) (n : ℕ) (x : α) : + ∥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 _, + 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, }, + cases le_total (∥hf.approx n x∥) c, + { rw min_eq_left _, + { simpa only [norm_one, one_mul] using h, }, + { rwa one_le_div (lt_of_le_of_ne (norm_nonneg _) (ne.symm h0)), }, }, + { rw min_eq_right _, + { rw [norm_div, norm_norm, mul_comm, mul_div, div_eq_mul_inv, mul_comm, ← mul_assoc, + inv_mul_cancel h0, one_mul, real.norm_of_nonneg hc], }, + { rwa div_le_one (lt_of_le_of_ne (norm_nonneg _) (ne.symm h0)), }, }, +end + end basic_properties_in_any_topological_space lemma fin_strongly_measurable_of_set_sigma_finite [topological_space β] [has_zero β] @@ -694,18 +763,18 @@ protected lemma dist {m : measurable_space α} {β : Type*} [pseudo_metric_space strongly_measurable (λ x, dist (f x) (g x)) := continuous_dist.comp_strongly_measurable (hf.prod_mk hg) -protected lemma norm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} - (hf : strongly_measurable f) : +protected lemma norm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] + {f : α → β} (hf : strongly_measurable f) : strongly_measurable (λ x, ∥f x∥) := continuous_norm.comp_strongly_measurable hf -protected lemma nnnorm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} - (hf : strongly_measurable f) : +protected lemma nnnorm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] + {f : α → β} (hf : strongly_measurable f) : strongly_measurable (λ x, ∥f x∥₊) := continuous_nnnorm.comp_strongly_measurable hf -protected lemma ennnorm {m : measurable_space α} {β : Type*} [normed_add_comm_group β] {f : α → β} - (hf : strongly_measurable f) : +protected lemma ennnorm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] + {f : α → β} (hf : strongly_measurable f) : measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := (ennreal.continuous_coe.comp_strongly_measurable hf.nnnorm).measurable @@ -830,6 +899,39 @@ begin exact hg_seq_tendsto x, end +/-- If a function `f` is strongly measurable w.r.t. a sub-σ-algebra `m` and the measure is σ-finite +on `m`, then there exists spanning measurable sets with finite measure on which `f` has bounded +norm. In particular, `f` is integrable on each of those sets. -/ +lemma exists_spanning_measurable_set_norm_le [seminormed_add_comm_group β] + {m m0 : measurable_space α} (hm : m ≤ m0) (hf : strongly_measurable[m] f) (μ : measure α) + [sigma_finite (μ.trim hm)] : + ∃ s : ℕ → set α, (∀ n, measurable_set[m] (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, ∥f x∥ ≤ n) + ∧ (⋃ i, s i) = set.univ := +begin + let sigma_finite_sets := spanning_sets (μ.trim hm), + let norm_sets := λ (n : ℕ), {x | ∥f x∥ ≤ n}, + have norm_sets_spanning : (⋃ n, norm_sets n) = set.univ, + { ext1 x, simp only [set.mem_Union, set.mem_set_of_eq, set.mem_univ, iff_true], + exact ⟨⌈∥f x∥⌉₊, nat.le_ceil (∥f x∥)⟩, }, + let sets := λ n, sigma_finite_sets n ∩ norm_sets n, + have h_meas : ∀ n, measurable_set[m] (sets n), + { refine λ n, measurable_set.inter _ _, + { exact measurable_spanning_sets (μ.trim hm) n, }, + { exact hf.norm.measurable_set_le strongly_measurable_const, }, }, + have h_finite : ∀ n, μ (sets n) < ∞, + { refine λ n, (measure_mono (set.inter_subset_left _ _)).trans_lt _, + exact (le_trim hm).trans_lt (measure_spanning_sets_lt_top (μ.trim hm) n), }, + refine ⟨sets, λ n, ⟨h_meas n, h_finite n, _⟩, _⟩, + { exact λ x hx, hx.2, }, + { have : (⋃ i, sigma_finite_sets i ∩ norm_sets i) + = (⋃ i, sigma_finite_sets i) ∩ (⋃ i, norm_sets i), + { refine set.Union_inter_of_monotone (monotone_spanning_sets (μ.trim hm)) (λ i j hij x, _), + simp only [norm_sets, set.mem_set_of_eq], + refine λ hif, hif.trans _, + exact_mod_cast hij, }, + rw [this, norm_sets_spanning, Union_spanning_sets (μ.trim hm), set.inter_univ], }, +end + end strongly_measurable /-! ## Finitely strongly measurable functions -/ @@ -1266,22 +1368,22 @@ protected lemma dist {β : Type*} [pseudo_metric_space β] {f g : α → β} ae_strongly_measurable (λ x, dist (f x) (g x)) μ := continuous_dist.comp_ae_strongly_measurable (hf.prod_mk hg) -protected lemma norm {β : Type*} [normed_add_comm_group β] {f : α → β} +protected lemma norm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : ae_strongly_measurable (λ x, ∥f x∥) μ := continuous_norm.comp_ae_strongly_measurable hf -protected lemma nnnorm {β : Type*} [normed_add_comm_group β] {f : α → β} +protected lemma nnnorm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : ae_strongly_measurable (λ x, ∥f x∥₊) μ := continuous_nnnorm.comp_ae_strongly_measurable hf -protected lemma ennnorm {β : Type*} [normed_add_comm_group β] {f : α → β} +protected lemma ennnorm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := (ennreal.continuous_coe.comp_ae_strongly_measurable hf.nnnorm).ae_measurable -protected lemma edist {β : Type*} [normed_add_comm_group β] {f g : α → β} +protected lemma edist {β : Type*} [seminormed_add_comm_group β] {f g : α → β} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : ae_measurable (λ a, edist (f a) (g a)) μ := (continuous_edist.comp_ae_strongly_measurable (hf.prod_mk hg)).ae_measurable @@ -1712,7 +1814,7 @@ end ae_fin_strongly_measurable section second_countable_topology variables {G : Type*} {p : ℝ≥0∞} {m m0 : measurable_space α} {μ : measure α} - [normed_add_comm_group G] [measurable_space G] [borel_space G] [second_countable_topology G] + [seminormed_add_comm_group G] [measurable_space G] [borel_space G] [second_countable_topology G] {f : α → G} /-- In a space with second countable topology and a sigma-finite measure, `fin_strongly_measurable` diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 44c48877bb234..37eb88792a7d1 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -1002,3 +1002,31 @@ begin end end thickened_indicator + +section bilinear_map +namespace measure_theory + +variables {f : β → ℝ} {m m0 : measurable_space β} {μ : measure β} + +lemma integrable.simple_func_mul (g : simple_func β ℝ) (hf : integrable f μ) : + integrable (g * f) μ := +begin + refine simple_func.induction (λ c s hs, _) (λ g₁ g₂ h_disj h_int₁ h_int₂, + (h_int₁.add h_int₂).congr (by rw [simple_func.coe_add, add_mul])) g, + simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, + simple_func.coe_zero, set.piecewise_eq_indicator], + have : set.indicator s (function.const β c) * f = s.indicator (c • f), + { ext1 x, + by_cases hx : x ∈ s, + { simp only [hx, pi.mul_apply, set.indicator_of_mem, pi.smul_apply, algebra.id.smul_eq_mul] }, + { simp only [hx, pi.mul_apply, set.indicator_of_not_mem, not_false_iff, zero_mul], }, }, + rw [this, integrable_indicator_iff hs], + exact (hf.smul c).integrable_on, +end + +lemma integrable.simple_func_mul' (hm : m ≤ m0) (g : @simple_func β m ℝ) (hf : integrable f μ) : + integrable (g * f) μ := +by { rw ← simple_func.coe_to_larger_space_eq hm g, exact hf.simple_func_mul (g.to_larger_space hm) } + +end measure_theory +end bilinear_map From f34e3402254ebfefadbdd6566ec8e2c992707d95 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Tue, 26 Jul 2022 16:40:47 +0000 Subject: [PATCH 072/173] feat(measure_theory/constructions/borel_space): the set of points for which a measurable sequence of functions converges is measurable (#15307) Co-authored-by: RemyDegenne --- src/measure_theory/constructions/polish.lean | 34 ++++++++++++++++++-- src/order/filter/bases.lean | 5 +++ 2 files changed, 36 insertions(+), 3 deletions(-) diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 03147e5777603..9558e51250858 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 +open_locale topological_space measure_theory filter variables {α : Type*} [topological_space α] {ι : Type*} @@ -580,10 +580,11 @@ begin { rwa inj_on_iff_injective at f_inj } end -variables [measurable_space γ] [borel_space γ] + +variables [measurable_space γ] [hγb : borel_space γ] {β : Type*} [tβ : topological_space β] [t2_space β] [measurable_space β] [borel_space β] {s : set γ} {f : γ → β} -include tβ +include tβ hγb /-- The Lusin-Souslin theorem: if `s` is Borel-measurable in a Polish space, then its image under a continuous injective map is also Borel-measurable. -/ @@ -680,4 +681,31 @@ begin simp only [id.def, image_id'], end +omit hγb + +/-- The set of points for which a measurable sequence of functions converges is measurable. -/ +@[measurability] lemma measurable_set_exists_tendsto + [hγ : opens_measurable_space γ] [countable ι] {l : filter ι} + [l.is_countably_generated] {f : ι → β → γ} (hf : ∀ i, measurable (f i)) : + measurable_set {x | ∃ c, tendsto (λ n, f n x) l (𝓝 c)} := +begin + by_cases hl : l.ne_bot, + swap, { rw not_ne_bot at hl, simp [hl] }, + letI := upgrade_polish_space γ, + rcases l.exists_antitone_basis with ⟨u, hu⟩, + simp_rw ← cauchy_map_iff_exists_tendsto, + change measurable_set {x | _ ∧ _}, + have : ∀ x, ((map (λ i, f i x) l) ×ᶠ (map (λ i, f i x) l)).has_antitone_basis + (λ n, ((λ i, f i x) '' u n) ×ˢ ((λ i, f i x) '' u n)) := λ x, hu.map.prod hu.map, + simp_rw [and_iff_right (hl.map _), filter.has_basis.le_basis_iff (this _).to_has_basis + metric.uniformity_basis_dist_inv_nat_succ, set.set_of_forall], + refine measurable_set.bInter set.countable_univ (λ K _, _), + simp_rw set.set_of_exists, + refine measurable_set.bUnion set.countable_univ (λ N hN, _), + simp_rw [prod_image_image_eq, image_subset_iff, prod_subset_iff, set.set_of_forall], + exact measurable_set.bInter (to_countable (u N)) (λ i _, + measurable_set.bInter (to_countable (u N)) (λ j _, + measurable_set_lt (measurable.dist (hf i) (hf j)) measurable_const)), +end + end measure_theory diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index a5a149a05c99b..994ab6953cd52 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -707,6 +707,11 @@ includes `s i` for some `i`, and `s` is decreasing. -/ extends has_basis l (λ _, true) s : Prop := (antitone : antitone s) +lemma has_antitone_basis.map {l : filter α} {s : ι'' → set α} {m : α → β} + (hf : has_antitone_basis l s) : + has_antitone_basis (map m l) (λ n, m '' s n) := +⟨has_basis.map _ hf.to_has_basis, λ i j hij, image_subset _ $ hf.2 hij⟩ + end same_type section two_types From 34a1e2c5d5f5564e2a8e184a356f991a75cf31db Mon Sep 17 00:00:00 2001 From: jackmckoen Date: Tue, 26 Jul 2022 18:06:37 +0000 Subject: [PATCH 073/173] feat(ring_theory/valuation/valuation_subring): define maximal ideal of valuation subring and provide basic API (#14656) This PR defines the unique maximal ideal of a valuation subring as a subsemigroup of the field. We prove a few equivalent conditions for two valuations to be equivalent, and we use this to show two valuation subrings are equivalent iff they have the same maximal ideal. Co-authored-by: Vierkantor Co-authored-by: jackmckoen <104791831+jackmckoen@users.noreply.github.com> --- src/ring_theory/valuation/basic.lean | 91 ++++++++++++--- .../valuation/valuation_subring.lean | 109 +++++++++++++++--- 2 files changed, 166 insertions(+), 34 deletions(-) diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 0456353ae971e..0174b62776a02 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -67,7 +67,7 @@ noncomputable theory open function ideal -variables {F R : Type*} -- This will be a ring, assumed commutative in some sections +variables {K F R : Type*} [division_ring K] section variables (F R) (Γ₀ : Type*) [linear_ordered_comm_monoid_with_zero Γ₀] [ring R] @@ -170,12 +170,12 @@ lemma ext_iff {v₁ v₂ : valuation R Γ₀} : v₁ = v₂ ↔ ∀ r, v₁ r = def to_preorder : preorder R := preorder.lift v /-- If `v` is a valuation on a division ring then `v(x) = 0` iff `x = 0`. -/ -@[simp] lemma zero_iff [nontrivial Γ₀] {K : Type*} [division_ring K] - (v : valuation K Γ₀) {x : K} : v x = 0 ↔ x = 0 := +@[simp] lemma zero_iff [nontrivial Γ₀] (v : valuation K Γ₀) {x : K} : + v x = 0 ↔ x = 0 := v.to_monoid_with_zero_hom.map_eq_zero -lemma ne_zero_iff [nontrivial Γ₀] {K : Type*} [division_ring K] - (v : valuation K Γ₀) {x : K} : v x ≠ 0 ↔ x ≠ 0 := +lemma ne_zero_iff [nontrivial Γ₀] (v : valuation K Γ₀) {x : K} : + v x ≠ 0 ↔ x ≠ 0 := v.to_monoid_with_zero_hom.map_ne_zero theorem unit_map_eq (u : Rˣ) : @@ -217,11 +217,11 @@ end monoid section group variables [linear_ordered_comm_group_with_zero Γ₀] {R} {Γ₀} (v : valuation R Γ₀) {x y z : R} -@[simp] lemma map_inv {K : Type*} [division_ring K] - (v : valuation K Γ₀) {x : K} : v x⁻¹ = (v x)⁻¹ := +@[simp] lemma map_inv (v : valuation K Γ₀) {x : K} : + v x⁻¹ = (v x)⁻¹ := v.to_monoid_with_zero_hom.map_inv x -@[simp] lemma map_zpow {K : Type*} [division_ring K] (v : valuation K Γ₀) {x : K} {n : ℤ} : +@[simp] lemma map_zpow (v : valuation K Γ₀) {x : K} {n : ℤ} : v (x^n) = (v x)^n := v.to_monoid_with_zero_hom.map_zpow x n @@ -293,6 +293,10 @@ begin simpa only [v.map_one, v.map_neg] using v.map_add_eq_of_lt_left h end +lemma one_lt_val_iff (v : valuation K Γ₀) {x : K} (h : x ≠ 0) : + 1 < v x ↔ v x⁻¹ < 1 := +by simpa using (inv_lt_inv₀ (v.ne_zero_iff.2 h) one_ne_zero).symm + /-- The subgroup of elements whose valuation is less than a certain unit.-/ def lt_add_subgroup (v : valuation R Γ₀) (γ : Γ₀ˣ) : add_subgroup R := { carrier := {x | v x < γ}, @@ -358,7 +362,6 @@ lemma is_equiv_of_map_strict_mono [linear_ordered_comm_monoid_with_zero Γ₀] lemma is_equiv_of_val_le_one [linear_ordered_comm_group_with_zero Γ₀] [linear_ordered_comm_group_with_zero Γ'₀] - {K : Type*} [division_ring K] (v : valuation K Γ₀) (v' : valuation K Γ'₀) (h : ∀ {x:K}, v x ≤ 1 ↔ v' x ≤ 1) : v.is_equiv v' := begin @@ -382,7 +385,6 @@ end lemma is_equiv_iff_val_le_one [linear_ordered_comm_group_with_zero Γ₀] [linear_ordered_comm_group_with_zero Γ'₀] - {K : Type*} [division_ring K] (v : valuation K Γ₀) (v' : valuation K Γ'₀) : v.is_equiv v' ↔ ∀ {x : K}, v x ≤ 1 ↔ v' x ≤ 1 := ⟨λ h x, by simpa using h x 1, is_equiv_of_val_le_one _ _⟩ @@ -390,7 +392,6 @@ lemma is_equiv_iff_val_le_one lemma is_equiv_iff_val_eq_one [linear_ordered_comm_group_with_zero Γ₀] [linear_ordered_comm_group_with_zero Γ'₀] - {K : Type*} [division_ring K] (v : valuation K Γ₀) (v' : valuation K Γ'₀) : v.is_equiv v' ↔ ∀ {x : K}, v x = 1 ↔ v' x = 1 := begin @@ -419,6 +420,61 @@ begin { rw ← h at hx', exact le_of_eq hx' } } } end +lemma is_equiv_iff_val_lt_one + [linear_ordered_comm_group_with_zero Γ₀] + [linear_ordered_comm_group_with_zero Γ'₀] + (v : valuation K Γ₀) (v' : valuation K Γ'₀) : + v.is_equiv v' ↔ ∀ {x : K}, v x < 1 ↔ v' x < 1 := +begin + split, + { intros h x, + simp only [lt_iff_le_and_ne, and_congr ((is_equiv_iff_val_le_one _ _).1 h) + ((is_equiv_iff_val_eq_one _ _).1 h).not] }, + { rw is_equiv_iff_val_eq_one, + intros h x, + by_cases hx : x = 0, { simp only [(zero_iff _).2 hx, zero_ne_one] }, + split, + { intro hh, + 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)) } }, + { 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)) } } } +end + +lemma is_equiv_iff_val_sub_one_lt_one + [linear_ordered_comm_group_with_zero Γ₀] + [linear_ordered_comm_group_with_zero Γ'₀] + (v : valuation K Γ₀) (v' : valuation K Γ'₀) : + v.is_equiv v' ↔ ∀ {x : K}, v (x - 1) < 1 ↔ v' (x - 1) < 1 := +begin + rw is_equiv_iff_val_lt_one, + exact (equiv.sub_right 1).surjective.forall +end + +lemma is_equiv_tfae + [linear_ordered_comm_group_with_zero Γ₀] + [linear_ordered_comm_group_with_zero Γ'₀] + (v : valuation K Γ₀) (v' : valuation K Γ'₀) : + [v.is_equiv v', + ∀ {x}, v x ≤ 1 ↔ v' x ≤ 1, + ∀ {x}, v x = 1 ↔ v' x = 1, + ∀ {x}, v x < 1 ↔ v' x < 1, + ∀ {x}, v (x-1) < 1 ↔ v' (x-1) < 1].tfae := +begin + tfae_have : 1 ↔ 2, { apply is_equiv_iff_val_le_one }, + tfae_have : 1 ↔ 3, { apply is_equiv_iff_val_eq_one }, + tfae_have : 1 ↔ 4, { apply is_equiv_iff_val_lt_one }, + tfae_have : 1 ↔ 5, { apply is_equiv_iff_val_sub_one_lt_one }, + tfae_finish +end + end section supp @@ -605,12 +661,13 @@ valuation.ext_iff def to_preorder : preorder R := preorder.lift v /-- If `v` is an additive valuation on a division ring then `v(x) = ⊤` iff `x = 0`. -/ -@[simp] lemma top_iff [nontrivial Γ₀] {K : Type*} [division_ring K] - (v : add_valuation K Γ₀) {x : K} : v x = ⊤ ↔ x = 0 := +@[simp] lemma top_iff [nontrivial Γ₀] (v : add_valuation K Γ₀) {x : K} : + v x = ⊤ ↔ x = 0 := v.zero_iff -lemma ne_top_iff [nontrivial Γ₀] {K : Type*} [division_ring K] - (v : add_valuation K Γ₀) {x : K} : v x ≠ ⊤ ↔ x ≠ 0 := v.ne_zero_iff +lemma ne_top_iff [nontrivial Γ₀] (v : add_valuation K Γ₀) {x : K} : + v x ≠ ⊤ ↔ x ≠ 0 := +v.ne_zero_iff /-- A ring homomorphism `S → R` induces a map `add_valuation R Γ₀ → add_valuation S Γ₀`. -/ def comap {S : Type*} [ring S] (f : S →+* R) (v : add_valuation R Γ₀) : @@ -644,8 +701,8 @@ end monoid section group variables [linear_ordered_add_comm_group_with_top Γ₀] [ring R] (v : add_valuation R Γ₀) {x y z : R} -@[simp] lemma map_inv {K : Type*} [division_ring K] - (v : add_valuation K Γ₀) {x : K} : v x⁻¹ = - (v x) := +@[simp] lemma map_inv (v : add_valuation K Γ₀) {x : K} : + v x⁻¹ = - (v x) := v.map_inv lemma map_units_inv (x : Rˣ) : v (x⁻¹ : Rˣ) = - (v x) := diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 2ffc52f0df5e5..f26a3abd9eb69 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -72,9 +72,9 @@ instance : valuation_ring A := by_cases (b : K) = 0, { use 0, left, ext, simp [h] }, by_cases (a : K) = 0, { use 0, right, ext, simp [h] }, cases A.mem_or_inv_mem (a/b) with hh hh, - { use ⟨a/b,hh⟩, right, ext, field_simp, ring }, + { use ⟨a/b, hh⟩, right, ext, field_simp, ring }, { rw (show (a/b : K)⁻¹ = b/a, by field_simp) at hh, - use ⟨b/a,hh⟩, left, ext, field_simp, ring }, + use ⟨b/a, hh⟩, left, ext, field_simp, ring }, end } instance : algebra A K := @@ -84,16 +84,16 @@ show algebra A.to_subring K, by apply_instance lemma algebra_map_apply (a : A) : algebra_map A K a = a := rfl instance : is_fraction_ring A K := -{ map_units := λ ⟨y,hy⟩, +{ map_units := λ ⟨y, hy⟩, (units.mk0 (y : K) (λ c, non_zero_divisors.ne_zero hy $ subtype.ext c)).is_unit, surj := λ z, begin - by_cases z = 0, { use (0,1), simp [h] }, + by_cases z = 0, { use (0, 1), simp [h] }, cases A.mem_or_inv_mem z with hh hh, - { use (⟨z,hh⟩,1), simp }, - { refine ⟨⟨1,⟨⟨_,hh⟩,_⟩⟩, mul_inv_cancel h⟩, + { use (⟨z, hh⟩, 1), simp }, + { refine ⟨⟨1, ⟨⟨_, hh⟩, _⟩⟩, mul_inv_cancel h⟩, 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⟩, + eq_iff_exists := λ a b, ⟨ λ h, ⟨1, by { ext, simpa using h }⟩, λ ⟨c, h⟩, congr_arg coe ((mul_eq_mul_right_iff.1 h).resolve_right (non_zero_divisors.ne_zero c.2)) ⟩ } /-- The value group of the valuation associated to `A`. -/ @@ -106,13 +106,13 @@ def valuation : valuation K A.value_group := valuation_ring.valuation A K instance inhabited_value_group : inhabited A.value_group := ⟨A.valuation 0⟩ lemma valuation_le_one (a : A) : A.valuation a ≤ 1 := -(valuation_ring.mem_integer_iff A K _).2 ⟨a,rfl⟩ +(valuation_ring.mem_integer_iff A K _).2 ⟨a, rfl⟩ lemma mem_of_valuation_le_one (x : K) (h : A.valuation x ≤ 1) : x ∈ A := -let ⟨a,ha⟩ := (valuation_ring.mem_integer_iff A K x).1 h in ha ▸ a.2 +let ⟨a, ha⟩ := (valuation_ring.mem_integer_iff A K x).1 h in ha ▸ a.2 lemma valuation_le_one_iff (x : K) : A.valuation x ≤ 1 ↔ x ∈ A := -⟨mem_of_valuation_le_one _ _, λ ha, A.valuation_le_one ⟨x,ha⟩⟩ +⟨mem_of_valuation_le_one _ _, λ ha, A.valuation_le_one ⟨x, ha⟩⟩ lemma valuation_eq_iff (x y : K) : A.valuation x = A.valuation y ↔ ∃ a : Aˣ, (a : K) * y = x := quotient.eq' @@ -179,7 +179,7 @@ subring.subtype R.to_subring /-- The canonical map on value groups induced by a coarsening of valuation rings. -/ def map_of_le (R S : valuation_subring K) (h : R ≤ S) : R.value_group →*₀ S.value_group := -{ to_fun := quotient.map' id $ λ x y ⟨u,hu⟩, ⟨units.map (R.inclusion S h).to_monoid_hom u, hu⟩, +{ to_fun := quotient.map' id $ λ x y ⟨u, hu⟩, ⟨units.map (R.inclusion S h).to_monoid_hom u, hu⟩, map_zero' := rfl, map_one' := rfl, map_mul' := by { rintro ⟨⟩ ⟨⟩, refl } } @@ -187,7 +187,7 @@ def map_of_le (R S : valuation_subring K) (h : R ≤ S) : @[mono] lemma monotone_map_of_le (R S : valuation_subring K) (h : R ≤ S) : monotone (R.map_of_le S h) := -by { rintros ⟨⟩ ⟨⟩ ⟨a,ha⟩, exact ⟨R.inclusion S h a, ha⟩ } +by { rintros ⟨⟩ ⟨⟩ ⟨a, ha⟩, exact ⟨R.inclusion S h a, ha⟩ } @[simp] lemma map_of_le_comp_valuation (R S : valuation_subring K) (h : R ≤ S) : @@ -209,7 +209,7 @@ def of_prime (A : valuation_subring K) (P : ideal A) [P.is_prime] : valuation_subring K := of_le A (localization.subalgebra.of_field K P.prime_compl $ le_non_zero_divisors_of_no_zero_divisors $ not_not_intro P.zero_mem).to_subring $ - λ a ha, subalgebra.algebra_map_mem _ (⟨a,ha⟩ : A) + λ a ha, subalgebra.algebra_map_mem _ (⟨a, ha⟩ : A) instance of_prime_algebra (A : valuation_subring K) (P : ideal A) [P.is_prime] : algebra A (A.of_prime P) := subalgebra.algebra _ @@ -223,7 +223,7 @@ by apply localization.subalgebra.is_localization_of_field K lemma le_of_prime (A : valuation_subring K) (P : ideal A) [P.is_prime] : A ≤ of_prime A P := -λ a ha, subalgebra.algebra_map_mem _ (⟨a,ha⟩ : A) +λ a ha, subalgebra.algebra_map_mem _ (⟨a, ha⟩ : A) lemma of_prime_valuation_eq_one_iff_mem_prime_compl (A : valuation_subring K) @@ -243,7 +243,7 @@ lemma of_prime_ideal_of_le (R S : valuation_subring K) (h : R ≤ S) : of_prime R (ideal_of_le R S h) = S := begin ext x, split, - { rintro ⟨a,r,hr,rfl⟩, apply mul_mem, { exact h a.2 }, + { rintro ⟨a, r, hr, rfl⟩, apply mul_mem, { exact h a.2 }, { rw [← valuation_le_one_iff, valuation.map_inv, ← inv_one, inv_le_inv₀], { exact not_lt.1 ((not_iff_not.2 $ valuation_lt_one_iff S _).1 hr) }, { intro hh, erw [valuation.zero_iff, subring.coe_eq_zero_iff] at hh, @@ -255,7 +255,7 @@ begin { use [1, x⁻¹, hr], split, { change (⟨x⁻¹, h hr⟩ : S) ∉ nonunits S, erw [mem_nonunits_iff, not_not], - apply is_unit_of_mul_eq_one _ (⟨x,hx⟩ : S), + apply is_unit_of_mul_eq_one _ (⟨x, hx⟩ : S), ext, field_simp }, { field_simp } } }, end @@ -407,7 +407,7 @@ lemma unit_group_le_unit_group {A B : valuation_subring K} : A.unit_group ≤ B.unit_group ↔ A ≤ B := begin split, - { rintros h x hx, + { intros h x hx, rw [← A.valuation_le_one_iff x, le_iff_lt_or_eq] at hx, by_cases h_1 : x = 0, { simp only [h_1, zero_mem] }, by_cases h_2 : 1 + x = 0, @@ -435,4 +435,79 @@ unit_group_order_embedding.strict_mono end unit_group +section nonunits + +/-- The nonunits of a valuation subring of `K`, as a subsemigroup of `K`-/ +def nonunits : subsemigroup K := +{ carrier := { x | A.valuation x < 1 }, + mul_mem' := λ a b ha hb, (mul_lt_mul₀ ha hb).trans_eq $ mul_one _ } + +lemma mem_nonunits_iff {x : K} : x ∈ A.nonunits ↔ A.valuation x < 1 := iff.refl _ + +lemma nonunits_injective : + function.injective (nonunits : valuation_subring K → subsemigroup _) := +λ A B h, begin + rw [← A.valuation_subring_valuation, ← B.valuation_subring_valuation, + ← valuation.is_equiv_iff_valuation_subring, valuation.is_equiv_iff_val_lt_one], + exact set_like.ext_iff.1 h, +end + +lemma nonunits_inj {A B : valuation_subring K} : A.nonunits = B.nonunits ↔ A = B := +nonunits_injective.eq_iff + +lemma nonunits_le_nonunits {A B : valuation_subring K} : + B.nonunits ≤ A.nonunits ↔ A ≤ B := +begin + split, + { intros h x hx, + by_cases h_1 : x = 0, { simp only [h_1, zero_mem] }, + rw [← valuation_le_one_iff, ← not_lt, valuation.one_lt_val_iff _ h_1] at hx ⊢, + by_contra h_2, from hx (h h_2) }, + { intros h x hx, + by_contra h_1, from not_lt.2 (monotone_map_of_le _ _ h (not_lt.1 h_1)) hx } +end + +/-- The map on valuation subrings to their nonunits is a dual order embedding. -/ +def nonunits_order_embedding : + valuation_subring K ↪o (subsemigroup K)ᵒᵈ := +{ to_fun := λ A, A.nonunits, + inj' := nonunits_injective, + map_rel_iff' := λ A B, nonunits_le_nonunits } + +variables {A} + + /-- The elements of `A.nonunits` are those of the maximal ideal of `A` after coercion to `K`. + +See also `mem_nonunits_iff_exists_mem_maximal_ideal`, which gets rid of the coercion to `K`, +at the expense of a more complicated right hand side. + -/ +theorem coe_mem_nonunits_iff {a : A} : (a : K) ∈ A.nonunits ↔ a ∈ local_ring.maximal_ideal A := +(valuation_lt_one_iff _ _).symm + +lemma nonunits_le : A.nonunits ≤ A.to_subring.to_submonoid.to_subsemigroup := +λ a ha, (A.valuation_le_one_iff _).mp (A.mem_nonunits_iff.mp ha).le + +lemma nonunits_subset : (A.nonunits : set K) ⊆ A := nonunits_le + + /-- The elements of `A.nonunits` are those of the maximal ideal of `A`. + +See also `coe_mem_nonunits_iff`, which has a simpler right hand side but requires the element +to be in `A` already. + -/ +theorem mem_nonunits_iff_exists_mem_maximal_ideal {a : K} : + a ∈ A.nonunits ↔ ∃ ha, (⟨a, ha⟩ : A) ∈ local_ring.maximal_ideal A := +⟨λ h, ⟨nonunits_subset h, coe_mem_nonunits_iff.mp h⟩, + λ ⟨ha, h⟩, coe_mem_nonunits_iff.mpr h⟩ + + /-- `A.nonunits` agrees with the maximal ideal of `A`, after taking its image in `K`. -/ +theorem image_maximal_ideal : (coe : A → K) '' local_ring.maximal_ideal A = A.nonunits := +begin + ext a, + simp only [set.mem_image, set_like.mem_coe, mem_nonunits_iff_exists_mem_maximal_ideal], + erw subtype.exists, + simp_rw [subtype.coe_mk, exists_and_distrib_right, exists_eq_right], +end + +end nonunits + end valuation_subring From 55db562bf871aeec2527351fa20e6b028a6346ac Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 18:06:38 +0000 Subject: [PATCH 074/173] chore(data/set/{function,countable}): extract a lemma from a proof (#15668) --- src/data/set/countable.lean | 8 ++------ src/data/set/function.lean | 5 +++++ 2 files changed, 7 insertions(+), 6 deletions(-) diff --git a/src/data/set/countable.lean b/src/data/set/countable.lean index e6a26e0cba795..db06ee87cbebb 100644 --- a/src/data/set/countable.lean +++ b/src/data/set/countable.lean @@ -44,11 +44,7 @@ countable_coe_iff.symm.trans (countable_iff_exists_injective s) on `s`. -/ lemma countable_iff_exists_inj_on {s : set α} : s.countable ↔ ∃ f : α → ℕ, inj_on f s := -set.countable_iff_exists_injective.trans -⟨λ ⟨f, hf⟩, ⟨λ a, if h : a ∈ s then f ⟨a, h⟩ else 0, - λ a as b bs h, congr_arg subtype.val $ - hf $ by simpa [as, bs] using h⟩, - λ ⟨f, hf⟩, ⟨_, inj_on_iff_injective.1 hf⟩⟩ +set.countable_iff_exists_injective.trans exists_inj_on_iff_injective.symm /-- Convert `set.countable s` to `encodable s` (noncomputable). -/ protected def countable.to_encodable {s : set α} : s.countable → encodable s := @@ -155,7 +151,7 @@ let ⟨g, hg⟩ := countable_iff_exists_inj_on.1 hs in countable_iff_exists_inj_on.2 ⟨g ∘ f, hg.comp hf (maps_to_image _ _)⟩ lemma countable_Union {t : ι → set α} [countable ι] (ht : ∀ i, (t i).countable) : - (⋃a, t a).countable := + (⋃ i, t i).countable := by { haveI := λ a, (ht a).to_subtype, rw Union_eq_range_psigma, apply countable_range } @[simp] lemma countable_Union_iff [countable ι] {t : ι → set α} : diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 5b6aafb7b4148..43fad68aed2e5 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -431,6 +431,11 @@ lemma inj_on_iff_injective : inj_on f s ↔ injective (s.restrict f) := alias inj_on_iff_injective ↔ inj_on.injective _ +lemma exists_inj_on_iff_injective [nonempty β] : + (∃ f : α → β, inj_on f s) ↔ ∃ f : s → β, injective f := +⟨λ ⟨f, hf⟩, ⟨_, hf.injective⟩, + λ ⟨f, hf⟩, by { lift f to α → β using trivial, exact ⟨f, inj_on_iff_injective.2 hf⟩ }⟩ + lemma inj_on_preimage {B : set (set β)} (hB : B ⊆ 𝒫 (range f)) : inj_on (preimage f) B := λ s hs t ht hst, (preimage_eq_preimage' (hB hs) (hB ht)).1 hst From cdabf21f000a77f80507e062c5818419c8d37494 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Tue, 26 Jul 2022 20:33:25 +0000 Subject: [PATCH 075/173] feat(set/intervals/monotone): add some monotonicity lemmas (#14896) Co-authored-by: Eric Wieser --- src/data/set/basic.lean | 36 +++++++++++++ src/data/set/intervals/monotone.lean | 77 ++++++++++++++++++++++++++++ src/order/succ_pred/basic.lean | 12 +++++ 3 files changed, 125 insertions(+) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 5eb4f220426b2..8e361d17937bb 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1779,6 +1779,42 @@ theorem subsingleton_of_image {α β : Type*} {f : α → β} (hf : function.inj theorem univ_eq_true_false : univ = ({true, false} : set Prop) := eq.symm $ eq_univ_of_forall $ classical.cases (by simp) (by simp) +section preorder + +variables [preorder α] [preorder β] (f : α → β) + +/-! ### Monotonicity on singletons -/ + +protected lemma subsingleton.monotone_on (h : s.subsingleton) : + monotone_on f s := +λ a ha b hb _, (congr_arg _ (h ha hb)).le + +protected lemma subsingleton.antitone_on (h : s.subsingleton) : + antitone_on f s := +λ a ha b hb _, (congr_arg _ (h hb ha)).le + +protected lemma subsingleton.strict_mono_on (h : s.subsingleton) : + strict_mono_on f s := +λ a ha b hb hlt, (hlt.ne (h ha hb)).elim + +protected lemma subsingleton.strict_anti_on (h : s.subsingleton) : + strict_anti_on f s := +λ a ha b hb hlt, (hlt.ne (h ha hb)).elim + +@[simp] lemma monotone_on_singleton : monotone_on f {a} := +subsingleton_singleton.monotone_on f + +@[simp] lemma antitone_on_singleton : antitone_on f {a} := +subsingleton_singleton.antitone_on f + +@[simp] lemma strict_mono_on_singleton : strict_mono_on f {a} := +subsingleton_singleton.strict_mono_on f + +@[simp] lemma strict_anti_on_singleton : strict_anti_on f {a} := +subsingleton_singleton.strict_anti_on f + +end preorder + /-! ### Lemmas about range of a function. -/ section range variables {f : ι → α} diff --git a/src/data/set/intervals/monotone.lean b/src/data/set/intervals/monotone.lean index fc9b40cb4cc54..bc11fba75cc52 100644 --- a/src/data/set/intervals/monotone.lean +++ b/src/data/set/intervals/monotone.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import data.set.intervals.disjoint +import order.succ_pred.basic import tactic.field_simp /-! @@ -233,3 +234,79 @@ calc (⋃ x, Ioo (f x) (g x)) = (⋃ x, Ioi (f x)) ∩ ⋃ x, Iio (g x) : ... = Ioi a ∩ Iio b : congr_arg2 (∩) ha.Union_Ioi_eq hb.Union_Iio_eq end Union + +section succ_order + +open order + +variables {α β : Type*} [partial_order α] + +lemma strict_mono_on.Iic_id_le [succ_order α] [is_succ_archimedean α] [order_bot α] + {n : α} {φ : α → α} (hφ : strict_mono_on φ (set.Iic n)) : + ∀ m ≤ n, m ≤ φ m := +begin + revert hφ, + refine succ.rec_bot (λ n, strict_mono_on φ (set.Iic n) → ∀ m ≤ n, m ≤ φ m) + (λ _ _ hm, hm.trans bot_le) _ _, + rintro k ih hφ m hm, + by_cases hk : is_max k, + { rw succ_eq_iff_is_max.2 hk at hm, + exact ih (hφ.mono $ Iic_subset_Iic.2 (le_succ _)) _ hm }, + obtain (rfl | h) := le_succ_iff_eq_or_le.1 hm, + { specialize ih (strict_mono_on.mono hφ (λ x hx, le_trans hx (le_succ _))) k le_rfl, + refine le_trans (succ_mono ih) (succ_le_of_lt (hφ (le_succ _) le_rfl _)), + rw lt_succ_iff_eq_or_lt_of_not_is_max hk, + exact or.inl rfl }, + { 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 α] + {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) + +variables [preorder β] {ψ : α → β} + +/-- A function `ψ` on a `succ_order` is strictly monotone before some `n` if for all `m` such that +`m < n`, we have `ψ m < ψ (succ m)`. -/ +lemma strict_mono_on_Iic_of_lt_succ [succ_order α] [is_succ_archimedean α] + {n : α} (hψ : ∀ m, m < n → ψ m < ψ (succ m)) : + strict_mono_on ψ (set.Iic n) := +begin + intros x hx y hy hxy, + obtain ⟨i, rfl⟩ := hxy.le.exists_succ_iterate, + induction i with k ih, + { simpa using hxy }, + cases k, + { exact hψ _ (lt_of_lt_of_le hxy hy) }, + rw set.mem_Iic at *, + simp only [function.iterate_succ', function.comp_apply] at ih hxy hy ⊢, + by_cases hmax : is_max (succ^[k] x), + { rw succ_eq_iff_is_max.2 hmax at hxy ⊢, + exact ih (le_trans (le_succ _) hy) hxy }, + by_cases hmax' : is_max (succ (succ^[k] x)), + { rw succ_eq_iff_is_max.2 hmax' at hxy ⊢, + exact ih (le_trans (le_succ _) hy) hxy }, + refine lt_trans (ih (le_trans (le_succ _) hy) + (lt_of_le_of_lt (le_succ_iterate k _) (lt_succ_iff_not_is_max.2 hmax))) _, + rw [← function.comp_apply succ, ← function.iterate_succ'], + refine hψ _ (lt_of_lt_of_le _ hy), + rwa [function.iterate_succ', function.comp_apply, lt_succ_iff_not_is_max], +end + +lemma strict_anti_on_Iic_of_succ_lt [succ_order α] [is_succ_archimedean α] + {n : α} (hψ : ∀ m, m < n → ψ (succ m) < ψ m) : + 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 α] + {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 α] + {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 + +end succ_order diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index d272b8fe7e6dc..c3a1904915278 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -210,6 +210,12 @@ end lemma succ_mono : monotone (succ : α → α) := λ a b, succ_le_succ +lemma le_succ_iterate (k : ℕ) (x : α) : x ≤ (succ^[k] x) := +begin + conv_lhs { rw (by simp only [function.iterate_id, id.def] : x = (id^[k] x)) }, + exact monotone.le_iterate_of_le succ_mono le_succ k x, +end + lemma Iio_succ_of_not_is_max (ha : ¬ is_max a) : Iio (succ a) = Iic a := set.ext $ λ x, lt_succ_iff_of_not_is_max ha @@ -427,6 +433,12 @@ lemma le_pred_iff_of_not_is_min (ha : ¬ is_min a) : b ≤ pred a ↔ b < a := lemma pred_mono : monotone (pred : α → α) := λ a b, pred_le_pred +lemma pred_iterate_le (k : ℕ) (x : α) : (pred^[k] x) ≤ x := +begin + conv_rhs { rw (by simp only [function.iterate_id, id.def] : x = (id^[k] x)) }, + exact monotone.iterate_le_of_le pred_mono pred_le k x, +end + lemma Ioi_pred_of_not_is_min (ha : ¬ is_min a) : Ioi (pred a) = Ici a := set.ext $ λ x, pred_lt_iff_of_not_is_min ha From 5d1f9f503c6ab2e809ce9a95b1b8a2242a5f2284 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 26 Jul 2022 20:33:27 +0000 Subject: [PATCH 076/173] feat(analysis/inner_product/pi_L2): a finite orthonormal family is a basis of its span (#15481) We actually prove this for finite sub-families of a generic orthonormal basis because this is easier to use in practice --- src/analysis/inner_product_space/basic.lean | 23 ++++++++++--- src/analysis/inner_product_space/pi_L2.lean | 38 ++++++++++++++++++++- src/linear_algebra/basis.lean | 3 ++ 3 files changed, 59 insertions(+), 5 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 1b6bff179d103..66177b2cc0ef1 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1207,14 +1207,18 @@ def linear_equiv.isometry_of_inner (f : E ≃ₗ[𝕜] E') (h : ∀ x y, ⟪f x, (f.isometry_of_inner h).to_linear_equiv = f := rfl /-- A linear isometry preserves the property of being orthonormal. -/ -lemma orthonormal.comp_linear_isometry {v : ι → E} (hv : orthonormal 𝕜 v) (f : E →ₗᵢ[𝕜] E') : - orthonormal 𝕜 (f ∘ v) := +lemma linear_isometry.orthonormal_comp_iff {v : ι → E} (f : E →ₗᵢ[𝕜] E') : + orthonormal 𝕜 (f ∘ v) ↔ orthonormal 𝕜 v := begin classical, - simp_rw [orthonormal_iff_ite, linear_isometry.inner_map_map, ←orthonormal_iff_ite], - exact hv + simp_rw [orthonormal_iff_ite, linear_isometry.inner_map_map] end +/-- A linear isometry preserves the property of being orthonormal. -/ +lemma orthonormal.comp_linear_isometry {v : ι → E} (hv : orthonormal 𝕜 v) (f : E →ₗᵢ[𝕜] E') : + orthonormal 𝕜 (f ∘ v) := +by rwa f.orthonormal_comp_iff + /-- A linear isometric equivalence preserves the property of being orthonormal. -/ lemma orthonormal.comp_linear_isometry_equiv {v : ι → E} (hv : orthonormal 𝕜 v) (f : E ≃ₗᵢ[𝕜] E') : orthonormal 𝕜 (f ∘ v) := @@ -1820,6 +1824,17 @@ instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_sp /-- The inner product on submodules is the same as on the ambient space. -/ @[simp] lemma submodule.coe_inner (W : submodule 𝕜 E) (x y : W) : ⟪x, y⟫ = ⟪(x:E), ↑y⟫ := rfl +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) := +s.subtypeₗᵢ.orthonormal_comp_iff.mp hv + +lemma orthonormal_span {ι : Type*} {v : ι → E} (hv : orthonormal 𝕜 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)) + /-! ### Families of mutually-orthogonal subspaces of an inner product space -/ section orthogonal_family diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 195a9a4ccf34f..70f7499ce8496 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -50,7 +50,7 @@ For consequences in infinite dimension (Hilbert bases, etc.), see the file -/ -open real set filter is_R_or_C +open real set filter is_R_or_C submodule open_locale big_operators uniformity topological_space nnreal ennreal complex_conjugate direct_sum noncomputable theory @@ -338,6 +338,16 @@ begin refl, end +/-- Mapping an orthonormal basis along a `linear_isometry_equiv`. -/ +protected def map {G : Type*} [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) + (L : E ≃ₗᵢ[𝕜] G) : + orthonormal_basis ι 𝕜 G := +{ repr := L.symm.trans b.repr } + +@[simp] protected lemma map_apply {G : Type*} [inner_product_space 𝕜 G] + (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) (i : ι) : + b.map L i = L (b i) := rfl + /-- A basis that is orthonormal is an orthonormal basis. -/ def _root_.basis.to_orthonormal_basis (v : basis ι 𝕜 E) (hv : orthonormal 𝕜 v) : orthonormal_basis ι 𝕜 E := @@ -385,6 +395,32 @@ protected lemma coe_mk (hon : orthonormal 𝕜 v) (hsp: submodule.span 𝕜 (set ⇑(orthonormal_basis.mk hon hsp) = v := 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 ι') : + orthonormal_basis s 𝕜 (span 𝕜 (s.image v' : set E)) := +let + e₀' : basis s 𝕜 _ := basis.span (h.linear_independent.comp (coe : s → ι') subtype.coe_injective), + e₀ : orthonormal_basis s 𝕜 _ := orthonormal_basis.mk + begin + convert orthonormal_span (h.comp (coe : s → ι') subtype.coe_injective), + ext, + simp [e₀', basis.span_apply], + end e₀'.span_eq, + φ : span 𝕜 (s.image v' : set E) ≃ₗᵢ[𝕜] span 𝕜 (range (v' ∘ (coe : s → ι'))) := + linear_isometry_equiv.of_eq _ _ + begin + rw [finset.coe_image, image_eq_range], + refl + end +in +e₀.map φ.symm + +@[simp] protected lemma span_apply {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, + linear_isometry_equiv.coe_of_eq_apply] + open submodule /-- A finite orthonormal family of vectors whose span has trivial orthogonal complement is an diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index dfad04351c884..030f264424770 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -970,6 +970,9 @@ begin rwa h_x_eq_y end +protected lemma span_apply (i : ι) : (basis.span hli i : M) = v i := +congr_arg (coe : span R (range v) → M) $ basis.mk_apply (linear_independent_span hli) _ i + end span lemma group_smul_span_eq_top From 01889e442293bbd35e17fc708e95b726762d2810 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 26 Jul 2022 20:33:29 +0000 Subject: [PATCH 077/173] chore(data/list/of_fn): injectivity (#15695) --- src/data/fin/tuple/basic.lean | 6 ++++++ src/data/list/of_fn.lean | 13 +++++++++++++ 2 files changed, 19 insertions(+) diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index d9b7b9aaf2eaf..cfc60e1bedab6 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -701,4 +701,10 @@ begin simpa using h, end +/-- `fin.sigma_eq_of_eq_comp_cast` as an `iff`. -/ +lemma sigma_eq_iff_eq_comp_cast {α : Type*} {a b : Σ ii, fin ii → α} : + a = b ↔ ∃ (h : a.fst = b.fst), a.snd = b.snd ∘ fin.cast h := +⟨λ h, h ▸ ⟨rfl, funext $ subtype.rec $ by exact λ i hi, rfl⟩, + λ ⟨h, h'⟩, sigma_eq_of_eq_comp_cast _ h'⟩ + end fin diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 313b8c1d98cf3..4c695e90a6889 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -188,4 +188,17 @@ lemma forall_iff_forall_tuple {P : list α → Prop} : (∀ l : list α, P l) ↔ ∀ n (f : fin n → α), P (list.of_fn f) := equiv_sigma_tuple.symm.surjective.forall.trans sigma.forall +/-- `fin.sigma_eq_iff_eq_comp_cast` may be useful to work with the RHS of this expression. -/ +lemma of_fn_inj' {m n : ℕ} {f : fin m → α} {g : fin n → α} : + of_fn f = of_fn g ↔ (⟨m, f⟩ : Σ n, fin n → α) = ⟨n, g⟩ := +iff.symm $ equiv_sigma_tuple.symm.injective.eq_iff.symm + +/-- Note we can only state this when the two functions are indexed by defeq `n`. -/ +lemma of_fn_injective {n : ℕ} : function.injective (of_fn : (fin n → α) → list α) := +λ f g h, eq_of_heq $ by injection of_fn_inj'.mp h + +/-- A special case of `list.of_fn_inj'` for when the two functions are indexed by defeq `n`. -/ +@[simp] lemma of_fn_inj {n : ℕ} {f g : fin n → α} : of_fn f = of_fn g ↔ f = g := +of_fn_injective.eq_iff + end list From 28fcd346e08afe9e4e2ba461298665fa54b63f3a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Jul 2022 20:33:31 +0000 Subject: [PATCH 078/173] =?UTF-8?q?chore(topology/sequences):=20drop=20loc?= =?UTF-8?q?al=20notation=20`x=20=E2=9F=B6=20a`=20(#15696)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit As requested by @PatrickMassot some time ago. --- src/topology/sequences.lean | 47 +++++++++++++++++-------------------- 1 file changed, 22 insertions(+), 25 deletions(-) diff --git a/src/topology/sequences.lean b/src/topology/sequences.lean index 1ec0d97328d2a..273d64bc0f54a 100644 --- a/src/topology/sequences.lean +++ b/src/topology/sequences.lean @@ -26,8 +26,6 @@ open_locale topological_space variables {X Y : Type*} -local notation x ` ⟶ ` a := tendsto x at_top (𝓝 a) - /-! ### Sequential closures, sequential continuity, and sequential spaces. -/ section topological_space variables [topological_space X] [topological_space Y] @@ -35,7 +33,7 @@ variables [topological_space X] [topological_space Y] /-- The sequential closure of a set `s : set X` in a topological space `X` is the set of all `a : X` which arise as limit of sequences in `s`. -/ def seq_closure (s : set X) : set X := -{a | ∃ x : ℕ → X, (∀ n : ℕ, x n ∈ s) ∧ (x ⟶ a)} +{a | ∃ x : ℕ → X, (∀ n : ℕ, x n ∈ s) ∧ tendsto x at_top (𝓝 a)} lemma subset_seq_closure (s : set X) : s ⊆ seq_closure s := λ a ha, ⟨const ℕ a, λ n, ha, tendsto_const_nhds⟩ @@ -46,11 +44,12 @@ def is_seq_closed (s : set X) : Prop := s = seq_closure s /-- A convenience lemma for showing that a set is sequentially closed. -/ lemma is_seq_closed_of_def {s : set X} - (h : ∀ (x : ℕ → X) (a : X), (∀ n : ℕ, x n ∈ s) → (x ⟶ a) → a ∈ s) : is_seq_closed s := + (h : ∀ (x : ℕ → X) (a : X), (∀ n : ℕ, x n ∈ s) → tendsto x at_top (𝓝 a) → a ∈ s) : + is_seq_closed s := show s = seq_closure s, from subset.antisymm (subset_seq_closure s) (show ∀ a, a ∈ seq_closure s → a ∈ s, from - (assume a ⟨x, _, _⟩, show a ∈ s, from h x a ‹∀ n : ℕ, ((x n) ∈ s)› ‹(x ⟶ a)›)) + (assume a ⟨x, hxs, hxa⟩, show a ∈ s, from h x a hxs hxa)) /-- The sequential closure of a set is contained in the closure of that set. The converse is not true. -/ @@ -66,9 +65,8 @@ calc seq_closure s ⊆ closure s : seq_closure_subset_closure s /-- The limit of a convergent sequence in a sequentially closed set is in that set.-/ lemma is_seq_closed.mem_of_tendsto {s : set X} (hs : is_seq_closed s) {x : ℕ → X} - (hmem : ∀ n, x n ∈ s) {a : X} (ha : (x ⟶ a)) : a ∈ s := -have a ∈ seq_closure s, from - show ∃ x : ℕ → X, (∀ n : ℕ, x n ∈ s) ∧ (x ⟶ a), from ⟨x, ‹∀ n, x n ∈ s›, ‹(x ⟶ a)›⟩, + (hmem : ∀ n, x n ∈ s) {a : X} (ha : tendsto x at_top (𝓝 a)) : a ∈ s := +have a ∈ seq_closure s, from ⟨x, hmem, ha⟩, eq.subst (eq.symm ‹is_seq_closed s›) ‹a ∈ seq_closure s› /-- A sequential space is a space in which 'sequences are enough to probe the topology'. This can be @@ -92,19 +90,17 @@ alias is_seq_closed_iff_is_closed ↔ is_seq_closed.is_closed _ /-- In a sequential space, a point belongs to the closure of a set iff it is a limit of a sequence taking values in this set. -/ lemma mem_closure_iff_seq_limit [sequential_space X] {s : set X} {a : X} : - a ∈ closure s ↔ ∃ x : ℕ → X, (∀ n : ℕ, x n ∈ s) ∧ (x ⟶ a) := + a ∈ closure s ↔ ∃ x : ℕ → X, (∀ n : ℕ, x n ∈ s) ∧ tendsto x at_top (𝓝 a) := by { rw ← sequential_space.seq_closure_eq_closure, exact iff.rfl } /-- A function between topological spaces is sequentially continuous if it commutes with limit of convergent sequences. -/ def seq_continuous (f : X → Y) : Prop := -∀ (x : ℕ → X), ∀ {a : X}, (x ⟶ a) → (f ∘ x ⟶ f a) +∀ (x : ℕ → X), ∀ {a : X}, tendsto x at_top (𝓝 a) → tendsto (f ∘ x) at_top (𝓝 (f a)) /- A continuous function is sequentially continuous. -/ protected lemma continuous.seq_continuous {f : X → Y} (hf : continuous f) : seq_continuous f := -assume x a (_ : x ⟶ a), -have tendsto f (𝓝 a) (𝓝 (f a)), from continuous.tendsto ‹continuous f› a, -show (f ∘ x) ⟶ (f a), from tendsto.comp this ‹(x ⟶ a)› +λ x a, (hf.tendsto _).comp /-- In a sequential space, continuity and sequential continuity coincide. -/ lemma continuous_iff_seq_continuous {f : X → Y} [sequential_space X] : @@ -116,10 +112,10 @@ iff.intro continuous_iff_is_closed.mpr (assume s _, is_seq_closed_iff_is_closed.mp $ h ‹is_closed s›), assume s (_ : is_closed s), is_seq_closed_of_def $ - assume (x : ℕ → X) a (_ : ∀ n, f (x n) ∈ s) (_ : x ⟶ a), - have (f ∘ x) ⟶ (f a), from ‹seq_continuous f› x ‹(x ⟶ a)›, + assume (x : ℕ → X) a (_ : ∀ n, f (x n) ∈ s) (_ : tendsto x at_top (𝓝 a)), + have tendsto (f ∘ x) at_top (𝓝 (f a)), from ‹seq_continuous f› x ‹_›, show f a ∈ s, - from ‹is_closed s›.is_seq_closed.mem_of_tendsto ‹∀ n, f (x n) ∈ s› ‹(f∘x ⟶ f a)›) + from ‹is_closed s›.is_seq_closed.mem_of_tendsto ‹∀ n, f (x n) ∈ s› ‹_›) alias continuous_iff_seq_continuous ↔ _ seq_continuous.continuous @@ -167,7 +163,7 @@ variables [topological_space X] /-- A set `s` is sequentially compact if every sequence taking values in `s` has a converging subsequence. -/ def is_seq_compact (s : set X) := -∀ ⦃x : ℕ → X⦄, (∀ n, x n ∈ s) → ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) +∀ ⦃x : ℕ → X⦄, (∀ n, x n ∈ s) → ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) /-- A space `X` is sequentially compact if every sequence in `X` has a converging subsequence. -/ @@ -176,12 +172,12 @@ class seq_compact_space (X : Type*) [topological_space X] : Prop := lemma is_seq_compact.subseq_of_frequently_in {s : set X} (hs : is_seq_compact s) {x : ℕ → X} (hx : ∃ᶠ n in at_top, x n ∈ s) : - ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := let ⟨ψ, hψ, huψ⟩ := extraction_of_frequently_at_top hx, ⟨a, a_in, φ, hφ, h⟩ := hs huψ in ⟨a, a_in, ψ ∘ φ, hψ.comp hφ, h⟩ lemma seq_compact_space.tendsto_subseq [seq_compact_space X] (x : ℕ → X) : - ∃ a (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ a (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := let ⟨a, _, φ, mono, h⟩ := seq_compact_space.seq_compact_univ (λ n, mem_univ (x n)) in ⟨a, φ, mono, h⟩ @@ -196,11 +192,11 @@ let ⟨a, a_in, ha⟩ := @hs (map x at_top) _ lemma is_compact.tendsto_subseq' {s : set X} {x : ℕ → X} (hs : is_compact s) (hx : ∃ᶠ n in at_top, x n ∈ s) : - ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := hs.is_seq_compact.subseq_of_frequently_in hx lemma is_compact.tendsto_subseq {s : set X} {x : ℕ → X} (hs : is_compact s) (hx : ∀ n, x n ∈ s) : - ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ (a ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := hs.is_seq_compact hx @[priority 100] -- see Note [lower instance priority] @@ -208,7 +204,7 @@ instance first_countable_topology.seq_compact_of_compact [compact_space X] : seq ⟨compact_univ.is_seq_compact⟩ lemma compact_space.tendsto_subseq [compact_space X] (x : ℕ → X) : - ∃ a (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ a (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := seq_compact_space.tendsto_subseq x end first_countable_topology @@ -237,7 +233,8 @@ begin { push_neg at H, choose x hx using H, exact ⟨x, forall_and_distrib.mp hx⟩ }, clear H, - obtain ⟨x₀, x₀_in, φ, φ_mono, hlim⟩ : ∃ (x₀ ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ (x ∘ φ ⟶ x₀), + obtain ⟨x₀, x₀_in, φ, φ_mono, hlim⟩ : + ∃ (x₀ ∈ s) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 x₀), from hs x_in, clear hs, obtain ⟨i₀, x₀_in⟩ : ∃ i₀, x₀ ∈ c i₀, { rcases hc₂ x₀_in with ⟨_, ⟨i₀, rfl⟩, x₀_in_c⟩, @@ -350,7 +347,7 @@ every bounded sequence has a converging subsequence. This version assumes only that the sequence is frequently in some bounded set. -/ lemma tendsto_subseq_of_frequently_bounded (hs : bounded s) {x : ℕ → X} (hx : ∃ᶠ n in at_top, x n ∈ s) : - ∃ a ∈ closure s, ∃ φ : ℕ → ℕ, strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ a ∈ closure s, ∃ φ : ℕ → ℕ, strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := have hcs : is_seq_compact (closure s), from hs.is_compact_closure.is_seq_compact, have hu' : ∃ᶠ n in at_top, x n ∈ closure s, from hx.mono (λ n hn, subset_closure hn), hcs.subseq_of_frequently_in hu' @@ -358,7 +355,7 @@ hcs.subseq_of_frequently_in hu' /-- A version of Bolzano-Weistrass: in a proper metric space (eg. $ℝ^n$), every bounded sequence has a converging subsequence. -/ lemma tendsto_subseq_of_bounded (hs : bounded s) {x : ℕ → X} (hx : ∀ n, x n ∈ s) : - ∃ a ∈ closure s, ∃ φ : ℕ → ℕ, strict_mono φ ∧ (x ∘ φ ⟶ a) := + ∃ a ∈ closure s, ∃ φ : ℕ → ℕ, strict_mono φ ∧ tendsto (x ∘ φ) at_top (𝓝 a) := tendsto_subseq_of_frequently_bounded hs $ frequently_of_forall hx end metric_seq_compact From 4af053164ce135e0504f06acb156b9c794a2dc39 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 26 Jul 2022 23:26:00 +0000 Subject: [PATCH 079/173] feat(analysis/inner_product_space): in finite dimension, hilbert basis = orthonormal basis (#15540) This allows transferring general facts about hilbert bases to orthonormal bases --- .../inner_product_space/l2_space.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index a88341d32c461..ca1dd6790cb0d 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -5,6 +5,7 @@ Authors: Heather Macbeth -/ import analysis.inner_product_space.projection import analysis.normed_space.lp_space +import analysis.inner_product_space.pi_L2 /-! # Hilbert sum of a family of inner product spaces @@ -403,6 +404,22 @@ protected lemma tsum_inner_mul_inner (b : hilbert_basis ι 𝕜 E) (x y : E) : ∑' i, ⟪x, b i⟫ * ⟪b i, y⟫ = ⟪x, y⟫ := (b.has_sum_inner_mul_inner x y).tsum_eq +-- Note : this should be `b.repr` composed with an identification of `lp (λ i : ι, 𝕜) p` with +-- `pi_Lp p (λ i : ι, 𝕜)` (in this case with `p = 2`), but we don't have this yet (July 2022). +/-- A finite Hilbert basis is an orthonormal basis. -/ +protected def to_orthonormal_basis [fintype ι] (b : hilbert_basis ι 𝕜 E) : + orthonormal_basis ι 𝕜 E := +orthonormal_basis.mk b.orthonormal +begin + have := (span 𝕜 (finset.univ.image b : set E)).closed_of_finite_dimensional, + simpa only [finset.coe_image, finset.coe_univ, set.image_univ, hilbert_basis.dense_span] using + this.submodule_topological_closure_eq.symm +end + +@[simp] lemma coe_to_orthonormal_basis [fintype ι] (b : hilbert_basis ι 𝕜 E) : + (b.to_orthonormal_basis : ι → E) = b := +orthonormal_basis.coe_mk _ _ + variables {v : ι → E} (hv : orthonormal 𝕜 v) include hv cplt @@ -437,6 +454,19 @@ hilbert_basis.coe_mk hv _ omit hv +-- Note : this should be `b.repr` composed with an identification of `lp (λ i : ι, 𝕜) p` with +-- `pi_Lp p (λ i : ι, 𝕜)` (in this case with `p = 2`), but we don't have this yet (July 2022). +/-- An orthonormal basis is an Hilbert basis. -/ +protected def _root_.orthonormal_basis.to_hilbert_basis [fintype ι] (b : orthonormal_basis ι 𝕜 E) : + hilbert_basis ι 𝕜 E := +hilbert_basis.mk b.orthonormal $ +by simpa only [← orthonormal_basis.coe_to_basis, b.to_basis.span_eq, eq_top_iff] + using @subset_closure E _ _ + +@[simp] lemma _root_.orthonormal_basis.coe_to_hilbert_basis [fintype ι] + (b : orthonormal_basis ι 𝕜 E) : (b.to_hilbert_basis : ι → E) = b := +hilbert_basis.coe_mk _ _ + /-- A Hilbert space admits a Hilbert basis extending a given orthonormal subset. -/ lemma _root_.orthonormal.exists_hilbert_basis_extension {s : set E} (hs : orthonormal 𝕜 (coe : s → E)) : From 4787b7f5a48578e5c08dea6867f9511f1dd8465c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 26 Jul 2022 23:26:01 +0000 Subject: [PATCH 080/173] feat(set_theory/zfc/basic): `inj` lemmas (#15545) We rename two existing lemmas from `*_inj` to `*_injective` to match mathlib convention, and add the corresponding `inj` lemmas. --- src/set_theory/zfc/basic.lean | 21 +++++++++++++-------- 1 file changed, 13 insertions(+), 8 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 779c333d8bdec..3fb1c9e86102a 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -594,8 +594,10 @@ mem_sUnion.2 ⟨z, hz, hy⟩ @[simp] theorem sUnion_singleton {x : Set.{u}} : ⋃₀ ({x} : Set) = x := ext $ λ y, by simp_rw [mem_sUnion, exists_prop, mem_singleton, exists_eq_left] -theorem singleton_inj {x y : Set.{u}} (H : ({x} : Set) = {y}) : x = y := -let this := congr_arg sUnion H in by rwa [sUnion_singleton, sUnion_singleton] at this +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 /-- The binary union operation -/ protected def union (x y : Set.{u}) : Set.{u} := ⋃₀ {x, y} @@ -684,13 +686,13 @@ begin { rintro (rfl|rfl); [left, right]; assumption } end -theorem pair_inj {x y x' y' : Set.{u}} (H : pair x y = pair x' y') : x = x' ∧ y = y' := -begin +theorem pair_injective : function.injective2 pair := +λ x x' y y' H, begin have ae := ext_iff.2 H, simp only [pair, mem_pair] at ae, obtain rfl : x = x', { cases (ae {x}).1 (by simp) with h h, - { exact singleton_inj h }, + { exact singleton_injective h }, { have m : x' ∈ ({x} : Set), { simp [h] }, rw mem_singleton.mp m } }, @@ -708,6 +710,9 @@ begin { simp [yy'] } } end +@[simp] theorem pair_inj {x y x' y' : Set} : pair x y = pair x' y' ↔ x = x' ∧ y = y' := +pair_injective.eq_iff + /-- The cartesian product, `{(a, b) | a ∈ x, b ∈ y}` -/ def prod : Set.{u} → Set.{u} → Set.{u} := pair_sep (λ a b, true) @@ -716,7 +721,7 @@ by simp [prod] @[simp] theorem pair_mem_prod {x y a b : Set.{u}} : pair a b ∈ prod x y ↔ a ∈ x ∧ b ∈ y := ⟨λ h, let ⟨a', a'x, b', b'y, e⟩ := mem_prod.1 h in - match a', b', pair_inj e, a'x, b'y with ._, ._, ⟨rfl, rfl⟩, ax, bY := ⟨ax, bY⟩ end, + match a', b', pair_injective e, a'x, b'y with ._, ._, ⟨rfl, rfl⟩, ax, bY := ⟨ax, bY⟩ end, λ ⟨ax, bY⟩, mem_prod.2 ⟨a, ax, b, bY, rfl⟩⟩ /-- `is_func x y f` is the assertion that `f` is a subset of `x × y` which relates to each element @@ -746,7 +751,7 @@ mem_image theorem map_unique {f : Set.{u} → Set.{u}} [H : definable 1 f] {x z : Set.{u}} (zx : z ∈ x) : ∃! w, pair z w ∈ map f x := -⟨f z, image.mk _ _ zx, λ y yx, let ⟨w, wx, we⟩ := mem_image.1 yx, ⟨wz, fy⟩ := pair_inj we in +⟨f z, image.mk _ _ zx, λ y yx, let ⟨w, wx, we⟩ := mem_image.1 yx, ⟨wz, fy⟩ := pair_injective we in by rw[←fy, wz]⟩ @[simp] theorem map_is_func {f : Set → Set} [H : definable 1 f] {x y : Set} : @@ -888,7 +893,7 @@ namespace Set {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 - ⟨λ ⟨w, wz, pr⟩, let ⟨wy, fw⟩ := Set.pair_inj pr in by rw[←fw, wy], + ⟨λ ⟨w, wz, pr⟩, let ⟨wy, fw⟩ := Set.pair_injective pr in by rw[←fw, wy], λ e, by { subst e, exact ⟨_, h, rfl⟩ }⟩ }) variables (x : Set.{u}) (h : ∅ ∉ x) From e942d1cbae8da63ae3ff9e7858ae233809ab5278 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 26 Jul 2022 23:26:02 +0000 Subject: [PATCH 081/173] style(analysis/inner_product_space/l2_space): small style improvement (#15698) Spotted by @hrmacbeth --- src/analysis/inner_product_space/l2_space.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index ca1dd6790cb0d..11eb17b1f7b03 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -391,9 +391,9 @@ 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).map _ (innerSL x).continuous, + convert (b.has_sum_repr y).mapL (innerSL x), ext i, - rw [function.comp_apply, innerSL_apply, b.repr_apply_apply, inner_smul_right, mul_comm] + rw [innerSL_apply, b.repr_apply_apply, inner_smul_right, mul_comm] end protected lemma summable_inner_mul_inner (b : hilbert_basis ι 𝕜 E) (x y : E) : From 0b0e8236a4c9435af1f6dfe75a0e335230a9a7fc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Jul 2022 00:42:28 +0000 Subject: [PATCH 082/173] refactor(algebra/big_operators/basic): Use `to_additive` (#15702) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove `sum_range_sub` and `sum_range_sub'` using `to_additive`. Change the statement of `prod_range_div` to use `f (i + 1) / f i` rather than `f (i + 1) * (f i)⁻¹`. Rename `sum_range_sub_of_monotone` to `sum_range_tsub`. --- src/algebra/big_operators/basic.lean | 48 +++++++++------------------- 1 file changed, 15 insertions(+), 33 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index a3793a032be2c..9636c53e79a5f 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1070,12 +1070,15 @@ lemma prod_induction_nonempty {M : Type*} [comm_monoid M] (f : α → M) (p : M multiset.prod_induction_nonempty p p_mul (by simp [nonempty_iff_ne_empty.mp hs_nonempty]) (multiset.forall_mem_map_iff.mpr p_s) -/-- -For any product along `{0, ..., n-1}` of a commutative-monoid-valued function, we can verify that -it's equal to a different function just by checking ratios of adjacent terms. +/-- For any product along `{0, ..., n - 1}` of a commutative-monoid-valued function, we can verify +that it's equal to a different function just by checking ratios of adjacent terms. + This is a multiplicative discrete analogue of the fundamental theorem of calculus. -/ -lemma prod_range_induction {M : Type*} [comm_monoid M] - (f s : ℕ → M) (h0 : s 0 = 1) (h : ∀ n, s (n + 1) = s n * f n) (n : ℕ) : +@[to_additive "For any sum along `{0, ..., n - 1}` of a commutative-monoid-valued function, we can +verify that it's equal to a different function just by checking differences of adjacent terms. + +This is a discrete analogue of the fundamental theorem of calculus."] +lemma prod_range_induction (f s : ℕ → β) (h0 : s 0 = 1) (h : ∀ n, s (n + 1) = s n * f n) (n : ℕ) : ∏ k in finset.range n, f k = s n := begin induction n with k hk, @@ -1083,46 +1086,25 @@ begin { simp only [hk, finset.prod_range_succ, h, mul_comm] } end -/-- -For any sum along `{0, ..., n-1}` of a commutative-monoid-valued function, -we can verify that it's equal to a different function -just by checking differences of adjacent terms. -This is a discrete analogue -of the fundamental theorem of calculus. --/ -lemma sum_range_induction {M : Type*} [add_comm_monoid M] - (f s : ℕ → M) (h0 : s 0 = 0) (h : ∀ n, s (n + 1) = s n + f n) (n : ℕ) : - ∑ k in finset.range n, f k = s n := -@prod_range_induction (multiplicative M) _ f s h0 h n - -/-- A telescoping sum along `{0, ..., n - 1}` of an additive commutative group valued function -reduces to the difference of the last and first terms.-/ -lemma sum_range_sub {G : Type*} [add_comm_group G] (f : ℕ → G) (n : ℕ) : - ∑ i in range n, (f (i+1) - f i) = f n - f 0 := -by { apply sum_range_induction; simp } - -lemma sum_range_sub' {G : Type*} [add_comm_group G] (f : ℕ → G) (n : ℕ) : - ∑ i in range n, (f i - f (i+1)) = f 0 - f n := -by { apply sum_range_induction; simp } - /-- A telescoping product along `{0, ..., n - 1}` of a commutative group valued function reduces to the ratio of the last and first factors. -/ -@[to_additive] +@[to_additive "A telescoping sum along `{0, ..., n - 1}` of an additive commutative group valued +function reduces to the difference of the last and first terms."] lemma prod_range_div {M : Type*} [comm_group M] (f : ℕ → M) (n : ℕ) : - ∏ i in range n, (f (i+1) * (f i)⁻¹) = f n * (f 0)⁻¹ := -by simpa only [← div_eq_mul_inv] using @sum_range_sub (additive M) _ f n + ∏ i in range n, (f (i + 1) / f i) = f n / f 0 := +by apply prod_range_induction; simp @[to_additive] lemma prod_range_div' {M : Type*} [comm_group M] (f : ℕ → M) (n : ℕ) : - ∏ i in range n, (f i * (f (i+1))⁻¹) = f 0 * (f n)⁻¹ := -by simpa only [← div_eq_mul_inv] using @sum_range_sub' (additive M) _ f n + ∏ i in range n, (f i / f (i + 1)) = f 0 / f n := +by apply prod_range_induction; simp /-- A telescoping sum along `{0, ..., n-1}` of an `ℕ`-valued function reduces to the difference of the last and first terms when the function we are summing is monotone. -/ -lemma sum_range_sub_of_monotone [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] +lemma sum_range_tsub [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] [contravariant_class α α (+) (≤)] {f : ℕ → α} (h : monotone f) (n : ℕ) : ∑ i in range n, (f (i+1) - f i) = f n - f 0 := begin From 57e6cff208df9b4f47cbf675bef8c816b40543c9 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Wed, 27 Jul 2022 04:27:17 +0000 Subject: [PATCH 083/173] chore(scripts): update nolints.txt (#15706) 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 80821b398ffc3..2331c52d28ab1 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -1584,9 +1584,6 @@ apply_nolint nonarchimedean_group.prod_subset to_additive_doc apply_nolint continuous.exists_forall_ge_of_has_compact_mul_support to_additive_doc apply_nolint continuous.exists_forall_le_of_has_compact_mul_support to_additive_doc --- topology/bases.lean -apply_nolint topological_space.second_countable_topology_fintype fintype_finite - -- topology/basic.lean apply_nolint closure_Union_of_fintype fintype_finite apply_nolint interior_Inter_of_fintype fintype_finite From 1176fe6b3b49b2827aad3be6389b7cd0a50818a9 Mon Sep 17 00:00:00 2001 From: kkytola Date: Wed, 27 Jul 2022 05:09:53 +0000 Subject: [PATCH 084/173] refactor(topology/algebra/order/basic): Add antitone versions of sup and inf lemmas for continuous monotone functions and move them to monotone/antitone namespaces. (#15218) This PR adds antitone versions of lemmas about supremum and infimum under monotone functions and moves those lemmas to monotone/antitone namespaces. Simultaneously, the type class assumption on the codomain is weakened from `order_topology` to `order_closed_topology`. Moreover, lemmas about limsup and liminf under monotone/antitone functions are added: `antitone.map_Limsup_of_continuous_at` etc. Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- .../rotation_number/translation_number.lean | 2 +- src/order/liminf_limsup.lean | 31 ++++ src/topology/algebra/order/basic.lean | 161 ++++++++++++++---- src/topology/algebra/order/liminf_limsup.lean | 106 ++++++++++++ src/topology/instances/ennreal.lean | 8 +- src/topology/metric_space/gluing.lean | 8 +- .../gromov_hausdorff_realized.lean | 8 +- 7 files changed, 280 insertions(+), 44 deletions(-) diff --git a/src/dynamics/circle/rotation_number/translation_number.lean b/src/dynamics/circle/rotation_number/translation_number.lean index 8b44a5d92d557..80a352a9ef4d5 100644 --- a/src/dynamics/circle/rotation_number/translation_number.lean +++ b/src/dynamics/circle/rotation_number/translation_number.lean @@ -842,7 +842,7 @@ begin { refine csupr_mono (this y) (λ g, _), exact mono _ (mono _ hxy) }, { simp only [map_add_one], - exact (map_csupr_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) + exact (monotone.map_csupr_of_continuous_at (continuous_at_id.add continuous_at_const) (monotone_id.add_const (1 : ℝ)) (this x)).symm }, { exact this x } end diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index f135e1fe1c8ed..73977d65726ed 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -465,6 +465,25 @@ theorem has_basis.liminf_eq_supr_infi {p : ι → Prop} {s : ι → set β} {f : (h : f.has_basis p s) : f.liminf u = ⨆ i (hi : p i), ⨅ a ∈ s i, u a := @has_basis.limsup_eq_infi_supr αᵒᵈ _ _ _ _ _ _ _ h +lemma limsup_eq_Inf_Sup {ι R : Type*} (F : filter ι) [complete_lattice R] (a : ι → R) : + F.limsup a = Inf ((λ I, Sup (a '' I)) '' F.sets) := +begin + refine le_antisymm _ _, + { rw limsup_eq, + refine Inf_le_Inf (λ x hx, _), + rcases (mem_image _ F.sets x).mp hx with ⟨I, ⟨I_mem_F, hI⟩⟩, + filter_upwards [I_mem_F] with i hi, + exact hI ▸ le_Sup (mem_image_of_mem _ hi), }, + { refine le_Inf_iff.mpr (λ b hb, Inf_le_of_le (mem_image_of_mem _ $ filter.mem_sets.mpr hb) + $ Sup_le _), + rintros _ ⟨_, h, rfl⟩, + exact h, }, +end + +lemma liminf_eq_Sup_Inf {ι R : Type*} (F : filter ι) [complete_lattice R] (a : ι → R) : + F.liminf a = Sup ((λ I, Inf (a '' I)) '' F.sets) := +@filter.limsup_eq_Inf_Sup ι (order_dual R) _ _ a + @[simp] lemma liminf_nat_add (f : ℕ → α) (k : ℕ) : at_top.liminf (λ i, f (i + k)) = at_top.liminf f := by { simp_rw liminf_eq_supr_infi_of_nat, exact supr_infi_ge_nat_add f k } @@ -495,6 +514,18 @@ end complete_lattice section conditionally_complete_linear_order +lemma frequently_lt_of_lt_Limsup {f : filter α} [conditionally_complete_linear_order α] {a : α} + (hf : f.is_cobounded (≤) . is_bounded_default) (h : a < f.Limsup) : ∃ᶠ n in f, a < n := +begin + contrapose! h, + simp only [not_frequently, not_lt] at h, + exact Limsup_le_of_le hf h, +end + +lemma frequently_lt_of_Liminf_lt {f : filter α} [conditionally_complete_linear_order α] {a : α} + (hf : f.is_cobounded (≥) . is_bounded_default) (h : f.Liminf < a) : ∃ᶠ n in f, n < a := +@frequently_lt_of_lt_Limsup (order_dual α) f _ a hf h + lemma eventually_lt_of_lt_liminf {f : filter α} [conditionally_complete_linear_order β] {u : α → β} {b : β} (h : b < liminf f u) (hu : f.is_bounded_under (≥) u . is_bounded_default) : ∀ᶠ a in f, b < u a := diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 9d43af5ec7b81..6bd6d1c141859 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -2732,7 +2732,7 @@ end densely_ordered section complete_linear_order variables [complete_linear_order α] [topological_space α] [order_topology α] - [complete_linear_order β] [topological_space β] [order_topology β] [nonempty γ] + [complete_linear_order β] [topological_space β] [order_closed_topology β] [nonempty γ] lemma Sup_mem_closure {α : Type u} [topological_space α] [complete_linear_order α] [order_topology α] {s : set α} (hs : s.nonempty) : @@ -2756,72 +2756,137 @@ lemma is_closed.Inf_mem {α : Type u} [topological_space α] [complete_linear_or /-- A monotone function continuous at the supremum of a nonempty set sends this supremum to the supremum of the image of this set. -/ -lemma map_Sup_of_continuous_at_of_monotone' {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) +lemma monotone.map_Sup_of_continuous_at' {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) (Mf : monotone f) (hs : s.nonempty) : f (Sup s) = Sup (f '' s) := --This is a particular case of the more general is_lub.is_lub_of_tendsto ((is_lub_Sup _).is_lub_of_tendsto (λ x hx y hy xy, Mf xy) hs $ Cf.mono_left inf_le_left).Sup_eq.symm -/-- A monotone function `s` sending `bot` to `bot` and continuous at the supremum of a set sends +/-- A monotone function `f` sending `bot` to `bot` and continuous at the supremum of a set sends this supremum to the supremum of the image of this set. -/ -lemma map_Sup_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) +lemma monotone.map_Sup_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) (Mf : monotone f) (fbot : f ⊥ = ⊥) : f (Sup s) = Sup (f '' s) := begin cases s.eq_empty_or_nonempty with h h, { simp [h, fbot] }, - { exact map_Sup_of_continuous_at_of_monotone' Cf Mf h } + { exact Mf.map_Sup_of_continuous_at' Cf h } end /-- A monotone function continuous at the indexed supremum over a nonempty `Sort` sends this indexed supremum to the indexed supremum of the composition. -/ -lemma map_supr_of_continuous_at_of_monotone' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} +lemma monotone.map_supr_of_continuous_at' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} (Cf : continuous_at f (supr g)) (Mf : monotone f) : f (⨆ i, g i) = ⨆ i, f (g i) := -by rw [supr, map_Sup_of_continuous_at_of_monotone' Cf Mf (range_nonempty g), ← range_comp, supr] +by rw [supr, Mf.map_Sup_of_continuous_at' Cf (range_nonempty g), ← range_comp, supr] /-- If a monotone function sending `bot` to `bot` is continuous at the indexed supremum over a `Sort`, then it sends this indexed supremum to the indexed supremum of the composition. -/ -lemma map_supr_of_continuous_at_of_monotone {ι : Sort*} {f : α → β} {g : ι → α} +lemma monotone.map_supr_of_continuous_at {ι : Sort*} {f : α → β} {g : ι → α} (Cf : continuous_at f (supr g)) (Mf : monotone f) (fbot : f ⊥ = ⊥) : f (⨆ i, g i) = ⨆ i, f (g i) := -by rw [supr, map_Sup_of_continuous_at_of_monotone Cf Mf fbot, ← range_comp, supr] +by rw [supr, Mf.map_Sup_of_continuous_at Cf fbot, ← range_comp, supr] /-- A monotone function continuous at the infimum of a nonempty set sends this infimum to the infimum of the image of this set. -/ -lemma map_Inf_of_continuous_at_of_monotone' {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) +lemma monotone.map_Inf_of_continuous_at' {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (hs : s.nonempty) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual hs +@monotone.map_Sup_of_continuous_at' αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual hs -/-- A monotone function `s` sending `top` to `top` and continuous at the infimum of a set sends +/-- A monotone function `f` sending `top` to `top` and continuous at the infimum of a set sends this infimum to the infimum of the image of this set. -/ -lemma map_Inf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) +lemma monotone.map_Inf_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ftop +@monotone.map_Sup_of_continuous_at αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ftop /-- A monotone function continuous at the indexed infimum over a nonempty `Sort` sends this indexed infimum to the indexed infimum of the composition. -/ -lemma map_infi_of_continuous_at_of_monotone' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} +lemma monotone.map_infi_of_continuous_at' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_supr_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι _ f g Cf Mf.dual +@monotone.map_supr_of_continuous_at' αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι _ f g Cf Mf.dual /-- If a monotone function sending `top` to `top` is continuous at the indexed infimum over a `Sort`, then it sends this indexed infimum to the indexed infimum of the composition. -/ -lemma map_infi_of_continuous_at_of_monotone {ι : Sort*} {f : α → β} {g : ι → α} +lemma monotone.map_infi_of_continuous_at {ι : Sort*} {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (infi g) = infi (f ∘ g) := -@map_supr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι f g Cf Mf.dual ftop +@monotone.map_supr_of_continuous_at αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι f g Cf Mf.dual ftop + +/-- An antitone function continuous at the supremum of a nonempty set sends this supremum to +the infimum of the image of this set. -/ +lemma antitone.map_Sup_of_continuous_at' {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) + (Af : antitone f) (hs : s.nonempty) : + f (Sup s) = Inf (f '' s) := +monotone.map_Sup_of_continuous_at' + (show continuous_at (order_dual.to_dual ∘ f) (Sup s), from Cf) Af hs + +/-- An antitone function `f` sending `bot` to `top` and continuous at the supremum of a set sends +this supremum to the infimum of the image of this set. -/ +lemma antitone.map_Sup_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) + (Af : antitone f) (fbot : f ⊥ = ⊤) : + f (Sup s) = Inf (f '' s) := +monotone.map_Sup_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (Sup s), from Cf) Af fbot + +/-- An antitone function continuous at the indexed supremum over a nonempty `Sort` sends this +indexed supremum to the indexed infimum of the composition. -/ +lemma antitone.map_supr_of_continuous_at' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} + (Cf : continuous_at f (supr g)) (Af : antitone f) : + f (⨆ i, g i) = ⨅ i, f (g i) := +monotone.map_supr_of_continuous_at' + (show continuous_at (order_dual.to_dual ∘ f) (supr g), from Cf) Af + +/-- An antitone function sending `bot` to `top` is continuous at the indexed supremum over +a `Sort`, then it sends this indexed supremum to the indexed supremum of the composition. -/ +lemma antitone.map_supr_of_continuous_at {ι : Sort*} {f : α → β} {g : ι → α} + (Cf : continuous_at f (supr g)) (Af : antitone f) (fbot : f ⊥ = ⊤) : + f (⨆ i, g i) = ⨅ i, f (g i) := +monotone.map_supr_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (supr g), from Cf) Af fbot + +/-- An antitone function continuous at the infimum of a nonempty set sends this infimum to +the supremum of the image of this set. -/ +lemma antitone.map_Inf_of_continuous_at' {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) + (Af : antitone f) (hs : s.nonempty) : + f (Inf s) = Sup (f '' s) := +monotone.map_Inf_of_continuous_at' + (show continuous_at (order_dual.to_dual ∘ f) (Inf s), from Cf) Af hs + +/-- An antitone function `f` sending `top` to `bot` and continuous at the infimum of a set sends +this infimum to the supremum of the image of this set. -/ +lemma antitone.map_Inf_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) + (Af : antitone f) (ftop : f ⊤ = ⊥) : + f (Inf s) = Sup (f '' s) := +monotone.map_Inf_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (Inf s), from Cf) Af ftop + +/-- An antitone function continuous at the indexed infimum over a nonempty `Sort` sends this indexed +infimum to the indexed supremum of the composition. -/ +lemma antitone.map_infi_of_continuous_at' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} + (Cf : continuous_at f (infi g)) (Af : antitone f) : + f (⨅ i, g i) = ⨆ i, f (g i) := +monotone.map_infi_of_continuous_at' + (show continuous_at (order_dual.to_dual ∘ f) (infi g), from Cf) Af + +/-- If an antitone function sending `top` to `bot` is continuous at the indexed infimum over +a `Sort`, then it sends this indexed infimum to the indexed supremum of the composition. -/ +lemma antitone.map_infi_of_continuous_at {ι : Sort*} {f : α → β} {g : ι → α} + (Cf : continuous_at f (infi g)) (Af : antitone f) (ftop : f ⊤ = ⊥) : + f (infi g) = supr (f ∘ g) := +monotone.map_infi_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (infi g), from Cf) Af ftop end complete_linear_order section conditionally_complete_linear_order variables [conditionally_complete_linear_order α] [topological_space α] [order_topology α] - [conditionally_complete_linear_order β] [topological_space β] [order_topology β] [nonempty γ] + [conditionally_complete_linear_order β] [topological_space β] [order_closed_topology β] + [nonempty γ] lemma cSup_mem_closure {s : set α} (hs : s.nonempty) (B : bdd_above s) : Sup s ∈ closure s := (is_lub_cSup hs B).mem_closure hs @@ -2839,7 +2904,7 @@ lemma is_closed.cInf_mem {s : set α} (hc : is_closed s) (hs : s.nonempty) (B : /-- If a monotone function is continuous at the supremum of a nonempty bounded above set `s`, then it sends this supremum to the supremum of the image of `s`. -/ -lemma map_cSup_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) +lemma monotone.map_cSup_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) (Mf : monotone f) (ne : s.nonempty) (H : bdd_above s) : f (Sup s) = Sup (f '' s) := begin @@ -2850,28 +2915,61 @@ end /-- If a monotone function is continuous at the indexed supremum of a bounded function on a nonempty `Sort`, then it sends this supremum to the supremum of the composition. -/ -lemma map_csupr_of_continuous_at_of_monotone {f : α → β} {g : γ → α} +lemma monotone.map_csupr_of_continuous_at {f : α → β} {g : γ → α} (Cf : continuous_at f (⨆ i, g i)) (Mf : monotone f) (H : bdd_above (range g)) : f (⨆ i, g i) = ⨆ i, f (g i) := -by rw [supr, map_cSup_of_continuous_at_of_monotone Cf Mf (range_nonempty _) H, ← range_comp, supr] +by rw [supr, Mf.map_cSup_of_continuous_at Cf (range_nonempty _) H, ← range_comp, supr] /-- If a monotone function is continuous at the infimum of a nonempty bounded below set `s`, then it sends this infimum to the infimum of the image of `s`. -/ -lemma map_cInf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) +lemma monotone.map_cInf_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ne : s.nonempty) (H : bdd_below s) : f (Inf s) = Inf (f '' s) := -@map_cSup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ne H +@monotone.map_cSup_of_continuous_at αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ne H /-- A continuous monotone function sends indexed infimum to indexed infimum in conditionally complete linear order, under a boundedness assumption. -/ -lemma map_cinfi_of_continuous_at_of_monotone {f : α → β} {g : γ → α} +lemma monotone.map_cinfi_of_continuous_at {f : α → β} {g : γ → α} (Cf : continuous_at f (⨅ i, g i)) (Mf : monotone f) (H : bdd_below (range g)) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_csupr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ _ _ _ Cf Mf.dual H +@monotone.map_csupr_of_continuous_at αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ _ _ _ Cf Mf.dual H + +/-- If an antitone function is continuous at the supremum of a nonempty bounded above set `s`, +then it sends this supremum to the infimum of the image of `s`. -/ +lemma antitone.map_cSup_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Sup s)) + (Af : antitone f) (ne : s.nonempty) (H : bdd_above s) : + f (Sup s) = Inf (f '' s) := +monotone.map_cSup_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (Sup s), from Cf) Af ne H + +/-- If an antitone function is continuous at the indexed supremum of a bounded function on +a nonempty `Sort`, then it sends this supremum to the infimum of the composition. -/ +lemma antitone.map_csupr_of_continuous_at {f : α → β} {g : γ → α} + (Cf : continuous_at f (⨆ i, g i)) (Af : antitone f) (H : bdd_above (range g)) : + f (⨆ i, g i) = ⨅ i, f (g i) := +monotone.map_csupr_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (⨆ i, g i), from Cf) Af H + +/-- If an antitone function is continuous at the infimum of a nonempty bounded below set `s`, +then it sends this infimum to the supremum of the image of `s`. -/ +lemma antitone.map_cInf_of_continuous_at {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) + (Af : antitone f) (ne : s.nonempty) (H : bdd_below s) : + f (Inf s) = Sup (f '' s) := +monotone.map_cInf_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (Inf s), from Cf) Af ne H + +/-- A continuous antitone function sends indexed infimum to indexed supremum in conditionally +complete linear order, under a boundedness assumption. -/ +lemma antitone.map_cinfi_of_continuous_at {f : α → β} {g : γ → α} + (Cf : continuous_at f (⨅ i, g i)) (Af : antitone f) (H : bdd_below (range g)) : + f (⨅ i, g i) = ⨆ i, f (g i) := +monotone.map_cinfi_of_continuous_at + (show continuous_at (order_dual.to_dual ∘ f) (⨅ i, g i), from Cf) Af H /-- A monotone map has a limit to the left of any point `x`, equal to `Sup (f '' (Iio x))`. -/ -lemma monotone.tendsto_nhds_within_Iio - {α : Type*} [linear_order α] [topological_space α] [order_topology α] +lemma monotone.tendsto_nhds_within_Iio {α β : Type*} + [linear_order α] [topological_space α] [order_topology α] + [conditionally_complete_linear_order β] [topological_space β] [order_topology β] {f : α → β} (Mf : monotone f) (x : α) : tendsto f (𝓝[<] x) (𝓝 (Sup (f '' (Iio x)))) := begin @@ -2888,11 +2986,12 @@ begin end /-- A monotone map has a limit to the right of any point `x`, equal to `Inf (f '' (Ioi x))`. -/ -lemma monotone.tendsto_nhds_within_Ioi - {α : Type*} [linear_order α] [topological_space α] [order_topology α] +lemma monotone.tendsto_nhds_within_Ioi {α β : Type*} + [linear_order α] [topological_space α] [order_topology α] + [conditionally_complete_linear_order β] [topological_space β] [order_topology β] {f : α → β} (Mf : monotone f) (x : α) : tendsto f (𝓝[>] x) (𝓝 (Inf (f '' (Ioi x)))) := -@monotone.tendsto_nhds_within_Iio βᵒᵈ _ _ _ αᵒᵈ _ _ _ f Mf.dual x +@monotone.tendsto_nhds_within_Iio αᵒᵈ βᵒᵈ _ _ _ _ _ _ f Mf.dual x end conditionally_complete_linear_order diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index e0af73d699854..75901a2fa3d0f 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -185,3 +185,109 @@ end end conditionally_complete_linear_order end liminf_limsup + +section monotone + +variables {ι R S : Type*} {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] + +/-- An antitone function between complete linear ordered spaces sends a `filter.Limsup` +to the `filter.liminf` of the image if it is continuous at the `Limsup`. -/ +lemma antitone.map_Limsup_of_continuous_at {F : filter R} [ne_bot F] + {f : R → S} (f_decr : antitone f) (f_cont : continuous_at f (F.Limsup)) : + f (F.Limsup) = F.liminf f := +begin + apply le_antisymm, + { have A : {a : R | ∀ᶠ (n : R) in F, n ≤ a}.nonempty, from ⟨⊤, by simp⟩, + rw [Limsup, (f_decr.map_Inf_of_continuous_at' f_cont A)], + apply le_of_forall_lt, + assume c hc, + simp only [liminf, Liminf, lt_Sup_iff, eventually_map, set.mem_set_of_eq, exists_prop, + set.mem_image, exists_exists_and_eq_and] at hc ⊢, + rcases hc with ⟨d, hd, h'd⟩, + refine ⟨f d, _, h'd⟩, + filter_upwards [hd] with x hx using f_decr hx }, + { rcases eq_or_lt_of_le (bot_le : ⊥ ≤ F.Limsup) with h|Limsup_ne_bot, + { rw ← h, + apply liminf_le_of_frequently_le, + apply frequently_of_forall, + assume x, + exact f_decr bot_le }, + by_cases h' : ∃ c, c < F.Limsup ∧ set.Ioo c F.Limsup = ∅, + { rcases h' with ⟨c, c_lt, hc⟩, + have B : ∃ᶠ n in F, F.Limsup ≤ n, + { apply (frequently_lt_of_lt_Limsup (by is_bounded_default) c_lt).mono, + assume x hx, + by_contra', + have : (set.Ioo c F.Limsup).nonempty := ⟨x, ⟨hx, this⟩⟩, + simpa [hc] }, + apply liminf_le_of_frequently_le, + exact B.mono (λ x hx, f_decr hx) }, + by_contra' H, + obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.Limsup, set.Ioc l F.Limsup ⊆ {x : R | f x < F.liminf f}, + from exists_Ioc_subset_of_mem_nhds ((tendsto_order.1 f_cont.tendsto).2 _ H) + ⟨⊥, Limsup_ne_bot⟩, + obtain ⟨m, l_m, m_lt⟩ : (set.Ioo l F.Limsup).nonempty, + { contrapose! h', + refine ⟨l, l_lt, by rwa set.not_nonempty_iff_eq_empty at h'⟩ }, + have B : F.liminf f ≤ f m, + { apply liminf_le_of_frequently_le, + apply (frequently_lt_of_lt_Limsup (by is_bounded_default) m_lt).mono, + assume x hx, + exact f_decr hx.le }, + have I : f m < F.liminf f := h'l ⟨l_m, m_lt.le⟩, + exact lt_irrefl _ (B.trans_lt I) } +end + +/-- A continuous antitone function between complete linear ordered spaces sends a `filter.limsup` +to the `filter.liminf` of the images. -/ +lemma antitone.map_limsup_of_continuous_at + {f : R → S} (f_decr : antitone f) (a : ι → R) (f_cont : continuous_at f (F.limsup a)) : + f (F.limsup a) = F.liminf (f ∘ a) := +f_decr.map_Limsup_of_continuous_at f_cont + +/-- An antitone function between complete linear ordered spaces sends a `filter.Liminf` +to the `filter.limsup` of the image if it is continuous at the `Liminf`. -/ +lemma antitone.map_Liminf_of_continuous_at {F : filter R} [ne_bot F] + {f : R → S} (f_decr : antitone f) (f_cont : continuous_at f (F.Liminf)) : + f (F.Liminf) = F.limsup f := +@antitone.map_Limsup_of_continuous_at + (order_dual R) (order_dual S) _ _ _ _ _ _ _ _ f f_decr.dual f_cont + +/-- A continuous antitone function between complete linear ordered spaces sends a `filter.liminf` +to the `filter.limsup` of the images. -/ +lemma antitone.map_liminf_of_continuous_at + {f : R → S} (f_decr : antitone f) (a : ι → R) (f_cont : continuous_at f (F.liminf a)) : + f (F.liminf a) = F.limsup (f ∘ a) := +f_decr.map_Liminf_of_continuous_at f_cont + +/-- A monotone function between complete linear ordered spaces sends a `filter.Limsup` +to the `filter.limsup` of the image if it is continuous at the `Limsup`. -/ +lemma monotone.map_Limsup_of_continuous_at {F : filter R} [ne_bot F] + {f : R → S} (f_incr : monotone f) (f_cont : continuous_at f (F.Limsup)) : + f (F.Limsup) = F.limsup f := +@antitone.map_Limsup_of_continuous_at R (order_dual S) _ _ _ _ _ _ _ _ f f_incr f_cont + +/-- A continuous monotone function between complete linear ordered spaces sends a `filter.limsup` +to the `filter.limsup` of the images. -/ +lemma monotone.map_limsup_of_continuous_at + {f : R → S} (f_incr : monotone f) (a : ι → R) (f_cont : continuous_at f (F.limsup a)) : + f (F.limsup a) = F.limsup (f ∘ a) := +f_incr.map_Limsup_of_continuous_at f_cont + +/-- A monotone function between complete linear ordered spaces sends a `filter.Liminf` +to the `filter.liminf` of the image if it is continuous at the `Liminf`. -/ +lemma monotone.map_Liminf_of_continuous_at {F : filter R} [ne_bot F] + {f : R → S} (f_incr : monotone f) (f_cont : continuous_at f (F.Liminf)) : + f (F.Liminf) = F.liminf f := +@antitone.map_Liminf_of_continuous_at R (order_dual S) _ _ _ _ _ _ _ _ f f_incr f_cont + +/-- A continuous monotone function between complete linear ordered spaces sends a `filter.liminf` +to the `filter.liminf` of the images. -/ +lemma monotone.map_liminf_of_continuous_at + {f : R → S} (f_incr : monotone f) (a : ι → R) (f_cont : continuous_at f (F.liminf a)) : + f (F.liminf a) = F.liminf (f ∘ a) := +f_incr.map_Liminf_of_continuous_at f_cont + +end monotone diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index a5d81dc5fc0c8..ce78ebe87cd99 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -461,8 +461,8 @@ begin casesI is_empty_or_nonempty ι, { rw [infi_of_empty, infi_of_empty, mul_top, if_neg], exact mt h0 (not_nonempty_iff.2 ‹_›) }, - { exact (map_infi_of_continuous_at_of_monotone' (ennreal.continuous_at_const_mul H) - ennreal.mul_left_mono).symm } } + { exact (ennreal.mul_left_mono.map_infi_of_continuous_at' + (ennreal.continuous_at_const_mul H)).symm } } end lemma infi_mul_left {ι} [nonempty ι] {f : ι → ℝ≥0∞} {a : ℝ≥0∞} @@ -519,7 +519,7 @@ protected lemma tendsto_inv_nat_nhds_zero : tendsto (λ n : ℕ, (n : ℝ≥0∞ ennreal.inv_top ▸ ennreal.tendsto_inv_iff.2 tendsto_nat_nhds_top lemma supr_add {ι : Sort*} {s : ι → ℝ≥0∞} [h : nonempty ι] : supr s + a = ⨆b, s b + a := -map_supr_of_continuous_at_of_monotone' (continuous_at_id.add continuous_at_const) $ +monotone.map_supr_of_continuous_at' (continuous_at_id.add continuous_at_const) $ monotone_id.add monotone_const lemma bsupr_add' {ι : Sort*} {p : ι → Prop} (h : ∃ i, p i) {f : ι → ℝ≥0∞} : @@ -593,7 +593,7 @@ begin by_cases hf : ∀ i, f i = 0, { obtain rfl : f = (λ _, 0), from funext hf, simp only [supr_zero_eq_zero, mul_zero] }, - { refine map_supr_of_continuous_at_of_monotone _ (monotone_id.const_mul' _) (mul_zero a), + { refine (monotone_id.const_mul' _).map_supr_of_continuous_at _ (mul_zero a), refine ennreal.tendsto.const_mul tendsto_id (or.inl _), exact mt supr_eq_zero.1 hf } end diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index a1d45f7937fe2..15e54523b73a8 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -106,7 +106,7 @@ private lemma glue_dist_triangle (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) have : (⨅ p, dist z (Φ p) + dist x (Ψ p)) ≤ (⨅ p, dist y (Φ p) + dist x (Ψ p)) + dist y z, { have : (⨅ p, dist y (Φ p) + dist x (Ψ p)) + dist y z = infi ((λt, t + dist y z) ∘ (λp, dist y (Φ p) + dist x (Ψ p))), - { refine map_cinfi_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ + { refine monotone.map_cinfi_of_continuous_at (continuous_at_id.add continuous_at_const) _ (B _ _), intros x y hx, simpa }, rw [this, comp], @@ -124,7 +124,7 @@ private lemma glue_dist_triangle (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) have : (⨅ p, dist z (Φ p) + dist x (Ψ p)) ≤ dist x y + ⨅ p, dist z (Φ p) + dist y (Ψ p), { have : dist x y + (⨅ p, dist z (Φ p) + dist y (Ψ p)) = infi ((λt, dist x y + t) ∘ (λp, dist z (Φ p) + dist y (Ψ p))), - { refine map_cinfi_of_continuous_at_of_monotone (continuous_at_const.add continuous_at_id) _ + { refine monotone.map_cinfi_of_continuous_at (continuous_at_const.add continuous_at_id) _ (B _ _), intros x y hx, simpa }, rw [this, comp], @@ -142,7 +142,7 @@ private lemma glue_dist_triangle (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) have : (⨅ p, dist x (Φ p) + dist z (Ψ p)) ≤ dist x y + ⨅ p, dist y (Φ p) + dist z (Ψ p), { have : dist x y + (⨅ p, dist y (Φ p) + dist z (Ψ p)) = infi ((λt, dist x y + t) ∘ (λp, dist y (Φ p) + dist z (Ψ p))), - { refine map_cinfi_of_continuous_at_of_monotone (continuous_at_const.add continuous_at_id) _ + { refine monotone.map_cinfi_of_continuous_at (continuous_at_const.add continuous_at_id) _ (B _ _), intros x y hx, simpa }, rw [this, comp], @@ -160,7 +160,7 @@ private lemma glue_dist_triangle (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) have : (⨅ p, dist x (Φ p) + dist z (Ψ p)) ≤ (⨅ p, dist x (Φ p) + dist y (Ψ p)) + dist y z, { have : (⨅ p, dist x (Φ p) + dist y (Ψ p)) + dist y z = infi ((λt, t + dist y z) ∘ (λp, dist x (Φ p) + dist y (Ψ p))), - { refine map_cinfi_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ + { refine monotone.map_cinfi_of_continuous_at (continuous_at_id.add continuous_at_const) _ (B _ _), intros x y hx, simpa }, rw [this, comp], diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index fc2e8017e8830..f0623b8f2058b 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -369,12 +369,12 @@ begin -- (here the addition of `dist f g`) preserve infimum and supremum have E1 : ∀ x, (⨅ y, g (inl x, inr y)) + dist f g = ⨅ y, g (inl x, inr y) + dist f g, { assume x, - refine map_cinfi_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ _, + refine monotone.map_cinfi_of_continuous_at (continuous_at_id.add continuous_at_const) _ _, { assume x y hx, simpa }, { show bdd_below (range (λ (y : Y), g (inl x, inr y))), from ⟨cg, forall_range_iff.2(λi, Hcg _)⟩ } }, have E2 : (⨆ x, ⨅ y, g (inl x, inr y)) + dist f g = ⨆ x, (⨅ y, g (inl x, inr y)) + dist f g, - { refine map_csupr_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ _, + { refine monotone.map_csupr_of_continuous_at (continuous_at_id.add continuous_at_const) _ _, { assume x y hx, simpa }, { simpa using HD_bound_aux1 _ 0 } }, -- deduce the result from the above two steps @@ -398,12 +398,12 @@ begin -- (here the addition of `dist f g`) preserve infimum and supremum have E1 : ∀ y, (⨅ x, g (inl x, inr y)) + dist f g = ⨅ x, g (inl x, inr y) + dist f g, { assume y, - refine map_cinfi_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ _, + refine monotone.map_cinfi_of_continuous_at (continuous_at_id.add continuous_at_const) _ _, { assume x y hx, simpa }, { show bdd_below (range (λx:X, g (inl x, inr y))), from ⟨cg, forall_range_iff.2 (λi, Hcg _)⟩ } }, have E2 : (⨆ y, ⨅ x, g (inl x, inr y)) + dist f g = ⨆ y, (⨅ x, g (inl x, inr y)) + dist f g, - { refine map_csupr_of_continuous_at_of_monotone (continuous_at_id.add continuous_at_const) _ _, + { refine monotone.map_csupr_of_continuous_at (continuous_at_id.add continuous_at_const) _ _, { assume x y hx, simpa }, { simpa using HD_bound_aux2 _ 0 } }, -- deduce the result from the above two steps From ddfc7d3cf940fcbc0d9b6800285f9896791dda00 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Jul 2022 06:57:41 +0000 Subject: [PATCH 085/173] feat(category_theory): a category with a small detecting set is well-powered (#15238) --- src/category_theory/generator.lean | 43 ++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index 522630650fe36..23f20b7069223 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -6,6 +6,8 @@ Authors: Markus Himmel import category_theory.balanced import category_theory.limits.opposites import category_theory.limits.shapes.zero_morphisms +import category_theory.subobject.lattice +import category_theory.subobject.well_powered import data.set.opposite /-! @@ -237,6 +239,43 @@ lemma is_codetecting_empty_of_groupoid [∀ {X Y : C} (f : X ⟶ Y), is_iso f] : end empty +section well_powered + +namespace subobject + +lemma eq_of_le_of_is_detecting {𝒢 : set C} (h𝒢 : is_detecting 𝒢) {X : C} (P Q : subobject X) + (h₁ : P ≤ Q) (h₂ : ∀ (G ∈ 𝒢) {f : G ⟶ X}, Q.factors f → P.factors f) : P = Q := +begin + suffices : is_iso (of_le _ _ h₁), + { exactI le_antisymm h₁ (le_of_comm (inv (of_le _ _ h₁)) (by simp)) }, + refine h𝒢 _ (λ G hG f, _), + have : P.factors (f ≫ Q.arrow) := h₂ _ hG ((factors_iff _ _).2 ⟨_, rfl⟩), + refine ⟨factor_thru _ _ this, _, λ g (hg : g ≫ _ = f), _⟩, + { simp only [← cancel_mono Q.arrow, category.assoc, of_le_arrow, factor_thru_arrow] }, + { simp only [← cancel_mono (subobject.of_le _ _ h₁), ← cancel_mono Q.arrow, hg, + category.assoc, of_le_arrow, factor_thru_arrow] } +end + +lemma inf_eq_of_is_detecting [has_pullbacks C] {𝒢 : set C} (h𝒢 : is_detecting 𝒢) {X : C} + (P Q : subobject X) (h : ∀ (G ∈ 𝒢) {f : G ⟶ X}, P.factors f → Q.factors f) : P ⊓ Q = P := +eq_of_le_of_is_detecting h𝒢 _ _ _root_.inf_le_left (λ G hG f hf, (inf_factors _).2 ⟨hf, h _ hG hf⟩) + +lemma eq_of_is_detecting [has_pullbacks C] {𝒢 : set C} (h𝒢 : is_detecting 𝒢) {X : C} + (P Q : subobject X) (h : ∀ (G ∈ 𝒢) {f : G ⟶ X}, P.factors f ↔ Q.factors f) : P = Q := +calc P = P ⊓ Q : eq.symm $ inf_eq_of_is_detecting h𝒢 _ _ $ λ G hG f hf, (h G hG).1 hf + ... = Q ⊓ P : inf_comm + ... = Q : inf_eq_of_is_detecting h𝒢 _ _ $ λ G hG f hf, (h G hG).2 hf + +end subobject + +/-- A category with pullbacks and a small detecting set is well-powered. -/ +lemma well_powered_of_is_detecting [has_pullbacks C] {𝒢 : set C} [small.{v} 𝒢] + (h𝒢 : is_detecting 𝒢) : well_powered C := +⟨λ X, @small_of_injective _ _ _ (λ P : subobject X, { f : Σ G : 𝒢, G.1 ⟶ X | P.factors f.2 }) $ + λ P Q h, subobject.eq_of_is_detecting h𝒢 _ _ (by simpa [set.ext_iff] using h)⟩ + +end well_powered + /-- We say that `G` is a separator if the functor `C(G, -)` is faithful. -/ def is_separator (G : C) : Prop := is_separating ({G} : set C) @@ -442,4 +481,8 @@ begin rwa [is_iso_iff_bijective, function.bijective_iff_exists_unique] } end +lemma well_powered_of_is_detector [has_pullbacks C] (G : C) (hG : is_detector G) : + well_powered C := +well_powered_of_is_detecting hG + end category_theory From 34aa5390984c55d7d0ce130ffc94749d423db81f Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Wed, 27 Jul 2022 07:39:55 +0000 Subject: [PATCH 086/173] feat(ring_theory): the Ore localization of a ring (#14348) Implements the Ore localization of a potentially non-commutative ring and shows its universal property. Co-authored-by: Kevin Klinge (@jyn-n) --- docs/references.bib | 11 + src/data/subtype.lean | 3 + src/ring_theory/ore_localization/basic.lean | 785 ++++++++++++++++++ src/ring_theory/ore_localization/ore_set.lean | 102 +++ 4 files changed, 901 insertions(+) create mode 100644 src/ring_theory/ore_localization/basic.lean create mode 100644 src/ring_theory/ore_localization/ore_set.lean diff --git a/docs/references.bib b/docs/references.bib index 73b46727a3e5e..6dd8ea652d9ab 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1670,6 +1670,17 @@ @Book{ simon2011 collection = {Cambridge Tracts in Mathematics} } +@Article{ skoda2006, + author = {{\v{S}}koda, Zoran}, + title = {Noncommutative localization in noncommutative geometry}, + journal = {London Math. Soc. Lecture Note Series}, + fjournal = {London Mathematical Society Lecture Note Series}, + volume = {330}, + pages = {220--313}, + year = {2006}, + url = {https://doi.org/10.48550/arXiv.math/0403276} +} + @Book{ soare1987, author = {Soare, Robert I.}, title = {Recursively enumerable sets and degrees}, diff --git a/src/data/subtype.lean b/src/data/subtype.lean index 7347ff7aa6a5d..3909be2e51588 100644 --- a/src/data/subtype.lean +++ b/src/data/subtype.lean @@ -88,6 +88,9 @@ ext_iff theorem mk_eq_mk {a h a' h'} : @mk α p a h = @mk α p a' h' ↔ a = a' := ext_iff +lemma coe_eq_of_eq_mk {a : {a // p a}} {b : α} (h : ↑a = b) : + a = ⟨b, h ▸ a.2⟩ := subtype.ext h + theorem coe_eq_iff {a : {a // p a}} {b : α} : ↑a = b ↔ ∃ h, a = ⟨b, h⟩ := ⟨λ h, h ▸ ⟨a.2, (coe_eta _ _).symm⟩, λ ⟨hb, ha⟩, ha.symm ▸ rfl⟩ diff --git a/src/ring_theory/ore_localization/basic.lean b/src/ring_theory/ore_localization/basic.lean new file mode 100644 index 0000000000000..7adf5f507509f --- /dev/null +++ b/src/ring_theory/ore_localization/basic.lean @@ -0,0 +1,785 @@ +/- +Copyright (c) 2022 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer, Kevin Klinge +-/ +import algebra.group_with_zero.basic +import group_theory.congruence +import group_theory.monoid_localization +import ring_theory.non_zero_divisors +import ring_theory.ore_localization.ore_set +import tactic.noncomm_ring + +/-! + +# Localization over right Ore sets. + +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 +common monoid localization. Finally we show that in a ring without zero divisors, taking the Ore +localization at `R - {0}` results in a division ring. + +## Notations + +Introduces the notation `R[S⁻¹]` for the Ore localization of a monoid `R` at a right Ore +subset `S`. Also defines a new heterogeneos division notation `r /ₒ s` for a numerator `r : R` and +a denominator `s : S`. + +## References + +* +* [Zoran Škoda, *Noncommutative localization in noncommutative geometry*][skoda2006] + + +## Tags +localization, Ore, non-commutative + +-/ + +universe u + +open ore_localization + +namespace ore_localization + +variables (R : Type*) [monoid R] (S : submonoid R) [ore_set S] + +/-- The setoid on `R × S` used for the Ore localization. -/ +def ore_eqv : setoid (R × S) := +{ r := λ rs rs', ∃ (u : S) (v : R), rs'.1 * u = rs.1 * v + ∧ (rs'.2 : R) * u = rs.2 * v, + iseqv := + begin + refine ⟨_, _, _⟩, + { rintro ⟨r,s⟩, use 1, use 1, simp [submonoid.one_mem] }, + { rintros ⟨r, s⟩ ⟨r', s'⟩ ⟨u, v, hru, hsu⟩, + rcases ore_condition (s : R) s' with ⟨r₂, s₂, h₁⟩, + rcases ore_condition r₂ u with ⟨r₃, s₃, h₂⟩, + have : (s : R) * ((v : R) * r₃) = (s : R) * (s₂ * s₃), + { assoc_rw [h₁, h₂, hsu], symmetry, apply mul_assoc }, + rcases ore_left_cancel (v * r₃) (s₂ * s₃) s this with ⟨w, hw⟩, + use s₂ * s₃ * w, use u * r₃ * w, split; simp only [submonoid.coe_mul], + { assoc_rw [hru, ←hw], simp [mul_assoc] }, + { assoc_rw [hsu, ←hw], simp [mul_assoc] } }, + { rintros ⟨r₁, s₁⟩ ⟨r₂, s₂⟩ ⟨r₃, s₃⟩ ⟨u, v, hur₁, hs₁u⟩ ⟨u', v', hur₂, hs₂u⟩, + rcases ore_condition v' u with ⟨r', s', h⟩, + use u' * s', use v * r', split; simp only [submonoid.coe_mul], + { assoc_rw [hur₂, h, hur₁, mul_assoc] }, + { assoc_rw [hs₂u, h, hs₁u, mul_assoc] } } + end } + +end ore_localization + +/-- The ore localization of a monoid and a submonoid fulfilling the ore condition. -/ +def ore_localization (R : Type*) [monoid R] (S : submonoid R) [ore_set S] := +quotient (ore_localization.ore_eqv R S) + +namespace ore_localization + +section monoid + +variables {R : Type*} [monoid R] {S : submonoid R} + +variables (R S) [ore_set S] + +notation R `[`:1075 S `⁻¹]`:1075 := ore_localization R S + +local attribute [instance] ore_eqv + +variables {R S} + +/-- The division in the ore localization `R[S⁻¹]`, as a fraction of an element of `R` and `S`. -/ +def ore_div (r : R) (s : S) : R[S⁻¹] := quotient.mk (r, s) + +infixl ` /ₒ `:70 := ore_div + +@[elab_as_eliminator] +protected lemma ind {β : R [S ⁻¹] → Prop} (c : ∀ (r : R) (s : S), β (r /ₒ s)) : ∀ q, β q := +by { apply quotient.ind, rintro ⟨r, s⟩, exact c r s } + +lemma ore_div_eq_iff {r₁ r₂ : R} {s₁ s₂ : S} : + r₁ /ₒ s₁ = r₂ /ₒ s₂ ↔ (∃ (u : S) (v : R), r₂ * u = r₁ * v ∧ (s₂ : R) * u = s₁ * v) := +quotient.eq' + +/-- A fraction `r /ₒ s` is equal to its expansion by an arbitrary factor `t` if `s * t ∈ S`. -/ +protected lemma expand (r : R) (s : S) (t : R) (hst : (s : R) * t ∈ S) : + r /ₒ s = (r * t) /ₒ (⟨s * t, hst⟩) := +by { apply quotient.sound, refine ⟨s, t * s, _, _⟩; dsimp; rw [mul_assoc]; refl } + +/-- A fraction is equal to its expansion by an factor from s. -/ +protected lemma expand' (r : R) (s s' : S) : r /ₒ s = (r * s') /ₒ (s * s') := +ore_localization.expand r s s' (by norm_cast; apply set_like.coe_mem) + +/-- Fractions which differ by a factor of the numerator can be proven equal if +those factors expand to equal elements of `R`. -/ +protected lemma eq_of_num_factor_eq {r r' r₁ r₂ : R} {s t : S} + (h : r * t = r' * t) : (r₁ * r * r₂) /ₒ s = (r₁ * r' * r₂) /ₒ s := +begin + rcases ore_condition r₂ t with ⟨r₂',t', hr₂⟩, + calc (r₁ * r * r₂) /ₒ s = (r₁ * r * r₂ * t') /ₒ (s * t') : ore_localization.expand _ _ t' _ + ... = ((r₁ * r) * (r₂ * t')) /ₒ (s * t') : by simp [←mul_assoc] + ... = ((r₁ * r) * (t * r₂')) /ₒ (s * t') : by rw hr₂ + ... = (r₁ * (r * t) * r₂') /ₒ (s * t') : by simp [←mul_assoc] + ... = (r₁ * (r' * t) * r₂') /ₒ (s * t') : by rw h + ... = (r₁ * r' * (t * r₂')) /ₒ (s * t') : by simp [←mul_assoc] + ... = (r₁ * r' * (r₂ * t')) /ₒ (s * t') : by rw hr₂ + ... = (r₁ * r' * r₂ * t') /ₒ (s * t') : by simp [←mul_assoc] + ... = (r₁ * r' * r₂) /ₒ s : by symmetry; apply ore_localization.expand +end + +/-- A function or predicate over `R` and `S` can be lifted to `R[S⁻¹]` if it is invariant +under expansion on the right. -/ +def lift_expand {C : Sort*} (P : R → S → C) + (hP : ∀ (r t : R) (s : S) (ht : ((s : R) * t) ∈ S), P r s = P (r * t) ⟨s * t, ht⟩) : + R[S⁻¹] → C := +quotient.lift (λ (p : R × S), P p.1 p.2) $ λ p q pq, +begin + cases p with r₁ s₁, cases q with r₂ s₂, rcases pq with ⟨u, v, hr₂, hs₂⟩, + dsimp at *, + have s₁vS : (s₁ : R) * v ∈ S, + { rw [←hs₂, ←S.coe_mul], exact set_like.coe_mem (s₂ * u) }, + replace hs₂ : s₂ * u = ⟨(s₁ : R) * v, s₁vS⟩, { ext, simp [hs₂] }, + rw [hP r₁ v s₁ s₁vS, hP r₂ u s₂ (by { norm_cast, rw hs₂, assumption }), hr₂], + simpa [← hs₂] +end + +@[simp] +lemma lift_expand_of {C : Sort*} {P : R → S → C} + {hP : ∀ (r t : R) (s : S) (ht : ((s : R) * t) ∈ S), P r s = P (r * t) ⟨s * t, ht⟩} + (r : R) (s : S) : + lift_expand P hP (r /ₒ s) = P r s := rfl + +/-- A version of `lift_expand` used to simultaneously lift functions with two arguments +in ``R[S⁻¹]`.-/ +def lift₂_expand {C : Sort*} (P : R → S → R → S → C) + (hP : ∀ (r₁ t₁ : R) (s₁ : S) (ht₁ : (s₁ : R) * t₁ ∈ S) + (r₂ t₂ : R) (s₂ : S) (ht₂ : (s₂ : R) * t₂ ∈ S), + P r₁ s₁ r₂ s₂ = P (r₁ * t₁) ⟨s₁ * t₁, ht₁⟩ (r₂ * t₂) ⟨s₂ * t₂, ht₂⟩) : R[S⁻¹] → R[S⁻¹] → C := +lift_expand + (λ r₁ s₁, lift_expand (P r₁ s₁) $ λ r₂ t₂ s₂ ht₂, by simp [hP r₁ 1 s₁ (by simp) r₂ t₂ s₂ ht₂]) $ + λ r₁ t₁ s₁ ht₁, + begin + ext x, induction x using ore_localization.ind with r₂ s₂, + rw [lift_expand_of, lift_expand_of, hP r₁ t₁ s₁ ht₁ r₂ 1 s₂ (by simp)], simp, + end + +@[simp] +lemma lift₂_expand_of {C : Sort*} {P : R → S → R → S → C} + {hP : ∀ (r₁ t₁ : R) (s₁ : S) (ht₁ : (s₁ : R) * t₁ ∈ S) + (r₂ t₂ : R) (s₂ : S) (ht₂ : (s₂ : R) * t₂ ∈ S), + P r₁ s₁ r₂ s₂ = P (r₁ * t₁) ⟨s₁ * t₁, ht₁⟩ (r₂ * t₂) ⟨s₂ * t₂, ht₂⟩} + (r₁ : R) (s₁ : S) (r₂ : R) (s₂ : S) : + lift₂_expand P hP (r₁ /ₒ s₁) (r₂ /ₒ s₂) = P r₁ s₁ r₂ s₂ := +rfl + +private def mul' (r₁ : R) (s₁ : S) (r₂ : R) (s₂ : S) : R[S⁻¹] := + (r₁ * ore_num r₂ s₁) /ₒ (s₂ * ore_denom r₂ s₁) + +private lemma mul'_char (r₁ r₂ : R) (s₁ s₂ : S) (u : S) (v : R) (huv : r₂ * (u : R) = s₁ * v) : + mul' r₁ s₁ r₂ s₂ = (r₁ * v) /ₒ (s₂ * u) := +begin + simp only [mul'], + have h₀ := ore_eq r₂ s₁, set v₀ := ore_num r₂ s₁, set u₀ := ore_denom r₂ s₁, + rcases ore_condition (u₀ : R) u with ⟨r₃, s₃, h₃⟩, + have := + calc (s₁ : R) * (v * r₃) = r₂ * u * r₃ : by assoc_rw ←huv; symmetry; apply mul_assoc + ... = r₂ * u₀ * s₃ : by assoc_rw ←h₃; refl + ... = s₁ * (v₀ * s₃) : by assoc_rw h₀; apply mul_assoc, + rcases ore_left_cancel _ _ _ this with ⟨s₄, hs₄⟩, + symmetry, rw ore_div_eq_iff, + use s₃ * s₄, use r₃ * s₄, simp only [submonoid.coe_mul], split, + { assoc_rw ←hs₄, simp only [mul_assoc] }, + { assoc_rw h₃, simp only [mul_assoc] } +end + +/-- The multiplication on the Ore localization of monoids. -/ +protected def mul : R[S⁻¹] → R[S⁻¹] → R[S⁻¹] := +lift₂_expand mul' $ λ r₂ p s₂ hp r₁ r s₁ hr, +begin + have h₁ := ore_eq r₁ s₂, set r₁' := ore_num r₁ s₂, set s₂' := ore_denom r₁ s₂, + rcases ore_condition (↑s₂ * r₁') ⟨s₂ * p, hp⟩ with ⟨p', s_star, h₂⟩, dsimp at h₂, + rcases ore_condition r (s₂' * s_star) with ⟨p_flat, s_flat, h₃⟩, simp only [S.coe_mul] at h₃, + have : r₁ * r * s_flat = s₂ * p * (p' * p_flat), + { rw [←mul_assoc, ←h₂, ←h₁, mul_assoc, h₃], simp only [mul_assoc] }, + rw mul'_char (r₂ * p) (r₁ * r) ⟨↑s₂ * p, hp⟩ ⟨↑s₁ * r, hr⟩ _ _ this, clear this, + have hsssp : ↑s₁ * ↑s₂' * ↑s_star * p_flat ∈ S, + { rw [mul_assoc, mul_assoc, ←mul_assoc ↑s₂', ←h₃, ←mul_assoc], + exact S.mul_mem hr (set_like.coe_mem s_flat) }, + have : (⟨↑s₁ * r, hr⟩ : S) * s_flat = ⟨s₁ * s₂' * s_star * p_flat, hsssp⟩, + { ext, simp only [set_like.coe_mk, submonoid.coe_mul], + rw [mul_assoc, h₃, ←mul_assoc, ←mul_assoc] }, + rw this, clear this, + rcases ore_left_cancel (p * p') (r₁' * ↑s_star) s₂ (by simp [←mul_assoc, h₂]) with ⟨s₂'', h₂''⟩, + rw [←mul_assoc, mul_assoc r₂, ore_localization.eq_of_num_factor_eq h₂''], + norm_cast at ⊢ hsssp, rw [←ore_localization.expand _ _ _ hsssp, ←mul_assoc], + apply ore_localization.expand +end + +instance : has_mul R[S⁻¹] := ⟨ore_localization.mul⟩ + +lemma ore_div_mul_ore_div {r₁ r₂ : R} {s₁ s₂ : S} : + (r₁ /ₒ s₁) * (r₂ /ₒ s₂) = (r₁ * ore_num r₂ s₁) /ₒ (s₂ * ore_denom r₂ s₁) := rfl + +/-- A characterization lemma for the multiplication on the Ore localization, allowing for a choice +of Ore numerator and Ore denominator. -/ +lemma ore_div_mul_char (r₁ r₂ : R) (s₁ s₂ : S) (r' : R) (s' : S) + (huv : r₂ * (s' : R) = s₁ * r') : (r₁ /ₒ s₁) * (r₂ /ₒ s₂) = (r₁ * r') /ₒ (s₂ * s') := +mul'_char r₁ r₂ s₁ s₂ s' r' huv + +/-- Another characterization lemma for the multiplication on the Ore localizaion delivering +Ore witnesses and conditions bundled in a sigma type. -/ +def ore_div_mul_char' (r₁ r₂ : R) (s₁ s₂ : S) : + Σ' r' : R, Σ' s' : S, r₂ * (s' : R) = s₁ * r' + ∧ (r₁ /ₒ s₁) * (r₂ /ₒ s₂) = (r₁ * r') /ₒ (s₂ * s') := +⟨ore_num r₂ s₁, ore_denom r₂ s₁, ore_eq r₂ s₁, ore_div_mul_ore_div⟩ + +instance : has_one R[S⁻¹] := ⟨1 /ₒ 1⟩ + +protected lemma one_def : (1 : R[S⁻¹]) = 1 /ₒ 1 := rfl + +instance : inhabited R[S⁻¹] := ⟨1⟩ + +@[simp] +protected lemma div_eq_one' {r : R} (hr : r ∈ S) : r /ₒ ⟨r, hr⟩ = 1 := +by { rw [ore_localization.one_def, ore_div_eq_iff], exact ⟨⟨r, hr⟩, 1, by simp, by simp⟩ } + +@[simp] +protected lemma div_eq_one {s : S} : (s : R) /ₒ s = 1 := +by { cases s; apply ore_localization.div_eq_one' } + +protected lemma one_mul (x : R[S⁻¹]) : 1 * x = x := +begin + induction x using ore_localization.ind with r s, + simp [ore_localization.one_def, ore_div_mul_char (1 : R) r (1 : S) s r 1 (by simp)] +end + +protected lemma mul_one (x : R[S⁻¹]) : x * 1 = x := +begin + induction x using ore_localization.ind with r s, + simp [ore_localization.one_def, ore_div_mul_char r 1 s 1 1 s (by simp)] +end + +protected lemma mul_assoc (x y z : R[S⁻¹]) : x * y * z = x * (y * z) := +begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + induction z using ore_localization.ind with r₃ s₃, + rcases ore_div_mul_char' r₁ r₂ s₁ s₂ with ⟨ra, sa, ha, ha'⟩, rw ha', clear ha', + rcases ore_div_mul_char' r₂ r₃ s₂ s₃ with ⟨rb, sb, hb, hb'⟩, rw hb', clear hb', + rcases ore_condition rb sa with ⟨rc, sc, hc⟩, + rw ore_div_mul_char (r₁ * ra) r₃ (s₂ * sa) s₃ rc (sb * sc) + (by { simp only [submonoid.coe_mul], assoc_rw [hb, hc] }), + rw [mul_assoc, ←mul_assoc s₃], + symmetry, apply ore_div_mul_char, + assoc_rw [hc, ←ha], apply mul_assoc +end + +instance : monoid R[S⁻¹] := +{ one_mul := ore_localization.one_mul, + mul_one := ore_localization.mul_one, + mul_assoc := ore_localization.mul_assoc, + .. ore_localization.has_mul, + .. ore_localization.has_one } + +protected lemma mul_inv (s s' : S) : ((s : R) /ₒ s') * (s' /ₒ s) = 1 := +by simp [ore_div_mul_char (s :R) s' s' s 1 1 (by simp)] + +@[simp] +protected lemma mul_one_div {r : R} {s t : S} : (r /ₒ s) * (1 /ₒ t) = r /ₒ (t * s) := +by simp [ore_div_mul_char r 1 s t 1 s (by simp)] + +@[simp] +protected lemma mul_cancel {r : R} {s t : S} : (r /ₒ s) * (s /ₒ t) = r /ₒ t := +by simp [ore_div_mul_char r s s t 1 1 (by simp)] + +@[simp] +protected lemma mul_cancel' {r₁ r₂ : R} {s t : S} : (r₁ /ₒ s) * ((s * r₂) /ₒ t) = (r₁ * r₂) /ₒ t := +by simp [ore_div_mul_char r₁ (s * r₂) s t r₂ 1 (by simp)] + +@[simp] +lemma div_one_mul {p r : R} {s : S} : + (r /ₒ 1) * (p /ₒ s) = (r * p) /ₒ s := --TODO use coercion r ↦ r /ₒ 1 +by simp [ore_div_mul_char r p 1 s p 1 (by simp)] + +/-- The fraction `s /ₒ 1` as a unit in `R[S⁻¹]`, where `s : S`. -/ +def numerator_unit (s : S) : units R[S⁻¹] := +{ val := (s : R) /ₒ 1, + inv := (1 : R) /ₒ s, + val_inv := ore_localization.mul_inv s 1, + inv_val := ore_localization.mul_inv 1 s } + +/-- The multiplicative homomorphism from `R` to `R[S⁻¹]`, mapping `r : R` to the +fraction `r /ₒ 1`. -/ +def numerator_hom : R →* R[S⁻¹] := +{ to_fun := λ r, r /ₒ 1, + map_one' := rfl, + map_mul' := λ r₁ r₂, div_one_mul.symm } + +lemma numerator_hom_apply {r : R} : numerator_hom r = r /ₒ (1 : S) := rfl + +lemma numerator_is_unit (s : S) : is_unit ((numerator_hom (s : R)) : R[S⁻¹]) := +⟨numerator_unit s, rfl⟩ + +section UMP + +variables {T : Type*} [monoid T] +variables (f : R →* T) (fS : S →* units T) +variables (hf : ∀ (s : S), f s = fS s) + +include f fS hf + +/-- The universal lift from a morphism `R →* T`, which maps elements of `S` to units of `T`, +to a morphism `R[S⁻¹] →* T`. -/ +def universal_mul_hom : R[S⁻¹] →* T := +{ to_fun := λ x, x.lift_expand (λ r s, (f r) * ((fS s)⁻¹ : units T)) $ λ r t s ht, + begin + have : ((fS ⟨s * t, ht⟩) : T) = fS s * f t, + { simp only [←hf, set_like.coe_mk, monoid_hom.map_mul] }, + conv_rhs { rw [monoid_hom.map_mul, ←mul_one (f r), ←units.coe_one, ←mul_left_inv (fS s)], + rw [units.coe_mul, ←mul_assoc, mul_assoc _ ↑(fS s), ←this, mul_assoc] }, + simp only [mul_one, units.mul_inv] + end, + map_one' := by rw [ore_localization.one_def, lift_expand_of]; simp, + map_mul' := λ x y, + begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + rcases ore_div_mul_char' r₁ r₂ s₁ s₂ with ⟨ra, sa, ha, ha'⟩, rw ha', clear ha', + rw [lift_expand_of, lift_expand_of, lift_expand_of], + conv_rhs { congr, skip, congr, + rw [←mul_one (f r₂), ←(fS sa).mul_inv, ←mul_assoc, ←hf, ←f.map_mul, ha, f.map_mul] }, + rw [mul_assoc, mul_assoc, mul_assoc, ←mul_assoc _ (f s₁), hf s₁, (fS s₁).inv_mul, one_mul, + f.map_mul, mul_assoc, fS.map_mul, ←units.coe_mul], refl + end } + +lemma universal_mul_hom_apply {r : R} {s : S} : + universal_mul_hom f fS hf (r /ₒ s) = (f r) * ((fS s)⁻¹ : units T) := +rfl + +lemma universal_mul_hom_commutes {r : R} : universal_mul_hom f fS hf (numerator_hom r) = f r := +by simp [numerator_hom_apply, universal_mul_hom_apply] + +/-- The universal morphism `universal_mul_hom` is unique. -/ +lemma universal_mul_hom_unique (φ : R[S⁻¹] →* T) (huniv : ∀ (r : R), φ (numerator_hom r) = f r) : + φ = universal_mul_hom f fS hf := +begin + ext, induction x using ore_localization.ind with r s, + rw [universal_mul_hom_apply, ←huniv r, numerator_hom_apply, ←mul_one (φ (r /ₒ s)), + ←units.coe_one, ←mul_right_inv (fS s), units.coe_mul, ←mul_assoc, + ←hf, ←huniv, ←φ.map_mul, numerator_hom_apply, ore_localization.mul_cancel], +end + +end UMP + +end monoid + +section comm_monoid + +variables {R : Type*} [comm_monoid R] {S : submonoid R} [ore_set S] + +lemma ore_div_mul_ore_div_comm {r₁ r₂ : R} {s₁ s₂ : S} : + (r₁ /ₒ s₁) * (r₂ /ₒ s₂) = (r₁ * r₂) /ₒ (s₁ * s₂) := +by rw [ore_div_mul_char r₁ r₂ s₁ s₂ r₂ s₁ (by simp [mul_comm]), mul_comm s₂] + +instance : comm_monoid R[S⁻¹] := +{ mul_comm := λ x y, + begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + rw [ore_div_mul_ore_div_comm, ore_div_mul_ore_div_comm, mul_comm r₁, mul_comm s₁], + end, + ..ore_localization.monoid } + +variables (R S) + +/-- The morphism `numerator_hom` is a monoid localization map in the case of commutative `R`. -/ +protected def localization_map : S.localization_map R[S⁻¹] := +{ to_fun := numerator_hom, + map_one' := rfl, + map_mul' := λ r₁ r₂, by simp, + map_units' := numerator_is_unit, + surj' := λ z, + begin + induction z using ore_localization.ind with r s, + use (r, s), dsimp, + rw [numerator_hom_apply, numerator_hom_apply], simp + end, + eq_iff_exists' := λ r₁ r₂, + begin + dsimp, split, + { 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 }, + { rintro ⟨s, h⟩, + rw [numerator_hom_apply, numerator_hom_apply, ore_div_eq_iff], + use s, use s, simp [h, one_mul] } + end } + +/-- If `R` is commutative, Ore localization and monoid localization are isomorphic. -/ +protected noncomputable def equiv_monoid_localization : localization S ≃* R[S⁻¹] := +localization.mul_equiv_of_quotient (ore_localization.localization_map R S) + +end comm_monoid + +section semiring + +variables {R : Type*} [semiring R] {S : submonoid R} [ore_set S] + +private def add'' (r₁ : R) (s₁ : S) (r₂ : R) (s₂ : S) : R[S⁻¹] := +(r₁ * ore_denom (s₁ : R) s₂ + r₂ * ore_num s₁ s₂) /ₒ (s₁ * ore_denom s₁ s₂) + +private lemma add''_char + (r₁ : R) (s₁ : S) (r₂ : R) (s₂ : S) + (rb : R) (sb : S) (hb : (s₁ : R) * sb = (s₂ : R) * rb) : +add'' r₁ s₁ r₂ s₂ = (r₁ * sb + r₂ * rb) /ₒ (s₁ * sb) := +begin + simp only [add''], + have ha := ore_eq (s₁ : R) s₂, + set! ra := ore_num (s₁ : R) s₂ with h, rw ←h at *, clear h, -- r tilde + set! sa := ore_denom (s₁ : R) s₂ with h, rw ←h at *, clear h, -- s tilde + rcases ore_condition (sa : R) sb with ⟨rc, sc, hc⟩, -- s*, r* + have : (s₂ : R) * (rb * rc) = s₂ * (ra * sc), + { rw [←mul_assoc, ←hb, mul_assoc, ←hc, ←mul_assoc, ←mul_assoc, ha] }, + rcases ore_left_cancel _ _ s₂ this with ⟨sd, hd⟩, -- s# + symmetry, rw ore_div_eq_iff, + use sc * sd, use rc * sd, + split; simp only [submonoid.coe_mul], + { noncomm_ring, assoc_rw [hd, hc], noncomm_ring }, + { assoc_rewrite [hc], noncomm_ring } +end + +local attribute [instance] ore_localization.ore_eqv + +private def add' (r₂ : R) (s₂ : S) : R[S⁻¹] → R[S⁻¹] := --plus tilde +quotient.lift (λ (r₁s₁ : R × S), add'' r₁s₁.1 r₁s₁.2 r₂ s₂) $ +begin + rintros ⟨r₁', s₁'⟩ ⟨r₁, s₁⟩ ⟨sb, rb, hb, hb'⟩, -- s*, r* + rcases ore_condition (s₁' : R) s₂ with ⟨rc, sc, hc⟩, --s~~, r~~ + rcases ore_condition rb sc with ⟨rd, sd, hd⟩, -- s#, r# + dsimp at *, rw add''_char _ _ _ _ rc sc hc, + have : ↑s₁ * ↑(sb * sd) = ↑s₂ * (rc * rd), + { simp only [submonoid.coe_mul], assoc_rewrite [hb', hd, hc], noncomm_ring }, + rw add''_char _ _ _ _ (rc * rd : R) (sb * sd : S) this, + simp only [submonoid.coe_mul], assoc_rw [hb, hd], + rw [←mul_assoc, ←add_mul, ore_div_eq_iff], + use 1, use rd, split, + { simp }, + { simp only [mul_one, submonoid.coe_one, submonoid.coe_mul] at ⊢ this, assoc_rw [hc, this] }, +end + +private lemma add'_comm (r₁ r₂ : R) (s₁ s₂ : S) : add' r₁ s₁ (r₂ /ₒ s₂) = add' r₂ s₂ (r₁ /ₒ s₁) := +begin + simp only [add', ore_div, add'', quotient.lift_mk, quotient.eq], + have hb := ore_eq ↑s₂ s₁, set rb := ore_num ↑s₂ s₁ with h, -- r~~ + rw ←h, clear h, set sb := ore_denom ↑s₂ s₁ with h, rw ←h, clear h, -- s~~ + have ha := ore_eq ↑s₁ s₂, set ra := ore_num ↑s₁ s₂ with h, -- r~ + rw ←h, clear h, set sa := ore_denom ↑s₁ s₂ with h, rw ←h, clear h, -- s~ + rcases ore_condition ra sb with ⟨rc, sc, hc⟩, -- r#, s# + have : (s₁ : R) * (rb * rc) = s₁ * (sa * sc), + { rw [←mul_assoc, ←hb, mul_assoc, ←hc, ←mul_assoc, ←ha, mul_assoc] }, + rcases ore_left_cancel _ _ s₁ this with ⟨sd, hd⟩, -- s+ + use sc * sd, use rc * sd, + dsimp, split, + { rw [add_mul, add_mul, add_comm], assoc_rw [←hd, hc], noncomm_ring }, + { rw [mul_assoc, ←mul_assoc ↑sa, ←hd, hb], noncomm_ring } +end + +/-- The addition on the Ore localization. -/ +private def add : R[S⁻¹] → R[S⁻¹] → R[S⁻¹] := +λ x, +quotient.lift (λ rs : R × S, add' rs.1 rs.2 x) +begin + rintros ⟨r₁, s₁⟩ ⟨r₂, s₂⟩ hyz, + induction x using ore_localization.ind with r₃ s₃, + dsimp, rw [add'_comm, add'_comm r₂], + simp [(/ₒ), quotient.sound hyz], +end + +instance : has_add R[S⁻¹] := ⟨add⟩ + +lemma ore_div_add_ore_div {r r' : R} {s s' : S} : + r /ₒ s + r' /ₒ s' = (r * ore_denom (s : R) s' + r' * ore_num s s') /ₒ (s * ore_denom s s') := +rfl + +/-- A characterization of the addition on the Ore localizaion, allowing for arbitrary Ore +numerator and Ore denominator. -/ +lemma ore_div_add_char {r r' : R} (s s' : S) (rb : R) (sb : S) + (h : (s : R) * sb = s' * rb) : + r /ₒ s + r' /ₒ s' = (r * sb + r' * rb) /ₒ (s * sb) := +add''_char r s r' s' rb sb h + +/-- Another characterization of the addition on the Ore localization, bundling up all witnesses +and conditions into a sigma type. -/ +def ore_div_add_char' (r r' : R) (s s' : S) : + Σ' r'' : R, Σ' s'' : S, (s : R) * s'' = s' * r'' ∧ + r /ₒ s + r' /ₒs' = (r * s'' + r' * r'') /ₒ (s * s'') := +⟨ore_num s s', ore_denom s s', ore_eq s s', ore_div_add_ore_div⟩ + +@[simp] lemma add_ore_div {r r' : R} {s : S} : (r /ₒ s) + (r' /ₒ s) = (r + r') /ₒ s := +by simp [ore_div_add_char s s 1 1 (by simp)] + +protected lemma add_assoc (x y z : R[S⁻¹]) : + (x + y) + z = x + (y + z) := +begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + induction z using ore_localization.ind with r₃ s₃, + rcases ore_div_add_char' r₁ r₂ s₁ s₂ with ⟨ra, sa, ha, ha'⟩, rw ha', clear ha', + rcases ore_div_add_char' r₂ r₃ s₂ s₃ with ⟨rb, sb, hb, hb'⟩, rw hb', clear hb', + rcases ore_div_add_char' (r₁ * sa + r₂ * ra) r₃ (s₁ * sa) s₃ with ⟨rc, sc, hc, q⟩, rw q, clear q, + rcases ore_div_add_char' r₁ (r₂ * sb + r₃ * rb) s₁ (s₂ * sb) with ⟨rd, sd, hd, q⟩, rw q, clear q, + noncomm_ring, rw add_comm (r₂ * _), + repeat { rw ←add_ore_div }, + congr' 1, + { rcases ore_condition (sd : R) (sa * sc) with ⟨re, se, he⟩, + { simp_rw ←submonoid.coe_mul at hb hc hd, + assoc_rw [subtype.coe_eq_of_eq_mk hc], + rw [←ore_localization.expand, subtype.coe_eq_of_eq_mk hd, ←mul_assoc, + ←ore_localization.expand, subtype.coe_eq_of_eq_mk hb], + apply ore_localization.expand } }, + congr' 1, + { rw [←ore_localization.expand', ←mul_assoc, ←mul_assoc, + ←ore_localization.expand', ←ore_localization.expand'] }, + { simp_rw [←submonoid.coe_mul] at ha hd, + rw [subtype.coe_eq_of_eq_mk hd, ←mul_assoc, ←mul_assoc, + ←mul_assoc, ←ore_localization.expand, ←ore_localization.expand', + subtype.coe_eq_of_eq_mk ha, ←ore_localization.expand], + apply ore_localization.expand' } +end + +private def zero : R[S⁻¹] := 0 /ₒ 1 + +instance : has_zero R[S⁻¹] := ⟨zero⟩ + +protected lemma zero_def : (0 : R[S⁻¹]) = 0 /ₒ 1 := rfl + +@[simp] +lemma zero_div_eq_zero (s : S) : 0 /ₒ s = 0 := +by { rw [ore_localization.zero_def, ore_div_eq_iff], exact ⟨s, 1, by simp⟩ } + +protected lemma zero_add (x : R[S⁻¹]) : 0 + x = x := +begin + induction x using ore_localization.ind, + rw [←zero_div_eq_zero, add_ore_div], simp +end + +protected lemma add_comm (x y : R[S⁻¹]) : x + y = y + x := +begin + induction x using ore_localization.ind, + induction y using ore_localization.ind, + change add' _ _ (_ /ₒ _) = _, apply add'_comm +end + +instance : add_comm_monoid R[S⁻¹] := +{ add_comm := ore_localization.add_comm, + add_assoc := ore_localization.add_assoc, + zero := zero, + zero_add := ore_localization.zero_add, + add_zero := λ x, by rw [ore_localization.add_comm, ore_localization.zero_add], + .. ore_localization.has_add } + +protected lemma zero_mul (x : R[S⁻¹]) : 0 * x = 0 := +begin + induction x using ore_localization.ind with r s, + rw [ore_localization.zero_def, ore_div_mul_char 0 r 1 s r 1 (by simp)], simp +end + +protected lemma mul_zero (x : R[S⁻¹]) : x * 0 = 0 := +begin + induction x using ore_localization.ind with r s, + rw [ore_localization.zero_def, ore_div_mul_char r 0 s 1 0 1 (by simp)], simp +end + +protected lemma left_distrib (x y z : R[S⁻¹]) : x * (y + z) = x * y + x * z := +begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + induction z using ore_localization.ind with r₃ s₃, + rcases ore_div_add_char' r₂ r₃ s₂ s₃ with ⟨ra, sa, ha, q⟩, rw q, clear q, + rw ore_localization.expand' r₂ s₂ sa, + rcases ore_div_mul_char' r₁ (r₂ * sa) s₁ (s₂ * sa) with ⟨rb, sb, hb, q⟩, rw q, clear q, + have hs₃rasb : ↑s₃ * (ra * sb) ∈ S, { rw [←mul_assoc, ←ha], norm_cast, apply set_like.coe_mem }, + rw ore_localization.expand _ _ _ hs₃rasb, + have ha' : (↑(s₂ * sa * sb)) = ↑s₃ * (ra * sb), { simp [ha, ←mul_assoc] }, + rw ←subtype.coe_eq_of_eq_mk ha', + rcases ore_div_mul_char' r₁ (r₃ * (ra * sb)) s₁ (s₂ * sa * sb) with ⟨rc, sc, hc, hc'⟩, rw hc', + rw ore_div_add_char (s₂ * sa * sb) (s₂ * sa * sb * sc) 1 sc (by simp), + rw ore_localization.expand' (r₂ * ↑sa + r₃ * ra) (s₂ * sa) (sb * sc), + conv_lhs { congr, skip, congr, + rw [add_mul, S.coe_mul, ←mul_assoc, hb, ←mul_assoc, mul_assoc r₃, hc, mul_assoc, ←mul_add] }, + rw ore_localization.mul_cancel', simp only [mul_one, submonoid.coe_mul, mul_add, ←mul_assoc], +end + +lemma right_distrib (x y z : R[S⁻¹]) : (x + y) * z = x * z + y * z := +begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + induction z using ore_localization.ind with r₃ s₃, + rcases ore_div_add_char' r₁ r₂ s₁ s₂ with ⟨ra, sa, ha, ha'⟩, rw ha', clear ha', norm_cast at ha, + rw ore_localization.expand' r₁ s₁ sa, + rw ore_localization.expand r₂ s₂ ra (by rw ←ha; apply set_like.coe_mem), + rw ←subtype.coe_eq_of_eq_mk ha, + repeat { rw ore_div_mul_ore_div }, simp only [add_mul, add_ore_div] +end + +instance : semiring R[S⁻¹] := +{ zero_mul := ore_localization.zero_mul, + mul_zero := ore_localization.mul_zero, + left_distrib := ore_localization.left_distrib, + right_distrib := right_distrib, + .. ore_localization.add_comm_monoid, + .. ore_localization.monoid } + +section UMP +variables {T : Type*} [semiring T] +variables (f : R →+* T) (fS : S →* units T) +variables (hf : ∀ (s : S), f s = fS s) + +include f fS hf + +/-- The universal lift from a ring homomorphism `f : R →+* T`, which maps elements in `S` to +units of `T`, to a ring homomorphism `R[S⁻¹] →+* T`. This extends the construction on +monoids. -/ +def universal_hom : R[S⁻¹] →+* T := +{ map_zero' := + begin + rw [monoid_hom.to_fun_eq_coe, ore_localization.zero_def, universal_mul_hom_apply], + simp + end, + map_add' := λ x y, + begin + induction x using ore_localization.ind with r₁ s₁, + induction y using ore_localization.ind with r₂ s₂, + rcases ore_div_add_char' r₁ r₂ s₁ s₂ with ⟨r₃, s₃, h₃, h₃'⟩, rw h₃', clear h₃', + simp only [universal_mul_hom_apply, ring_hom.coe_monoid_hom, + ring_hom.to_monoid_hom_eq_coe, monoid_hom.to_fun_eq_coe], + simp only [mul_inv_rev, monoid_hom.map_mul, ring_hom.map_add, ring_hom.map_mul, units.coe_mul], + rw [add_mul, ←mul_assoc, mul_assoc (f r₁), hf, ←units.coe_mul], + simp only [mul_one, mul_right_inv, units.coe_one], + congr' 1, rw [mul_assoc], congr' 1, + norm_cast at h₃, have h₃' := subtype.coe_eq_of_eq_mk h₃, + rw [←units.coe_mul, ←mul_inv_rev, ←fS.map_mul, h₃'], + have hs₂r₃ : ↑s₂ * r₃ ∈ S, { rw ←h₃, exact set_like.coe_mem (s₁ * s₃)}, + apply (units.inv_mul_cancel_left (fS s₂) _).symm.trans, + conv_lhs { congr, skip, + rw [←units.mul_inv_cancel_left (fS ⟨s₂ * r₃, hs₂r₃⟩) (fS s₂), mul_assoc, mul_assoc], + congr, skip, rw [←hf, ←mul_assoc (f s₂), ←f.map_mul], + conv { congr, skip, congr, rw [←h₃] }, + rw [hf, ←mul_assoc, ←h₃', units.inv_mul] }, + rw [one_mul, ←h₃', units.mul_inv, mul_one], + end, + .. universal_mul_hom f.to_monoid_hom fS hf } + +lemma universal_hom_apply {r : R} {s : S} : + universal_hom f fS hf (r /ₒ s) = (f r) * ((fS s)⁻¹ : units T) := rfl + +lemma universal_hom_commutes {r : R} : universal_hom f fS hf (numerator_hom r) = f r := +by simp [numerator_hom_apply, universal_hom_apply] + +lemma universal_hom_unique (φ : R[S⁻¹] →+* T) (huniv : ∀ (r : R), φ (numerator_hom r) = f r) : + φ = universal_hom f fS hf := +ring_hom.coe_monoid_hom_injective $ +universal_mul_hom_unique (ring_hom.to_monoid_hom f) fS hf ↑φ huniv + +end UMP + +end semiring + +section ring +variables {R : Type*} [ring R] {S : submonoid R} [ore_set S] + +/-- Negation on the Ore localization is defined via negation on the numerator. -/ +protected def neg : R[S⁻¹] → R[S⁻¹] := +lift_expand (λ (r : R) (s : S), (- r) /ₒ s) $ + λ r t s ht, by rw [neg_mul_eq_neg_mul, ←ore_localization.expand] + +instance : has_neg R[S⁻¹] := ⟨ore_localization.neg⟩ + +@[simp] protected lemma neg_def (r : R) (s : S) : - (r /ₒ s) = (- r) /ₒ s := rfl + +protected lemma add_left_neg (x : R[S⁻¹]) : (- x) + x = 0 := +by induction x using ore_localization.ind with r s; simp + +instance : ring R[S⁻¹] := +{ add_left_neg := ore_localization.add_left_neg, + .. ore_localization.semiring, + .. ore_localization.has_neg } + +open_locale non_zero_divisors + +lemma numerator_hom_inj (hS : S ≤ R⁰) : function.injective (numerator_hom : R → R[S⁻¹]) := +λ r₁ r₂ h, +begin + rw [numerator_hom_apply, numerator_hom_apply, ore_div_eq_iff] at h, + rcases h with ⟨u, v, h₁, h₂⟩, + simp only [S.coe_one, one_mul] at h₂, + rwa [←h₂, mul_cancel_right_mem_non_zero_divisor (hS (set_like.coe_mem u)), eq_comm] at h₁, +end + +lemma nontrivial_of_non_zero_divisors [nontrivial R] (hS : S ≤ R⁰) : nontrivial R[S⁻¹] := +⟨⟨0, 1, λ h, + begin + rw [ore_localization.one_def, ore_localization.zero_def] at h, + apply non_zero_divisors.coe_ne_zero 1 (numerator_hom_inj hS h).symm + end⟩⟩ + +end ring + +section division_ring + +open_locale non_zero_divisors +open_locale classical + +variables {R : Type*} [ring R] [nontrivial R] [ore_set R⁰] + +instance : nontrivial R[R⁰⁻¹] := nontrivial_of_non_zero_divisors (refl R⁰) + +variables [no_zero_divisors R] + +noncomputable theory + +/-- The inversion of Ore fractions for a ring without zero divisors, satisying `0⁻¹ = 0` and +`(r /ₒ r')⁻¹ = r' /ₒ r` for `r ≠ 0`. -/ +protected def inv : R[R⁰⁻¹] → R[R⁰⁻¹] := +lift_expand (λ r s, if hr: r = (0 : R) then (0 : R[R⁰⁻¹]) + else (s /ₒ ⟨r, λ _, eq_zero_of_ne_zero_of_mul_right_eq_zero hr⟩)) +begin + intros r t s hst, + by_cases hr : r = 0, + { simp [hr] }, + { by_cases ht : t = 0, + { exfalso, apply non_zero_divisors.coe_ne_zero ⟨_, hst⟩, simp [ht, mul_zero] }, + { simp only [hr, ht, set_like.coe_mk, dif_neg, not_false_iff, or_self, mul_eq_zero], + apply ore_localization.expand } } +end + +instance : has_inv R[R⁰⁻¹] := ⟨ore_localization.inv⟩ + +protected lemma inv_def {r : R} {s : R⁰} : + (r /ₒ s)⁻¹ = if hr: r = (0 : R) then (0 : R[R⁰⁻¹]) + else (s /ₒ ⟨r, λ _, eq_zero_of_ne_zero_of_mul_right_eq_zero hr⟩) := rfl + +protected lemma mul_inv_cancel (x : R[R⁰⁻¹]) (h : x ≠ 0) : x * x⁻¹ = 1 := +begin + induction x using ore_localization.ind with r s, + rw [ore_localization.inv_def, ore_localization.one_def], + by_cases hr : r = 0, + { exfalso, apply h, simp [hr] }, + { simp [hr], apply ore_localization.div_eq_one' } +end + +protected lemma inv_zero : (0 : R[R⁰⁻¹])⁻¹ = 0 := +by { rw [ore_localization.zero_def, ore_localization.inv_def], simp } + +instance : division_ring R[(R⁰)⁻¹] := +{ mul_inv_cancel := ore_localization.mul_inv_cancel, + inv_zero := ore_localization.inv_zero, + .. ore_localization.nontrivial, + .. ore_localization.has_inv, + .. ore_localization.ring } + +end division_ring + +end ore_localization diff --git a/src/ring_theory/ore_localization/ore_set.lean b/src/ring_theory/ore_localization/ore_set.lean new file mode 100644 index 0000000000000..10fecbb77e85c --- /dev/null +++ b/src/ring_theory/ore_localization/ore_set.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2022 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer, Kevin Klinge +-/ +import group_theory.subgroup.basic + +/-! + +# (Right) Ore sets + +This defines right Ore sets on arbitrary monoids. + +## References + +* https://ncatlab.org/nlab/show/Ore+set + +-/ + +namespace ore_localization + +section monoid + +/-- A submonoid `S` of a monoid `R` is (right) Ore if common factors on the left can be turned +into common factors on the right, and if each pair of `r : R` and `s : S` admits an Ore numerator +`v : R` and an Ore denominator `u : S` such that `r * u = s * v`. -/ +class ore_set {R : Type*} [monoid R] (S : submonoid R) := +(ore_left_cancel : ∀ (r₁ r₂ : R) (s : S), ↑s * r₁ = s * r₂ → ∃ s' : S, r₁ * s' = r₂ * s') +(ore_num : R → S → R) +(ore_denom : R → S → S) +(ore_eq : ∀ (r : R) (s : S), r * ore_denom r s = s * ore_num r s) + +variables {R : Type*} [monoid R] {S : submonoid R} [ore_set S] + +/-- Common factors on the left can be turned into common factors on the right, a weak form of +cancellability. -/ +lemma ore_left_cancel (r₁ r₂ : R) (s : S) (h : ↑s * r₁ = s * r₂) : ∃ s' : S, r₁ * s' = r₂ * s' := +ore_set.ore_left_cancel r₁ r₂ s h + +/-- The Ore numerator of a fraction. -/ +def ore_num (r : R) (s : S) : R := ore_set.ore_num r s + +/-- The Ore denominator of a fraction. -/ +def ore_denom (r : R) (s : S) : S := ore_set.ore_denom r s + +/-- The Ore condition of a fraction, expressed in terms of `ore_num` and `ore_denom`. -/ +lemma ore_eq (r : R) (s : S) : r * (ore_denom r s) = s * (ore_num r s) := ore_set.ore_eq r s + +/-- The Ore condition bundled in a sigma type. This is useful in situations where we want to obtain +both witnesses and the condition for a given fraction. -/ +def ore_condition (r : R) (s : S) : Σ' r' : R, Σ' s' : S, r * s' = s * r' := +⟨ore_num r s, ore_denom r s, ore_eq r s⟩ + +/-- The trivial submonoid is an Ore set. -/ +instance ore_set_bot : ore_set (⊥ : submonoid R) := +{ ore_left_cancel := λ _ _ s h, + ⟨s, begin + rcases s with ⟨s, hs⟩, + rw submonoid.mem_bot at hs, + subst hs, + rw [set_like.coe_mk, one_mul, one_mul] at h, + subst h + end⟩, + ore_num := λ r _, r, + ore_denom := λ _ s, s, + ore_eq := λ _ s, by { rcases s with ⟨s, hs⟩, rw submonoid.mem_bot at hs, simp [hs] } } + +/-- Every submonoid of a commutative monoid is an Ore set. -/ +@[priority 100] +instance ore_set_comm {R} [comm_monoid R] (S : submonoid R) : ore_set S := +{ ore_left_cancel := λ m n s h, ⟨s, by rw [mul_comm n s, mul_comm m s, h]⟩, + ore_num := λ r _, r, + ore_denom := λ _ s, s, + ore_eq := λ r s, by rw mul_comm } + +end monoid + +/-- Cancellability in monoids with zeros can act as a replacement for the `ore_left_cancel` +condition of an ore set. -/ +def ore_set_of_cancel_monoid_with_zero + {R : Type*} [cancel_monoid_with_zero R] {S : submonoid R} + (ore_num : R → S → R) (ore_denom : R → S → S) + (ore_eq : ∀ (r : R) (s : S), r * (ore_denom r s) = s * (ore_num r s)) : + ore_set S := +{ ore_left_cancel := λ r₁ r₂ s h, ⟨s, mul_eq_mul_right_iff.mpr (mul_eq_mul_left_iff.mp h)⟩, + ore_num := ore_num, + ore_denom := ore_denom, + ore_eq := ore_eq } + +/-- In rings without zero divisors, the first (cancellability) condition is always fulfilled, +it suffices to give a proof for the Ore condition itself. -/ +def ore_set_of_no_zero_divisors + {R : Type*} [ring R] [no_zero_divisors R] {S : submonoid R} + (ore_num : R → S → R) (ore_denom : R → S → S) + (ore_eq : ∀ (r : R) (s : S), r * (ore_denom r s) = s * (ore_num r s)) : + ore_set S := +begin + letI : cancel_monoid_with_zero R := no_zero_divisors.to_cancel_monoid_with_zero, + exact ore_set_of_cancel_monoid_with_zero ore_num ore_denom ore_eq +end + +end ore_localization From 36473a4790bc220b0a3aab02048922cbc6feb73f Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Jul 2022 07:39:57 +0000 Subject: [PATCH 087/173] feat(category_theory): full_of_surjective (#14974) --- .../functor/fully_faithful.lean | 35 ++++++++++++++----- 1 file changed, 26 insertions(+), 9 deletions(-) diff --git a/src/category_theory/functor/fully_faithful.lean b/src/category_theory/functor/fully_faithful.lean index 30b08d29c6ada..91635c9022566 100644 --- a/src/category_theory/functor/fully_faithful.lean +++ b/src/category_theory/functor/fully_faithful.lean @@ -11,13 +11,18 @@ import logic.equiv.basic We define typeclasses `full` and `faithful`, decorating functors. -Use `F.map_injective` to retrieve the fact that `F.map` is injective when `[faithful F]`, -and `F.preimage` to obtain preimages of morphisms when `[full F]`. - -We prove some basic "cancellation" lemmas for full and/or faithful functors. - -See `category_theory.equivalence` for the fact that a functor is an equivalence if and only if -it is fully faithful and essentially surjective. +## Main definitions and results +* Use `F.map_injective` to retrieve the fact that `F.map` is injective when `[faithful F]`. +* Similarly, `F.map_surjective` states that `F.map` is surjective when `[full F]`. +* Use `F.preimage` to obtain preimages of morphisms when `[full F]`. +* We prove some basic "cancellation" lemmas for full and/or faithful functors, as well as a + construction for "dividing" a functor by a faithful functor, see `faithful.div`. +* `full F` carries data, so definitional properties of the preimage can be used when using + `F.preimage`. To obtain an instance of `full F` non-constructively, you can use `full_of_exists` + and `full_of_surjective`. + +See `category_theory.equivalence.of_fully_faithful_ess_surj` for the fact that a functor is an +equivalence if and only if it is fully faithful and essentially surjective. -/ @@ -55,8 +60,7 @@ restate_axiom faithful.map_injective' namespace functor variables {X Y : C} -lemma map_injective (F : C ⥤ D) [faithful F] : - function.injective $ @functor.map _ _ _ _ F X Y := +lemma map_injective (F : C ⥤ D) [faithful F] : function.injective $ @functor.map _ _ _ _ F X Y := faithful.map_injective F lemma map_iso_injective (F : C ⥤ D) [faithful F] : @@ -70,6 +74,19 @@ full.preimage.{v₁ v₂} f F.map (preimage F f) = f := by unfold preimage; obviously +lemma map_surjective (F : C ⥤ D) [full F] : function.surjective (@functor.map _ _ _ _ F X Y) := +λ f, ⟨F.preimage f, F.image_preimage f⟩ + +/-- Deduce that `F` is full from the existence of preimages, using choice. -/ +noncomputable def full_of_exists (F : C ⥤ D) + (h : ∀ (X Y : C) (f : F.obj X ⟶ F.obj Y), ∃ p, F.map p = f) : full F := +by { choose p hp using h, exact ⟨p, hp⟩ } + +/-- Deduce that `F` is full from surjectivity of `F.map`, using choice. -/ +noncomputable def full_of_surjective (F : C ⥤ D) + (h : ∀ (X Y : C), function.surjective (@functor.map _ _ _ _ F X Y)) : full F := +full_of_exists _ h + end functor section From 7e85d9026e06efca7eac7d0396fdd3afdfea45e0 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 11:14:36 +0000 Subject: [PATCH 088/173] feat(ring_theory/ring_hom/finite): Finite type is a local property (#15379) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebra/algebra/subalgebra/basic.lean | 35 +++++++++ src/ring_theory/adjoin/basic.lean | 27 +++++-- src/ring_theory/local_properties.lean | 93 ++++++++++++++++------- src/ring_theory/ring_hom/finite_type.lean | 92 ++++++++++++++++++++++ 4 files changed, 213 insertions(+), 34 deletions(-) create mode 100644 src/ring_theory/ring_hom/finite_type.lean diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index 542838a9a4a4d..e37628792f246 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -5,6 +5,8 @@ Authors: Kenny Lau, Yury Kudryashov -/ import algebra.algebra.basic import data.set.Union_lift +import linear_algebra.finsupp +import ring_theory.ideal.operations /-! # Subalgebras over Commutative Semiring @@ -1064,6 +1066,39 @@ set_like.ext' (set.centralizer_univ A) end centralizer +/-- Suppose we are given `∑ i, lᵢ * sᵢ = 1` in `S`, and `S'` a subalgebra of `S` that contains +`lᵢ` and `sᵢ`. To check that an `x : S` falls in `S'`, we only need to show that +`r ^ n • x ∈ M'` for some `n` for each `r : s`. -/ +lemma mem_of_span_eq_top_of_smul_pow_mem {S : Type*} [comm_ring S] [algebra R S] + (S' : subalgebra R S) (s : set S) (l : s →₀ S) (hs : finsupp.total s S S coe l = 1) + (hs' : s ⊆ S') (hl : ∀ i, l i ∈ S') (x : S) + (H : ∀ r : s, ∃ (n : ℕ), (r ^ n : S) • x ∈ S') : x ∈ S' := +begin + let s' : set S' := coe ⁻¹' s, + let e : s' ≃ s := ⟨λ x, ⟨x.1, x.2⟩, λ x, ⟨⟨_, hs' x.2⟩, x.2⟩, λ ⟨⟨_, _⟩, _⟩, rfl, λ ⟨_, _⟩, rfl⟩, + let l' : s →₀ S' := ⟨l.support, λ x, ⟨_, hl x⟩, + λ _, finsupp.mem_support_iff.trans $ iff.not $ by { rw ← subtype.coe_inj, refl }⟩, + have : ideal.span s' = ⊤, + { rw [ideal.eq_top_iff_one, ideal.span, finsupp.mem_span_iff_total], + refine ⟨finsupp.equiv_map_domain e.symm l', subtype.ext $ eq.trans _ hs⟩, + rw finsupp.total_equiv_map_domain, + exact finsupp.apply_total _ (algebra.of_id S' S).to_linear_map _ _ }, + obtain ⟨s'', hs₁, hs₂⟩ := (ideal.span_eq_top_iff_finite _).mp this, + replace H : ∀ r : s'', ∃ (n : ℕ), (r ^ n : S) • x ∈ S' := λ r, H ⟨r, hs₁ r.2⟩, + choose n₁ n₂ using H, + let N := s''.attach.sup n₁, + have hs' := ideal.span_pow_eq_top _ hs₂ N, + have : ∀ {x : S}, x ∈ (algebra.of_id S' S).range.to_submodule ↔ x ∈ S' := + λ x, ⟨by { rintro ⟨x, rfl⟩, exact x.2 }, λ h, ⟨⟨x, h⟩, rfl⟩⟩, + rw ← this, + apply (algebra.of_id S' S).range.to_submodule.mem_of_span_top_of_smul_mem _ hs', + rintro ⟨_, r, hr, rfl⟩, + convert submodule.smul_mem _ (r ^ (N - n₁ ⟨r, hr⟩)) (this.mpr $ n₂ ⟨r, hr⟩) using 1, + simp only [_root_.coe_coe, subtype.coe_mk, + subalgebra.smul_def, smul_smul, ← pow_add, subalgebra.coe_pow], + rw tsub_add_cancel_of_le (finset.le_sup (s''.mem_attach _) : n₁ ⟨r, hr⟩ ≤ N), +end + end subalgebra section nat diff --git a/src/ring_theory/adjoin/basic.lean b/src/ring_theory/adjoin/basic.lean index 51357c46d4265..63a64bcbc86e7 100644 --- a/src/ring_theory/adjoin/basic.lean +++ b/src/ring_theory/adjoin/basic.lean @@ -270,8 +270,13 @@ begin congr' 1 with z, simp [submonoid.closure_union, submonoid.mem_sup, set.mem_mul] end -lemma pow_smul_mem_adjoin_smul (r : R) (s : set A) {x : A} (hx : x ∈ adjoin R s) : - ∃ n₀ : ℕ, ∀ n ≥ n₀, r ^ n • x ∈ adjoin R (r • s) := +variable {R} + +lemma pow_smul_mem_of_smul_subset_of_mem_adjoin + [comm_semiring B] [algebra R B] [algebra A B] [is_scalar_tower R A B] + (r : A) (s : set B) (B' : subalgebra R B) (hs : r • s ⊆ B') {x : B} (hx : x ∈ adjoin R s) + (hr : algebra_map A B r ∈ B') : + ∃ n₀ : ℕ, ∀ n ≥ n₀, r ^ n • x ∈ B' := begin change x ∈ (adjoin R s).to_submodule at hx, rw [adjoin_eq_span, finsupp.mem_span_iff_total] at hx, @@ -280,16 +285,24 @@ begin use l.support.sup n₁, intros n hn, rw finsupp.smul_sum, - refine (adjoin R (r • s)).to_submodule.sum_mem _, + refine B'.to_submodule.sum_mem _, intros a ha, have : n ≥ n₁ a := le_trans (finset.le_sup ha) hn, dsimp only, - rw [← tsub_add_cancel_of_le this, pow_add, ← smul_smul, smul_smul _ (l a), mul_comm, - ← smul_smul, adjoin_eq_span], - refine submodule.smul_mem _ _ _, - exact submodule.smul_mem _ _ (submodule.subset_span (n₂ a)) + rw [← tsub_add_cancel_of_le this, pow_add, ← smul_smul, + ← is_scalar_tower.algebra_map_smul A (l a) (a : B), smul_smul (r ^ n₁ a), + mul_comm, ← smul_smul, smul_def, map_pow, is_scalar_tower.algebra_map_smul], + apply subalgebra.mul_mem _ (subalgebra.pow_mem _ hr _) _, + refine subalgebra.smul_mem _ _ _, + change _ ∈ B'.to_submonoid, + rw ← submonoid.closure_eq B'.to_submonoid, + apply submonoid.closure_mono hs (n₂ a), end +lemma pow_smul_mem_adjoin_smul (r : R) (s : set A) {x : A} (hx : x ∈ adjoin R s) : + ∃ n₀ : ℕ, ∀ n ≥ n₀, r ^ n • x ∈ adjoin R (r • s) := +pow_smul_mem_of_smul_subset_of_mem_adjoin r s _ subset_adjoin hx (subalgebra.algebra_map_mem _ _) + end comm_semiring section ring diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index a76ad1b435a1e..70f52e77a63c5 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -111,9 +111,24 @@ def ring_hom.holds_for_localization_away : Prop := ∀ ⦃R : Type u⦄ (S : Type u) [comm_ring R] [comm_ring S] [by exactI algebra R S] (r : R) [by exactI is_localization.away r S], by exactI P (algebra_map R S) -/-- A property `P` of ring homs satisfies `ring_hom.of_localization_span_target` +/-- A property `P` of ring homs satisfies `ring_hom.of_localization_finite_span_target` if `P` holds for `R →+* S` whenever there exists a finite set `{ r }` that spans `S` such that -`P` holds for `R →+* Sᵣ`. -/ +`P` holds for `R →+* Sᵣ`. + +Note that this is equivalent to `ring_hom.of_localization_span_target` via +`ring_hom.of_localization_span_target_iff_finite`, but this is easier to prove. -/ +def ring_hom.of_localization_finite_span_target : Prop := +∀ ⦃R S : Type u⦄ [comm_ring R] [comm_ring S] (f : by exactI R →+* S) + (s : finset S) (hs : by exactI ideal.span (s : set S) = ⊤) + (H : by exactI (∀ (r : s), P ((algebra_map S (localization.away (r : S))).comp f))), + by exactI P f + +/-- A property `P` of ring homs satisfies `ring_hom.of_localization_span_target` +if `P` holds for `R →+* S` whenever there exists a set `{ r }` that spans `S` such that +`P` holds for `R →+* Sᵣ`. + +Note that this is equivalent to `ring_hom.of_localization_finite_span_target` via +`ring_hom.of_localization_span_target_iff_finite`, but this has less restrictions when applying. -/ def ring_hom.of_localization_span_target : Prop := ∀ ⦃R S : Type u⦄ [comm_ring R] [comm_ring S] (f : by exactI R →+* S) (s : set S) (hs : by exactI ideal.span s = ⊤) @@ -148,6 +163,19 @@ begin exact h s' h₂ (λ x, hs' ⟨_, h₁ x.prop⟩) } end +lemma ring_hom.of_localization_span_target_iff_finite : + ring_hom.of_localization_span_target @P ↔ ring_hom.of_localization_finite_span_target @P := +begin + delta ring_hom.of_localization_span_target ring_hom.of_localization_finite_span_target, + apply forall₅_congr, -- TODO: Using `refine` here breaks `resetI`. + introsI, + split, + { intros h s, exact h s }, + { intros h s hs hs', + obtain ⟨s', h₁, h₂⟩ := (ideal.span_eq_top_iff_finite s).mp hs, + exact h s' h₂ (λ x, hs' ⟨_, h₁ x.prop⟩) } +end + variables {P f R' S'} lemma _root_.ring_hom.property_is_local.respects_iso (hP : ring_hom.property_is_local @P) : @@ -543,6 +571,37 @@ lemma localization_away_map_finite_type (r : R) [is_localization.away r R'] (is_localization.away.map R' S' f r).finite_type := localization_finite_type.away r hf +variable {S'} + +/-- +Let `S` be an `R`-algebra, `M` a submonoid of `S`, `S' = M⁻¹S`. +Suppose the image of some `x : S` falls in the adjoin of some finite `s ⊆ S'` over `R`, +and `A` is an `R`-subalgebra of `S` containing both `M` and the numerators of `s`. +Then, there exists some `m : M` such that `m • x` falls in `A`. +-/ +lemma is_localization.exists_smul_mem_of_mem_adjoin [algebra R S] + [algebra R S'] [is_scalar_tower R S S'] (M : submonoid S) + [is_localization M S'] (x : S) (s : finset S') (A : subalgebra R S) + (hA₁ : (is_localization.finset_integer_multiple M s : set S) ⊆ A) + (hA₂ : M ≤ A.to_submonoid) + (hx : algebra_map S S' x ∈ algebra.adjoin R (s : set S')) : + ∃ m : M, m • x ∈ A := +begin + let g : S →ₐ[R] S' := is_scalar_tower.to_alg_hom R S S', + let y := is_localization.common_denom_of_finset M s, + have hx₁ : (y : S) • ↑s = g '' _ := (is_localization.finset_integer_multiple_image _ s).symm, + obtain ⟨n, hn⟩ := algebra.pow_smul_mem_of_smul_subset_of_mem_adjoin (y : S) (s : set S') + (A.map g) (by { rw hx₁, exact set.image_subset _ hA₁ }) hx (set.mem_image_of_mem _ (hA₂ y.2)), + obtain ⟨x', hx', hx''⟩ := hn n (le_of_eq rfl), + 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, +end + /-- Let `S` be an `R`-algebra, `M` an submonoid of `R`, and `S' = M⁻¹S`. If the image of some `x : S` falls in the adjoin of some finite `s ⊆ S'` over `R`, @@ -556,30 +615,10 @@ lemma is_localization.lift_mem_adjoin_finset_integer_multiple [algebra R S] ∃ m : M, m • x ∈ algebra.adjoin R (is_localization.finset_integer_multiple (M.map (algebra_map R S : R →* S)) s : set S) := begin - -- mirrors the proof of `is_localization.smul_mem_finset_integer_multiple_span` - let g : S →ₐ[R] S' := alg_hom.mk' (algebra_map S S') - (λ c x, by simp [algebra.algebra_map_eq_smul_one]), - - let y := is_localization.common_denom_of_finset (M.map (algebra_map R S : R →* S)) s, - have hx₁ : (y : S) • ↑s = g '' _ := (is_localization.finset_integer_multiple_image _ s).symm, - obtain ⟨y', hy', e : algebra_map R S y' = y⟩ := y.prop, - have : algebra_map R S y' • (s : set S') = y' • s := - by simp_rw [algebra.algebra_map_eq_smul_one, smul_assoc, one_smul], - rw [← e, this] at hx₁, - replace hx₁ := congr_arg (algebra.adjoin R) hx₁, - obtain ⟨n, hn⟩ := algebra.pow_smul_mem_adjoin_smul _ y' (s : set S') hx, - specialize hn n (le_of_eq rfl), - erw [hx₁, ← g.map_smul, ← g.map_adjoin] at hn, - obtain ⟨x', hx', hx''⟩ := hn, - obtain ⟨⟨_, a, ha₁, rfl⟩, ha₂⟩ := (is_localization.eq_iff_exists - (M.map (algebra_map R S : R →* S)) S').mp hx'', - use (⟨a, ha₁⟩ : M) * (⟨y', hy'⟩ : M) ^ n, - convert (algebra.adjoin R (is_localization.finset_integer_multiple - (submonoid.map (algebra_map R S : R →* S) M) s : set S)).smul_mem hx' a using 1, - convert ha₂.symm, - { rw [mul_comm (y' ^ n • x), subtype.coe_mk, submonoid.smul_def, submonoid.coe_mul, ← smul_smul, - algebra.smul_def, submonoid_class.coe_pow], refl }, - { rw mul_comm, exact algebra.smul_def _ _ } + obtain ⟨⟨_, a, ha, rfl⟩, e⟩ := is_localization.exists_smul_mem_of_mem_adjoin + (M.map (algebra_map R S : R →* S)) x s (algebra.adjoin R _) algebra.subset_adjoin _ hx, + { exact ⟨⟨a, ha⟩, by simpa [submonoid.smul_def] using e⟩ }, +{ rintros _ ⟨a, ha, rfl⟩, exact subalgebra.algebra_map_mem _ a } end lemma finite_type_of_localization_span : ring_hom.of_localization_span @ring_hom.finite_type := @@ -613,7 +652,7 @@ begin rw [submonoid.smul_def, algebra.smul_def, is_scalar_tower.algebra_map_apply R S, subtype.coe_mk, ← map_mul] at hn₁, obtain ⟨⟨_, n₂, rfl⟩, hn₂⟩ := is_localization.lift_mem_adjoin_finset_integer_multiple - (submonoid.powers (r : R)) (localization.away (f r)) _ (s₁ r) hn₁, + (submonoid.powers (r : R)) _ (s₁ r) hn₁, rw [submonoid.smul_def, ← algebra.smul_def, smul_smul, subtype.coe_mk, ← pow_add] at hn₂, use n₂ + n₁, refine le_supr (λ (x : s), algebra.adjoin R (sf x : set S)) r _, diff --git a/src/ring_theory/ring_hom/finite_type.lean b/src/ring_theory/ring_hom/finite_type.lean new file mode 100644 index 0000000000000..ecce446a133de --- /dev/null +++ b/src/ring_theory/ring_hom/finite_type.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2021 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import ring_theory.local_properties +import ring_theory.localization.inv_submonoid + +/-! + +# The meta properties of finite-type ring homomorphisms. + +The main result is `ring_hom.finite_is_local`. + +-/ + +namespace ring_hom + +open_locale pointwise + +lemma finite_type_stable_under_composition : + stable_under_composition @finite_type := +by { introv R hf hg, exactI hg.comp hf } + +lemma finite_type_holds_for_localization_away : + holds_for_localization_away @finite_type := +begin + introv R _, + resetI, + suffices : algebra.finite_type R S, + { change algebra.finite_type _ _, convert this, ext, rw algebra.smul_def, refl }, + exact is_localization.finite_type_of_monoid_fg (submonoid.powers r) S, +end + +lemma finite_type_of_localization_span_target : of_localization_span_target @finite_type := +begin + -- Setup algebra intances. + rw of_localization_span_target_iff_finite, + introv R hs H, + resetI, + classical, + letI := f.to_algebra, + replace H : ∀ r : s, algebra.finite_type R (localization.away (r : S)), + { intro r, convert H r, ext, rw algebra.smul_def, refl }, + replace H := λ r, (H r).1, + constructor, + -- Suppose `s : finset S` spans `S`, and each `Sᵣ` is finitely generated as an `R`-algebra. + -- Say `t r : finset Sᵣ` generates `Sᵣ`. By assumption, we may find `lᵢ` such that + -- `∑ lᵢ * sᵢ = 1`. I claim that all `s` and `l` and the numerators of `t` and generates `S`. + choose t ht using H, + obtain ⟨l, hl⟩ := (finsupp.mem_span_iff_total S (s : set S) 1).mp + (show (1 : S) ∈ ideal.span (s : set S), by { rw hs, trivial }), + let sf := λ (x : s), is_localization.finset_integer_multiple (submonoid.powers (x : S)) (t x), + use s.attach.bUnion sf ∪ s ∪ l.support.image l, + rw eq_top_iff, + -- We need to show that every `x` falls in the subalgebra generated by those elements. + -- Since all `s` and `l` are in the subalgebra, it suffices to check that `sᵢ ^ nᵢ • x` falls in + -- the algebra for each `sᵢ` and some `nᵢ`. + rintro x -, + apply subalgebra.mem_of_span_eq_top_of_smul_pow_mem _ (s : set S) l hl _ _ x _, + { intros x hx, + apply algebra.subset_adjoin, + rw [finset.coe_union, finset.coe_union], + exact or.inl (or.inr hx) }, + { intros i, + by_cases h : l i = 0, { rw h, exact zero_mem _ }, + apply algebra.subset_adjoin, + rw [finset.coe_union, finset.coe_image], + exact or.inr (set.mem_image_of_mem _ (finsupp.mem_support_iff.mpr h)) }, + { intro r, + rw [finset.coe_union, finset.coe_union, finset.coe_bUnion], + -- Since all `sᵢ` and numerators of `t r` are in the algebra, it suffices to show that the + -- image of `x` in `Sᵣ` falls in the `R`-adjoin of `t r`, which is of course true. + obtain ⟨⟨_, n₂, rfl⟩, hn₂⟩ := is_localization.exists_smul_mem_of_mem_adjoin + (submonoid.powers (r : S)) x (t r) + (algebra.adjoin R _) _ _ _, + { exact ⟨n₂, hn₂⟩ }, + { intros x hx, + apply algebra.subset_adjoin, + refine or.inl (or.inl ⟨_, ⟨r, rfl⟩, _, ⟨s.mem_attach r, rfl⟩, hx⟩) }, + { rw [submonoid.powers_eq_closure, submonoid.closure_le, set.singleton_subset_iff], + apply algebra.subset_adjoin, + exact or.inl (or.inr r.2) }, + { rw ht, trivial } } +end + +lemma finite_is_local : + property_is_local @finite_type := +⟨localization_finite_type, finite_type_of_localization_span_target, + finite_type_stable_under_composition, finite_type_holds_for_localization_away⟩ + +end ring_hom From 0aa5432b2f21be20087539ea558b95de8750bfcb Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 11:14:37 +0000 Subject: [PATCH 089/173] refactor(topology/sheaves): Redefine sheaves in terms of Grothendieck topology. (#15384) We change the "official definition" of sheaves over topological spaces from the equalizer diagram condition to the condition in terms of the Grothendieck topology on `Opens(X)`. The benefit is that 1. The `X.Sheaf C` category is now defeq to `(opens.grothendieck_topology X).Sheaf C`, so that functors between categories of sheaves over sites could be specialized onto topological spaces without the abundant equivalences. 2. It allows sheaves over spaces to value in arbitrary categories that doesn't have all products. The original sheaf condition is now called `presheaf.is_sheaf_equalizer_products`, and `presheaf.is_sheaf_iff_is_sheaf_equalizer_products` (in `topology.sheaves.sheaf_condition.sites`) shows that the two are equivalent. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/AffineScheme.lean | 5 +- .../Gamma_Spec_adjunction.lean | 4 +- src/algebraic_geometry/open_immersion.lean | 2 +- .../projective_spectrum/structure_sheaf.lean | 21 ++-- src/algebraic_geometry/sheafed_space.lean | 5 +- src/algebraic_geometry/structure_sheaf.lean | 44 ++++---- src/category_theory/sites/sheaf.lean | 4 + src/topology/sheaves/forget.lean | 4 +- src/topology/sheaves/functors.lean | 4 +- src/topology/sheaves/limits.lean | 7 +- src/topology/sheaves/local_predicate.lean | 10 +- src/topology/sheaves/sheaf.lean | 104 +++++++++++------- .../sheaf_condition/equalizer_products.lean | 9 ++ .../sheaf_condition/opens_le_cover.lean | 19 +--- .../pairwise_intersections.lean | 19 ++-- .../sheaves/sheaf_condition/sites.lean | 74 +++---------- .../sheaf_condition/unique_gluing.lean | 4 +- src/topology/sheaves/sheafify.lean | 4 +- src/topology/sheaves/stalks.lean | 52 ++++----- 19 files changed, 198 insertions(+), 197 deletions(-) diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 343c6100d0084..79a128d633e17 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -491,9 +491,10 @@ begin erw [← X.presheaf.map_comp, Spec_Γ_naturality_assoc], congr' 1, simp only [← category.assoc], - transitivity _ ≫ (structure_sheaf (X.presheaf.obj $ op U)).1.germ ⟨_, _⟩, + transitivity _ ≫ (structure_sheaf (X.presheaf.obj $ op U)).presheaf.germ ⟨_, _⟩, { refl }, - convert ((structure_sheaf (X.presheaf.obj $ op U)).1.germ_res (hom_of_le le_top) ⟨_, _⟩) using 2, + convert ((structure_sheaf (X.presheaf.obj $ op U)).presheaf.germ_res (hom_of_le le_top) ⟨_, _⟩) + using 2, rw category.assoc, erw nat_trans.naturality, rw [← LocallyRingedSpace.Γ_map_op, ← LocallyRingedSpace.Γ.map_comp_assoc, ← op_comp], diff --git a/src/algebraic_geometry/Gamma_Spec_adjunction.lean b/src/algebraic_geometry/Gamma_Spec_adjunction.lean index 1678717709f5f..1c6a650e38820 100644 --- a/src/algebraic_geometry/Gamma_Spec_adjunction.lean +++ b/src/algebraic_geometry/Gamma_Spec_adjunction.lean @@ -191,7 +191,7 @@ def to_Γ_Spec : X ⟶ Spec.LocallyRingedSpace_obj (Γ.obj (op X)) := intro x, let p : prime_spectrum (Γ.obj (op X)) := X.to_Γ_Spec_fun x, constructor, /- show stalk map is local hom ↓ -/ - let S := (structure_sheaf _).val.stalk p, + let S := (structure_sheaf _).presheaf.stalk p, rintros (t : S) ht, obtain ⟨⟨r, s⟩, he⟩ := is_localization.surj p.as_ideal.prime_compl t, dsimp at he, @@ -272,7 +272,7 @@ begin apply LocallyRingedSpace.comp_ring_hom_ext, { ext (p : prime_spectrum R) x, erw ← is_localization.at_prime.to_map_mem_maximal_iff - ((structure_sheaf R).val.stalk p) p.as_ideal x, + ((structure_sheaf R).presheaf.stalk p) p.as_ideal x, refl }, { intro r, apply to_open_res }, end diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index c8f5a8b7251dd..e6be7a87c13f1 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -772,7 +772,7 @@ lemma of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) { 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 }, + (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, diff --git a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean index 7e838ea543a1b..e68917a294a1c 100644 --- a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean +++ b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean @@ -212,7 +212,7 @@ def Proj.structure_sheaf : sheaf CommRing (projective_spectrum.Top 𝒜) := -- We check the sheaf condition under `forget CommRing`. (is_sheaf_iff_is_sheaf_comp _ _).mpr (is_sheaf_of_iso (structure_presheaf_comp_forget 𝒜).symm - (structure_sheaf_in_Type 𝒜).property)⟩ + (structure_sheaf_in_Type 𝒜).cond)⟩ end projective_spectrum @@ -247,25 +247,26 @@ def open_to_localization (U : opens (projective_spectrum.Top 𝒜)) (x : project to a homogeneous prime ideal `x` to the *homogeneous localization* at `x`, formed by gluing the `open_to_localization` maps. -/ def stalk_to_fiber_ring_hom (x : projective_spectrum.Top 𝒜) : - (Proj.structure_sheaf 𝒜).1.stalk x ⟶ CommRing.of (at x) := + (Proj.structure_sheaf 𝒜).presheaf.stalk x ⟶ CommRing.of (at x) := limits.colimit.desc (((open_nhds.inclusion x).op) ⋙ (Proj.structure_sheaf 𝒜).1) { X := _, ι := { app := λ U, open_to_localization 𝒜 ((open_nhds.inclusion _).obj (unop U)) x (unop U).2, } } @[simp] lemma germ_comp_stalk_to_fiber_ring_hom (U : opens (projective_spectrum.Top 𝒜)) (x : U) : - (Proj.structure_sheaf 𝒜).1.germ x ≫ stalk_to_fiber_ring_hom 𝒜 x = + (Proj.structure_sheaf 𝒜).presheaf.germ x ≫ stalk_to_fiber_ring_hom 𝒜 x = open_to_localization 𝒜 U x x.2 := limits.colimit.ι_desc _ _ @[simp] lemma stalk_to_fiber_ring_hom_germ' (U : opens (projective_spectrum.Top 𝒜)) (x : projective_spectrum.Top 𝒜) (hx : x ∈ U) (s : (Proj.structure_sheaf 𝒜).1.obj (op U)) : - stalk_to_fiber_ring_hom 𝒜 x ((Proj.structure_sheaf 𝒜).1.germ ⟨x, hx⟩ s) = (s.1 ⟨x, hx⟩ : _) := + stalk_to_fiber_ring_hom 𝒜 x + ((Proj.structure_sheaf 𝒜).presheaf.germ ⟨x, hx⟩ s) = (s.1 ⟨x, hx⟩ : _) := ring_hom.ext_iff.1 (germ_comp_stalk_to_fiber_ring_hom 𝒜 U ⟨x, hx⟩ : _) s @[simp] lemma stalk_to_fiber_ring_hom_germ (U : opens (projective_spectrum.Top 𝒜)) (x : U) (s : (Proj.structure_sheaf 𝒜).1.obj (op U)) : - stalk_to_fiber_ring_hom 𝒜 x ((Proj.structure_sheaf 𝒜).1.germ x s) = s.1 x := + stalk_to_fiber_ring_hom 𝒜 x ((Proj.structure_sheaf 𝒜).presheaf.germ x s) = s.1 x := by { cases x, exact stalk_to_fiber_ring_hom_germ' 𝒜 U _ _ _ } lemma homogeneous_localization.mem_basic_open (x : projective_spectrum.Top 𝒜) (f : at x) : @@ -292,19 +293,19 @@ def section_in_basic_open (x : projective_spectrum.Top 𝒜) : stalk at `x` obtained by `section_in_basic_open`. This is the inverse of `stalk_to_fiber_ring_hom`. -/ def homogeneous_localization_to_stalk (x : projective_spectrum.Top 𝒜) : - (at x) → (Proj.structure_sheaf 𝒜).1.stalk x := -λ f, (Proj.structure_sheaf 𝒜).1.germ + (at x) → (Proj.structure_sheaf 𝒜).presheaf.stalk x := +λ f, (Proj.structure_sheaf 𝒜).presheaf.germ (⟨x, homogeneous_localization.mem_basic_open _ x f⟩ : projective_spectrum.basic_open _ f.denom) (section_in_basic_open _ x f) /--Using `homogeneous_localization_to_stalk`, we construct a ring isomorphism between stalk at `x` and homogeneous localization at `x` for any point `x` in `Proj`.-/ def Proj.stalk_iso' (x : projective_spectrum.Top 𝒜) : - (Proj.structure_sheaf 𝒜).1.stalk x ≃+* CommRing.of (at x) := + (Proj.structure_sheaf 𝒜).presheaf.stalk x ≃+* CommRing.of (at x) := ring_equiv.of_bijective (stalk_to_fiber_ring_hom _ x) ⟨λ z1 z2 eq1, begin - obtain ⟨u1, memu1, s1, rfl⟩ := (Proj.structure_sheaf 𝒜).1.germ_exist x z1, - obtain ⟨u2, memu2, s2, rfl⟩ := (Proj.structure_sheaf 𝒜).1.germ_exist x z2, + obtain ⟨u1, memu1, s1, rfl⟩ := (Proj.structure_sheaf 𝒜).presheaf.germ_exist x z1, + obtain ⟨u2, memu2, s2, rfl⟩ := (Proj.structure_sheaf 𝒜).presheaf.germ_exist x z2, obtain ⟨v1, memv1, i1, ⟨j1, ⟨a1, a1_mem⟩, ⟨b1, b1_mem⟩, hs1⟩⟩ := s1.2 ⟨x, memu1⟩, obtain ⟨v2, memv2, i2, ⟨j2, ⟨a2, a2_mem⟩, ⟨b2, b2_mem⟩, hs2⟩⟩ := s2.2 ⟨x, memu2⟩, obtain ⟨b1_nin_x, eq2⟩ := hs1 ⟨x, memv1⟩, diff --git a/src/algebraic_geometry/sheafed_space.lean b/src/algebraic_geometry/sheafed_space.lean index 8b6b807b50c3d..3f42c317cab80 100644 --- a/src/algebraic_geometry/sheafed_space.lean +++ b/src/algebraic_geometry/sheafed_space.lean @@ -119,8 +119,9 @@ The restriction of a sheafed space along an open embedding into the space. -/ def restrict {U : Top} (X : SheafedSpace C) {f : U ⟶ (X : Top.{v})} (h : open_embedding f) : SheafedSpace C := -{ is_sheaf := λ ι 𝒰, ⟨is_limit.of_iso_limit - ((is_limit.postcompose_inv_equiv _ _).inv_fun (X.is_sheaf _).some) +{ is_sheaf := (is_sheaf_iff_is_sheaf_equalizer_products _).mpr $ λ ι 𝒰, ⟨is_limit.of_iso_limit + ((is_limit.postcompose_inv_equiv _ _).inv_fun + ((is_sheaf_iff_is_sheaf_equalizer_products _).mp X.is_sheaf _).some) (sheaf_condition_equalizer_products.fork.iso_of_open_embedding h 𝒰).symm⟩, ..X.to_PresheafedSpace.restrict h } diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index efbf261969500..dd27bc65946b5 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -262,7 +262,7 @@ def Spec.structure_sheaf : sheaf CommRing (prime_spectrum.Top R) := -- We check the sheaf condition under `forget CommRing`. (is_sheaf_iff_is_sheaf_comp _ _).mpr (is_sheaf_of_iso (structure_presheaf_comp_forget R).symm - (structure_sheaf_in_Type R).property)⟩ + (structure_sheaf_in_Type R).cond)⟩ open Spec (structure_sheaf) @@ -402,20 +402,20 @@ subtype.eq $ funext $ λ x, eq.symm $ is_localization.mk'_one _ f /-- The canonical ring homomorphism interpreting an element of `R` as an element of the stalk of `structure_sheaf R` at `x`. -/ -def to_stalk (x : prime_spectrum.Top R) : CommRing.of R ⟶ (structure_sheaf R).1.stalk x := -(to_open R ⊤ ≫ (structure_sheaf R).1.germ ⟨x, ⟨⟩⟩ : _) +def to_stalk (x : prime_spectrum.Top R) : CommRing.of R ⟶ (structure_sheaf R).presheaf.stalk x := +(to_open R ⊤ ≫ (structure_sheaf R).presheaf.germ ⟨x, ⟨⟩⟩ : _) @[simp] lemma to_open_germ (U : opens (prime_spectrum.Top R)) (x : U) : - to_open R U ≫ (structure_sheaf R).1.germ x = + to_open R U ≫ (structure_sheaf R).presheaf.germ x = to_stalk R x := by { rw [← to_open_res R ⊤ U (hom_of_le le_top : U ⟶ ⊤), category.assoc, presheaf.germ_res], refl } @[simp] lemma germ_to_open (U : opens (prime_spectrum.Top R)) (x : U) (f : R) : - (structure_sheaf R).1.germ x (to_open R U f) = to_stalk R x f := + (structure_sheaf R).presheaf.germ x (to_open R U f) = to_stalk R x f := by { rw ← to_open_germ, refl } lemma germ_to_top (x : prime_spectrum.Top R) (f : R) : - (structure_sheaf R).1.germ (⟨x, trivial⟩ : (⊤ : opens (prime_spectrum.Top R))) + (structure_sheaf R).presheaf.germ (⟨x, trivial⟩ : (⊤ : opens (prime_spectrum.Top R))) (to_open R ⊤ f) = to_stalk R x f := rfl @@ -432,7 +432,7 @@ by { erw ← germ_to_open R (basic_open (f : R)) ⟨x, f.2⟩ (f : R), /-- The canonical ring homomorphism from the localization of `R` at `p` to the stalk of the structure sheaf at the point `p`. -/ def localization_to_stalk (x : prime_spectrum.Top R) : - CommRing.of (localization.at_prime x.as_ideal) ⟶ (structure_sheaf R).1.stalk x := + CommRing.of (localization.at_prime x.as_ideal) ⟶ (structure_sheaf R).presheaf.stalk x := show localization.at_prime x.as_ideal →+* _, from is_localization.lift (is_unit_to_stalk R x) @@ -443,7 +443,7 @@ is_localization.lift_eq _ f @[simp] lemma localization_to_stalk_mk' (x : prime_spectrum.Top R) (f : R) (s : (as_ideal x).prime_compl) : localization_to_stalk R x (is_localization.mk' _ f s : localization _) = - (structure_sheaf R).1.germ (⟨x, s.2⟩ : basic_open (s : R)) + (structure_sheaf R).presheaf.germ (⟨x, s.2⟩ : basic_open (s : R)) (const R f s (basic_open s) (λ _, id)) := (is_localization.lift_mk'_spec _ _ _ _).2 $ by erw [← germ_to_open R (basic_open s) ⟨x, s.2⟩, ← germ_to_open R (basic_open s) ⟨x, s.2⟩, @@ -478,25 +478,25 @@ rfl a prime ideal `p` to the localization of `R` at `p`, formed by gluing the `open_to_localization` maps. -/ def stalk_to_fiber_ring_hom (x : prime_spectrum.Top R) : - (structure_sheaf R).1.stalk x ⟶ CommRing.of (localization.at_prime x.as_ideal) := + (structure_sheaf R).presheaf.stalk x ⟶ CommRing.of (localization.at_prime x.as_ideal) := limits.colimit.desc (((open_nhds.inclusion x).op) ⋙ (structure_sheaf R).1) { X := _, ι := { app := λ U, open_to_localization R ((open_nhds.inclusion _).obj (unop U)) x (unop U).2, } } @[simp] lemma germ_comp_stalk_to_fiber_ring_hom (U : opens (prime_spectrum.Top R)) (x : U) : - (structure_sheaf R).1.germ x ≫ stalk_to_fiber_ring_hom R x = + (structure_sheaf R).presheaf.germ x ≫ stalk_to_fiber_ring_hom R x = open_to_localization R U x x.2 := limits.colimit.ι_desc _ _ @[simp] lemma stalk_to_fiber_ring_hom_germ' (U : opens (prime_spectrum.Top R)) (x : prime_spectrum.Top R) (hx : x ∈ U) (s : (structure_sheaf R).1.obj (op U)) : - stalk_to_fiber_ring_hom R x ((structure_sheaf R).1.germ ⟨x, hx⟩ s) = (s.1 ⟨x, hx⟩ : _) := + stalk_to_fiber_ring_hom R x ((structure_sheaf R).presheaf.germ ⟨x, hx⟩ s) = (s.1 ⟨x, hx⟩ : _) := ring_hom.ext_iff.1 (germ_comp_stalk_to_fiber_ring_hom R U ⟨x, hx⟩ : _) s @[simp] lemma stalk_to_fiber_ring_hom_germ (U : opens (prime_spectrum.Top R)) (x : U) (s : (structure_sheaf R).1.obj (op U)) : - stalk_to_fiber_ring_hom R x ((structure_sheaf R).1.germ x s) = s.1 x := + stalk_to_fiber_ring_hom R x ((structure_sheaf R).presheaf.germ x s) = s.1 x := by { cases x, exact stalk_to_fiber_ring_hom_germ' R U _ _ _ } @[simp] lemma to_stalk_comp_stalk_to_fiber_ring_hom (x : prime_spectrum.Top R) : @@ -510,15 +510,15 @@ ring_hom.ext_iff.1 (to_stalk_comp_stalk_to_fiber_ring_hom R x) _ /-- The ring isomorphism between the stalk of the structure sheaf of `R` at a point `p` corresponding to a prime ideal in `R` and the localization of `R` at `p`. -/ @[simps] def stalk_iso (x : prime_spectrum.Top R) : - (structure_sheaf R).1.stalk x ≅ CommRing.of (localization.at_prime x.as_ideal) := + (structure_sheaf R).presheaf.stalk x ≅ CommRing.of (localization.at_prime x.as_ideal) := { hom := stalk_to_fiber_ring_hom R x, inv := localization_to_stalk R x, - hom_inv_id' := (structure_sheaf R).1.stalk_hom_ext $ λ U hxU, + hom_inv_id' := (structure_sheaf R).presheaf.stalk_hom_ext $ λ U hxU, begin ext s, simp only [comp_apply], rw [id_apply, stalk_to_fiber_ring_hom_germ'], obtain ⟨V, hxV, iVU, f, g, hg, hs⟩ := exists_const _ _ s x hxU, erw [← res_apply R U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localization_to_stalk_mk'], - refine (structure_sheaf R).1.germ_ext V hxV (hom_of_le hg) iVU _, + refine (structure_sheaf R).presheaf.germ_ext V hxV (hom_of_le hg) iVU _, erw [← hs, res_const'] end, inv_hom_id' := @is_localization.ring_hom_ext R _ x.as_ideal.prime_compl @@ -813,15 +813,15 @@ def basic_open_iso (f : R) : (structure_sheaf R).1.obj (op (basic_open f)) ≅ CommRing.of (localization.away f) := (as_iso (show CommRing.of _ ⟶ _, from to_basic_open R f)).symm -instance stalk_algebra (p : prime_spectrum R) : algebra R ((structure_sheaf R).val.stalk p) := +instance stalk_algebra (p : prime_spectrum R) : algebra R ((structure_sheaf R).presheaf.stalk p) := (to_stalk R p).to_algebra @[simp] lemma stalk_algebra_map (p : prime_spectrum R) (r : R) : - algebra_map R ((structure_sheaf R).val.stalk p) r = to_stalk R p r := rfl + algebra_map R ((structure_sheaf R).presheaf.stalk p) r = to_stalk R p r := rfl /-- Stalk of the structure sheaf at a prime p as localization of R -/ instance is_localization.to_stalk (p : prime_spectrum R) : - is_localization.at_prime ((structure_sheaf R).val.stalk p) p.as_ideal := + is_localization.at_prime ((structure_sheaf R).presheaf.stalk p) p.as_ideal := begin convert (is_localization.is_localization_iff_of_ring_equiv _ (stalk_iso R p).symm .CommRing_iso_to_ring_equiv).mp localization.is_localization, @@ -885,13 +885,13 @@ as_iso (to_open R ⊤) @[simp, reassoc, elementwise] lemma to_stalk_stalk_specializes {R : Type*} [comm_ring R] {x y : prime_spectrum R} (h : x ⤳ y) : - to_stalk R y ≫ (structure_sheaf R).val.stalk_specializes h = to_stalk R x := -by { dsimp [ to_stalk], simpa } + to_stalk R y ≫ (structure_sheaf R).presheaf.stalk_specializes h = to_stalk R x := +by { dsimp[to_stalk], simpa [-to_open_germ], } @[simp, reassoc, elementwise] lemma localization_to_stalk_stalk_specializes {R : Type*} [comm_ring R] {x y : prime_spectrum R} (h : x ⤳ y) : - structure_sheaf.localization_to_stalk R y ≫ (structure_sheaf R).val.stalk_specializes h = + structure_sheaf.localization_to_stalk R y ≫ (structure_sheaf R).presheaf.stalk_specializes h = CommRing.of_hom (prime_spectrum.localization_map_of_specializes h) ≫ structure_sheaf.localization_to_stalk R x := begin @@ -907,7 +907,7 @@ end @[simp, reassoc, elementwise] lemma stalk_specializes_stalk_to_fiber {R : Type*} [comm_ring R] {x y : prime_spectrum R} (h : x ⤳ y) : - (structure_sheaf R).val.stalk_specializes h ≫ structure_sheaf.stalk_to_fiber_ring_hom R x = + (structure_sheaf R).presheaf.stalk_specializes h ≫ structure_sheaf.stalk_to_fiber_ring_hom R x = structure_sheaf.stalk_to_fiber_ring_hom R y ≫ prime_spectrum.localization_map_of_specializes h := begin diff --git a/src/category_theory/sites/sheaf.lean b/src/category_theory/sites/sheaf.lean index 11ad5aa1fb7dc..060e5ad495dd1 100644 --- a/src/category_theory/sites/sheaf.lean +++ b/src/category_theory/sites/sheaf.lean @@ -206,6 +206,10 @@ lemma is_sheaf.hom_ext {A : Type u₂} [category.{max v₁ u₁} A] e₁ = e₂ := (hP _ _ S.condition).is_separated_for.ext (λ Y f hf, h ⟨Y,f,hf⟩) +lemma is_sheaf_of_iso_iff {P P' : Cᵒᵖ ⥤ A} (e : P ≅ P') : is_sheaf J P ↔ is_sheaf J P' := +forall_congr $ λ a, ⟨presieve.is_sheaf_iso J (iso_whisker_right e _), + presieve.is_sheaf_iso J (iso_whisker_right e.symm _)⟩ + variable (J) end presheaf diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index 60e8e07ec58f3..4d1543b4af09e 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import category_theory.limits.preserves.shapes.products -import topology.sheaves.sheaf +import topology.sheaves.sheaf_condition.sites /-! # Checking the sheaf condition on the underlying presheaf of types. @@ -129,6 +129,8 @@ In fact we prove a stronger version with arbitrary complete target category. lemma is_sheaf_iff_is_sheaf_comp : presheaf.is_sheaf F ↔ presheaf.is_sheaf (F ⋙ G) := begin + rw [presheaf.is_sheaf_iff_is_sheaf_equalizer_products, + presheaf.is_sheaf_iff_is_sheaf_equalizer_products], split, { intros S ι U, -- We have that the sheaf condition fork for `F` is a limit fork, diff --git a/src/topology/sheaves/functors.lean b/src/topology/sheaves/functors.lean index c8c0de8e47d46..28a320de36595 100644 --- a/src/topology/sheaves/functors.lean +++ b/src/topology/sheaves/functors.lean @@ -72,7 +72,7 @@ variables [has_products.{v} C] The pushforward of a sheaf (by a continuous map) is a sheaf. -/ theorem pushforward_sheaf_of_sheaf - {F : presheaf C X} (h : F.is_sheaf) : (f _* F).is_sheaf := + {F : X.presheaf C} (h : F.is_sheaf) : (f _* F).is_sheaf := by rw is_sheaf_iff_is_sheaf_pairwise_intersections at h ⊢; exact sheaf_condition_pairwise_intersections.pushforward_sheaf_of_sheaf f h @@ -81,7 +81,7 @@ The pushforward functor. -/ def pushforward (f : X ⟶ Y) : X.sheaf C ⥤ Y.sheaf C := { obj := λ ℱ, ⟨f _* ℱ.1, pushforward_sheaf_of_sheaf f ℱ.2⟩, - map := λ _ _, pushforward_map f } + map := λ _ _ g, ⟨pushforward_map f g.1⟩ } end sheaf diff --git a/src/topology/sheaves/limits.lean b/src/topology/sheaves/limits.lean index 3612b1b8acc3f..f2dd492e36222 100644 --- a/src/topology/sheaves/limits.lean +++ b/src/topology/sheaves/limits.lean @@ -30,10 +30,7 @@ instance [has_colimits C] (X : Top) : has_colimits_of_size.{v} (presheaf C X) := limits.functor_category_has_colimits_of_size instance [has_limits C] (X : Top) : creates_limits (sheaf.forget C X) := -(@@creates_limits_of_nat_iso _ _ - (presheaf.Sheaf_spaces_equiv_sheaf_sites_inverse_forget C X)) - (@@category_theory.comp_creates_limits _ _ _ _ _ _ - Sheaf.category_theory.Sheaf_to_presheaf.category_theory.creates_limits.{u v v}) +Sheaf.category_theory.Sheaf_to_presheaf.category_theory.creates_limits.{u v v} instance [has_limits C] (X : Top) : has_limits_of_size.{v} (sheaf.{v} C X) := @@ -42,7 +39,7 @@ has_limits_of_has_limits_creates_limits (sheaf.forget C X) lemma is_sheaf_of_is_limit [has_limits C] {X : Top} (F : J ⥤ presheaf.{v} C X) (H : ∀ j, (F.obj j).is_sheaf) {c : cone F} (hc : is_limit c) : c.X.is_sheaf := begin - let F' : J ⥤ sheaf C X := { obj := λ j, ⟨F.obj j, H j⟩, map := F.map }, + let F' : J ⥤ sheaf C X := { obj := λ j, ⟨F.obj j, H j⟩, map := λ X Y f, ⟨F.map f⟩ }, let e : F' ⋙ sheaf.forget C X ≅ F := nat_iso.of_components (λ _, iso.refl _) (by tidy), exact presheaf.is_sheaf_of_iso ((is_limit_of_preserves (sheaf.forget C X) (limit.is_limit F')).cone_points_iso_of_nat_iso hc e) (limit F').2 diff --git a/src/topology/sheaves/local_predicate.lean b/src/topology/sheaves/local_predicate.lean index 69d86f56d3a23..4acca6972ccd2 100644 --- a/src/topology/sheaves/local_predicate.lean +++ b/src/topology/sheaves/local_predicate.lean @@ -215,7 +215,7 @@ def subsheaf_to_Types (P : local_predicate T) : sheaf (Type v) X := 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).1.stalk x ⟶ T x := + (subsheaf_to_Types P).presheaf.stalk x ⟶ T x := begin refine colimit.desc _ { X := T x, ι := { app := λ U f, _, naturality' := _ } }, @@ -224,7 +224,7 @@ begin end @[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).1.germ x f) = f.1 x := + stalk_to_fiber P x ((subsheaf_to_Types P).presheaf.germ x f) = f.1 x := begin dsimp [presheaf.germ, stalk_to_fiber], cases x, @@ -243,7 +243,7 @@ lemma stalk_to_fiber_surjective (P : local_predicate T) (x : X) begin rcases w t with ⟨U, f, h, rfl⟩, fsplit, - { exact (subsheaf_to_Types P).1.germ ⟨x, U.2⟩ ⟨f, h⟩, }, + { exact (subsheaf_to_Types P).presheaf.germ ⟨x, U.2⟩ ⟨f, h⟩, }, { exact stalk_to_fiber_germ _ U.1 ⟨x, U.2⟩ ⟨f, h⟩, } end @@ -261,8 +261,8 @@ 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), - tU = (subsheaf_to_Types P).1.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ ∧ - tV = (subsheaf_to_Types P).1.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ := _, + 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, exact e.1.trans e.2.symm, }, -- Then use induction to pick particular representatives of `tU tV : stalk x` diff --git a/src/topology/sheaves/sheaf.lean b/src/topology/sheaves/sheaf.lean index a63ac5bf08704..794e94d5c2aa8 100644 --- a/src/topology/sheaves/sheaf.lean +++ b/src/topology/sheaves/sheaf.lean @@ -6,35 +6,25 @@ Authors: Scott Morrison import topology.sheaves.sheaf_condition.equalizer_products import category_theory.full_subcategory import category_theory.limits.unit +import category_theory.sites.sheaf +import category_theory.sites.spaces /-! # Sheaves -We define sheaves on a topological space, with values in an arbitrary category with products. +We define sheaves on a topological space, with values in an arbitrary category. -The sheaf condition for a `F : presheaf C X` requires that the morphism -`F.obj U ⟶ ∏ F.obj (U i)` (where `U` is some open set which is the union of the `U i`) -is the equalizer of the two morphisms -`∏ F.obj (U i) ⟶ ∏ F.obj (U i ⊓ U j)`. +A presheaf on a topological space `X` is a sheaf presicely when it is a sheaf under the +grothendieck topology on `opens X`, which expands out to say: For each open cover `{ Uᵢ }` of +`U`, and a family of compatible functions `A ⟶ F(Uᵢ)` for an `A : X`, there exists an unique +gluing `A ⟶ F(U)` compatible with the restriction. + +See the docstring of `Top.presheaf.is_sheaf` for an explanation on the design descisions and a list +of equivalent conditions. We provide the instance `category (sheaf C X)` as the full subcategory of presheaves, and the fully faithful functor `sheaf.forget : sheaf C X ⥤ presheaf C X`. -## Equivalent conditions - -While the "official" definition is in terms of an equalizer diagram, -in `src/topology/sheaves/sheaf_condition/pairwise_intersections.lean` -and in `src/topology/sheaves/sheaf_condition/open_le_cover.lean` -we provide two equivalent conditions (and prove they are equivalent). - -The first is that `F.obj U` is the limit point of the diagram consisting of all the `F.obj (U i)` -and `F.obj (U i ⊓ U j)`. -(That is, we explode the equalizer of two products out into its component pieces.) - -The second is that `F.obj U` is the limit point of the diagram constisting of all the `F.obj V`, -for those `V : opens X` such that `V ≤ U i` for some `i`. -(This condition is particularly easy to state, and perhaps should become the "official" definition.) - -/ universes w v u @@ -49,7 +39,7 @@ open topological_space.opens namespace Top -variables {C : Type u} [category.{v} C] [has_products.{v} C] +variables {C : Type u} [category.{v} C] variables {X : Top.{w}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) namespace presheaf @@ -57,30 +47,60 @@ namespace presheaf open sheaf_condition_equalizer_products /-- -The sheaf condition for a `F : presheaf C X` requires that the morphism -`F.obj U ⟶ ∏ F.obj (U i)` (where `U` is some open set which is the union of the `U i`) -is the equalizer of the two morphisms -`∏ F.obj (U i) ⟶ ∏ F.obj (U i) ⊓ (U j)`. +The sheaf condition has several different equivalent formulations. +The official definition chosen here is in terms of grothendieck topologies so that the results on +sites could be applied here easily, and this condition does not require additional constraints on +the value category. +The equivalent formulations of the sheaf condition on `presheaf C X` are as follows : + +1. `Top.presheaf.is_sheaf`: (the official definition) + It is a sheaf with respect to the grothendieck topology on `opens X`, which is to say: + For each open cover `{ Uᵢ }` of `U`, and a family of compatible functions `A ⟶ F(Uᵢ)` for an + `A : X`, there exists an unique gluing `A ⟶ F(U)` compatible with the restriction. + +2. `Top.presheaf.is_sheaf_equalizer_products`: (requires `C` to have all products) + For each open cover `{ Uᵢ }` of `U`, `F(U) ⟶ ∏ F(Uᵢ)` is the equalizer of the two morphisms + `∏ F(Uᵢ) ⟶ ∏ F(Uᵢ ∩ Uⱼ)`. + See `Top.presheaf.is_sheaf_iff_is_sheaf_equalizer_products`. + +3. `Top.presheaf.is_sheaf_opens_le_cover`: + Each `F(U)` is the (inverse) limit of all `F(V)` with `V ⊆ U`. + See `Top.presheaf.is_sheaf_iff_is_sheaf_opens_le_cover`. + +4. `Top.presheaf.is_sheaf_pairwise_intersections`: + For each open cover `{ Uᵢ }` of `U`, `F(U)` is the limit of all `F(Uᵢ)` and all `F(Uᵢ ∩ Uⱼ)`. + See `Top.presheaf.is_sheaf_iff_is_sheaf_pairwise_intersections`. + +The following requires `C` to be concrete and complete, and `forget C` to reflect isomorphisms and +preserve limits. This applies to most "algebraic" categories, e.g. groups, abelian groups and rings. + +5. `Top.presheaf.is_sheaf_unique_gluing`: + (requires `C` to be concrete and complete; `forget C` to reflect isomorphisms and preserve limits) + For each open cover `{ Uᵢ }` of `U`, and a compatible family of elements `x : F(Uᵢ)`, there exists + a unique gluing `x : F(U)` that restricts to the given elements. + See `Top.presheaf.is_sheaf_iff_is_sheaf_unique_gluing`. + +6. The underlying sheaf of types is a sheaf. + See `Top.presheaf.is_sheaf_iff_is_sheaf_comp` and + `category_theory.presheaf.is_sheaf_iff_is_sheaf_forget`. -/ def is_sheaf (F : presheaf.{w v u} C X) : Prop := -∀ ⦃ι : Type v⦄ (U : ι → opens X), nonempty (is_limit (sheaf_condition_equalizer_products.fork F U)) +presheaf.is_sheaf (opens.grothendieck_topology X) F /-- The presheaf valued in `unit` over any topological space is a sheaf. -/ lemma is_sheaf_unit (F : presheaf (category_theory.discrete unit) X) : F.is_sheaf := -λ ι U, ⟨punit_cone_is_limit⟩ +λ x U S hS x hx, ⟨eq_to_hom (subsingleton.elim _ _), by tidy, by tidy⟩ + +lemma is_sheaf_iso_iff {F G : presheaf C X} (α : F ≅ G) : F.is_sheaf ↔ G.is_sheaf := +presheaf.is_sheaf_of_iso_iff α /-- Transfer the sheaf condition across an isomorphism of presheaves. -/ lemma is_sheaf_of_iso {F G : presheaf C X} (α : F ≅ G) (h : F.is_sheaf) : G.is_sheaf := -λ ι U, ⟨is_limit.of_iso_limit - ((is_limit.postcompose_inv_equiv _ _).symm (h U).some) - (sheaf_condition_equalizer_products.fork.iso_of_iso U α.symm).symm⟩ - -lemma is_sheaf_iso_iff {F G : presheaf C X} (α : F ≅ G) : F.is_sheaf ↔ G.is_sheaf := -⟨(λ h, is_sheaf_of_iso α h), (λ h, is_sheaf_of_iso α.symm h)⟩ +(is_sheaf_iso_iff α).1 h end presheaf @@ -91,7 +111,14 @@ A `sheaf C X` is a presheaf of objects from `C` over a (bundled) topological spa satisfying the sheaf condition. -/ @[derive category] -def sheaf : Type (max u v w) := { F : presheaf C X // F.is_sheaf } +def sheaf : Type (max u v w) := Sheaf (opens.grothendieck_topology X) C + +variables {C X} + +/-- The underlying presheaf of a sheaf -/ +abbreviation sheaf.presheaf (F : X.sheaf C) : Top.presheaf C X := F.1 + +variables (C X) -- Let's construct a trivial example, to keep the inhabited linter happy. instance sheaf_inhabited : inhabited (sheaf (category_theory.discrete punit) X) := @@ -104,11 +131,12 @@ The forgetful functor from sheaves to presheaves. -/ @[derive [full, faithful]] def forget : Top.sheaf C X ⥤ Top.presheaf C X := -full_subcategory_inclusion presheaf.is_sheaf +Sheaf_to_presheaf _ _ -@[simp] lemma id_app (F : sheaf C X) (t) : (𝟙 F : F ⟶ F).app t = 𝟙 _ := rfl -@[simp] lemma comp_app {F G H : sheaf C X} (f : F ⟶ G) (g : G ⟶ H) (t) : - (f ≫ g).app t = f.app t ≫ g.app t := rfl +-- Note: These can be proved by simp. +lemma id_app (F : sheaf C X) (t) : (𝟙 F : F ⟶ F).1.app t = 𝟙 _ := rfl +lemma comp_app {F G H : sheaf C X} (f : F ⟶ G) (g : G ⟶ H) (t) : + (f ≫ g).1.app t = f.1.app t ≫ g.1.app t := rfl end sheaf diff --git a/src/topology/sheaves/sheaf_condition/equalizer_products.lean b/src/topology/sheaves/sheaf_condition/equalizer_products.lean index a6285e62b494d..08226a3c3b28e 100644 --- a/src/topology/sheaves/sheaf_condition/equalizer_products.lean +++ b/src/topology/sheaves/sheaf_condition/equalizer_products.lean @@ -252,6 +252,15 @@ end open_embedding end sheaf_condition_equalizer_products +/-- +The sheaf condition for a `F : presheaf C X` requires that the morphism +`F.obj U ⟶ ∏ F.obj (U i)` (where `U` is some open set which is the union of the `U i`) +is the equalizer of the two morphisms +`∏ F.obj (U i) ⟶ ∏ F.obj (U i) ⊓ (U j)`. +-/ +def is_sheaf_equalizer_products (F : presheaf.{v' v u} C X) : Prop := +∀ ⦃ι : Type v⦄ (U : ι → opens X), nonempty (is_limit (sheaf_condition_equalizer_products.fork F U)) + end presheaf end Top diff --git a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean index da501d97c8057..7af25b4afdc98 100644 --- a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean +++ b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean @@ -305,10 +305,11 @@ end it satisfies the `is_sheaf_opens_le_cover` sheaf condition. The latter is not the official definition of sheaves on spaces, but has the advantage that it does not require `has_products C`. -/ -lemma is_sheaf_sites_iff_is_sheaf_opens_le_cover : - category_theory.presheaf.is_sheaf (opens.grothendieck_topology X) F ↔ F.is_sheaf_opens_le_cover := +lemma is_sheaf_iff_is_sheaf_opens_le_cover : + F.is_sheaf ↔ F.is_sheaf_opens_le_cover := begin - rw presheaf.is_sheaf_iff_is_limit, split, + refine (presheaf.is_sheaf_iff_is_limit _ _).trans _, + split, { intros h ι U, rw (is_limit_opens_le_equiv_generate₁ F U rfl).nonempty_congr, apply h, apply presieve_of_covering.mem_grothendieck_topology }, { intros h Y S, rw ← sieve.generate_sieve S, intro hS, @@ -317,18 +318,6 @@ end end -variable [has_products.{v} C] - -/-- -The sheaf condition in terms of an equalizer diagram is equivalent -to the reformulation in terms of a limit diagram over all `{ V : opens X // ∃ i, V ≤ U i }`. --/ -lemma is_sheaf_iff_is_sheaf_opens_le_cover (F : presheaf C X) : - F.is_sheaf ↔ F.is_sheaf_opens_le_cover := -iff.trans - (is_sheaf_iff_is_sheaf_pairwise_intersections F) - (is_sheaf_opens_le_cover_iff_is_sheaf_pairwise_intersections F).symm - end presheaf end Top diff --git a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean index 068a4d2fbc687..098a78851bc74 100644 --- a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean +++ b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean @@ -369,6 +369,7 @@ to the reformulation in terms of a limit diagram over `U i` and `U i ⊓ U j`. -/ lemma is_sheaf_iff_is_sheaf_pairwise_intersections (F : presheaf C X) : F.is_sheaf ↔ F.is_sheaf_pairwise_intersections := +(is_sheaf_iff_is_sheaf_equalizer_products F).trans $ iff.intro (λ h ι U, ⟨is_limit_map_cone_of_is_limit_sheaf_condition_fork F U (h U).some⟩) (λ h ι U, ⟨is_limit_sheaf_condition_fork_of_is_limit_map_cone F U (h U).some⟩) @@ -393,7 +394,7 @@ end Top.presheaf namespace Top.sheaf -variables {X : Top.{v}} {C : Type u} [category.{v} C] [has_products.{v} C] +variables {X : Top.{v}} {C : Type u} [category.{v} C] variables (F : X.sheaf C) (U V : opens X) open category_theory.limits @@ -414,8 +415,12 @@ pullback_cone.mk (F.1.map (hom_of_le le_sup_left).op) (F.1.map (hom_of_le le_sup variable (s : pullback_cone (F.1.map (hom_of_le inf_le_left : U ∩ V ⟶ _).op) (F.1.map (hom_of_le inf_le_right).op)) +variable [has_products.{v} C] + /-- (Implementation). -Every cone over `F(U) ⟶ F(U ∩ V)` and `F(V) ⟶ F(U ∩ V)` factors through `F(U ∪ V)`. -/ +Every cone over `F(U) ⟶ F(U ∩ V)` and `F(V) ⟶ F(U ∩ V)` factors through `F(U ∪ V)`. +TODO: generalize to `C` without products. +-/ def inter_union_pullback_cone_lift : s.X ⟶ F.1.obj (op (U ∪ V)) := begin let ι : ulift.{v} walking_pair → opens X := λ j, walking_pair.cases_on j.down U V, @@ -427,7 +432,7 @@ begin exacts [⟨⟨walking_pair.left⟩, h⟩, ⟨⟨walking_pair.right⟩, h⟩] }, { rintro ⟨⟨_ | _⟩, h⟩, exacts [or.inl h, or.inr h] } }, - refine (F.1.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 ι).some.lift + refine (F.presheaf.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 ι).some.lift ⟨s.X, { app := _, naturality' := _ }⟩ ≫ F.1.map (eq_to_hom hι).op, { apply opposite.rec, rintro ((_|_)|(_|_)), @@ -443,14 +448,14 @@ begin category.assoc], erw [← F.1.map_comp, ← F.1.map_comp], convert s.condition.symm }, - { convert s.condition } end lemma inter_union_pullback_cone_lift_left : inter_union_pullback_cone_lift F U V s ≫ F.1.map (hom_of_le le_sup_left).op = s.fst := begin + dsimp, erw [category.assoc, ←F.1.map_comp], - exact (F.1.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 _).some.fac _ + exact (F.presheaf.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 _).some.fac _ (op $ pairwise.single (ulift.up walking_pair.left)) end @@ -458,7 +463,7 @@ lemma inter_union_pullback_cone_lift_right : inter_union_pullback_cone_lift F U V s ≫ F.1.map (hom_of_le le_sup_right).op = s.snd := begin erw [category.assoc, ←F.1.map_comp], - exact (F.1.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 _).some.fac _ + exact (F.presheaf.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 _).some.fac _ (op $ pairwise.single (ulift.up walking_pair.right)) end @@ -482,7 +487,7 @@ begin { apply inter_union_pullback_cone_lift_right }, { intros m h₁ h₂, rw ← cancel_mono (F.1.map (eq_to_hom hι.symm).op), - apply (F.1.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 ι).some.hom_ext, + apply (F.presheaf.is_sheaf_iff_is_sheaf_pairwise_intersections.mp F.2 ι).some.hom_ext, apply opposite.rec, rintro ((_|_)|(_|_)); rw [category.assoc, category.assoc], { erw ← F.1.map_comp, diff --git a/src/topology/sheaves/sheaf_condition/sites.lean b/src/topology/sheaves/sheaf_condition/sites.lean index 0acfebdfea8a7..7cb8f2bfb00d0 100644 --- a/src/topology/sheaves/sheaf_condition/sites.lean +++ b/src/topology/sheaves/sheaf_condition/sites.lean @@ -12,10 +12,10 @@ import category_theory.sites.dense_subsite # The sheaf condition in terms of sites. -The theory of sheaves on sites is developed independently from sheaves on spaces in +The theory of sheaves on sites was developed independently from sheaves on spaces in `category_theory/sites`. In this file, we connect the two theories: We show that for a topological space `X`, a presheaf `F : (opens X)ᵒᵖ ⥤ C` is a sheaf on the site `opens X` if and only if it is -a sheaf on `X` in the usual sense. +a sheaf on `X` in terms of `is_sheaf_equalizer_products`. Recall that a presheaf `F : (opens X)ᵒᵖ ⥤ C` is called a *sheaf* on the space `X`, if for every family of opens `U : ι → opens X`, the object `F.obj (op (supr U))` is the limit of some fork @@ -26,10 +26,8 @@ and `presieve_of_covering`, which translate between the two concepts. We then pr naturality lemmas relating the two fork diagrams to each other. ## Main statements -* `is_sheaf_sites_iff_is_sheaf_spaces`. A presheaf `F : (opens X)ᵒᵖ ⥤ C` is a sheaf on the site - `opens X` if and only if it is a sheaf on the space `X`. -* `Sheaf_sites_eq_sheaf_spaces`. The type of sheaves on the site `opens X` is *equal* to the type - of sheaves on the space `X`. +* `is_sheaf_iff_is_sheaf_equalizer_products`. A presheaf `F : (opens X)ᵒᵖ ⥤ C` is a sheaf on the + site `opens X` if and only if it is a sheaf on the space `X`. -/ @@ -231,9 +229,10 @@ as_iso (postcompose_diagram_fork_hom F U R hR) end covering_of_presieve -lemma is_sheaf_sites_of_is_sheaf_spaces (Fsh : F.is_sheaf) : - presheaf.is_sheaf (opens.grothendieck_topology X) F := +lemma is_sheaf_of_is_sheaf_equalizer_products (Fsh : F.is_sheaf_equalizer_products) : + F.is_sheaf := begin + delta is_sheaf, rw presheaf.is_sheaf_iff_is_sheaf', intros U R hR, refine ⟨_⟩, @@ -424,11 +423,12 @@ end presieve_of_covering open presieve_of_covering -lemma is_sheaf_spaces_of_is_sheaf_sites - (Fsh : presheaf.is_sheaf (opens.grothendieck_topology X) F) : - F.is_sheaf := +lemma is_sheaf_equalizer_products_of_is_sheaf + (Fsh : F.is_sheaf) : + F.is_sheaf_equalizer_products := begin intros ι U, + delta is_sheaf at Fsh, rw presheaf.is_sheaf_iff_is_sheaf' at Fsh, -- We know that the sites diagram for `presieve_of_covering U` is a limit fork obtain ⟨h_limit⟩ := Fsh (supr U) (presieve_of_covering U) @@ -454,47 +454,9 @@ begin erw [← category.assoc, hm], }, end -lemma is_sheaf_sites_iff_is_sheaf_spaces : - presheaf.is_sheaf (opens.grothendieck_topology X) F ↔ F.is_sheaf := -iff.intro (is_sheaf_spaces_of_is_sheaf_sites F) (is_sheaf_sites_of_is_sheaf_spaces F) - -variables (C X) - -/-- Turn a sheaf on the site `opens X` into a sheaf on the space `X`. -/ -@[simps] -def Sheaf_sites_to_sheaf_spaces : Sheaf (opens.grothendieck_topology X) C ⥤ sheaf C X := -{ obj := λ F, ⟨F.1, is_sheaf_spaces_of_is_sheaf_sites F.1 F.2⟩, - map := λ F G f, f.val } - -/-- Turn a sheaf on the space `X` into a sheaf on the site `opens X`. -/ -@[simps] -def Sheaf_spaces_to_sheaf_sites : sheaf C X ⥤ Sheaf (opens.grothendieck_topology X) C := -{ obj := λ F, ⟨F.1, is_sheaf_sites_of_is_sheaf_spaces F.1 F.2⟩, - map := λ F G f, ⟨f⟩ } - -/-- -The equivalence of categories between sheaves on the site `opens X` and sheaves on the space `X`. --/ -@[simps] -def Sheaf_spaces_equiv_sheaf_sites : Sheaf (opens.grothendieck_topology X) C ≌ sheaf C X := -{ functor := Sheaf_sites_to_sheaf_spaces C X, - inverse := Sheaf_spaces_to_sheaf_sites C X, - unit_iso := nat_iso.of_components (λ t, ⟨⟨𝟙 _⟩, ⟨𝟙 _⟩, by { ext1, simp }, by { ext1, simp }⟩) $ - by { intros, ext1, dsimp, simp }, - counit_iso := nat_iso.of_components (λ t, ⟨𝟙 _, 𝟙 _, by { ext, simp }, by { ext, simp }⟩) $ - by { intros, ext, dsimp, simp } } - -/-- The two forgetful functors are isomorphic via `Sheaf_spaces_equiv_sheaf_sites`. -/ -def Sheaf_spaces_equiv_sheaf_sites_functor_forget : - (Sheaf_spaces_equiv_sheaf_sites C X).functor ⋙ sheaf.forget C X ≅ Sheaf_to_presheaf _ _ := -nat_iso.of_components (λ F, (iso.refl F.1)) - (λ F G f, by { erw [category.comp_id, category.id_comp], refl }) - -/-- The two forgetful functors are isomorphic via `Sheaf_spaces_equiv_sheaf_sites`. -/ -def Sheaf_spaces_equiv_sheaf_sites_inverse_forget : - (Sheaf_spaces_equiv_sheaf_sites C X).inverse ⋙ Sheaf_to_presheaf _ _ ≅ sheaf.forget C X := -nat_iso.of_components (λ F, (iso.refl F.1)) - (λ F G f, by { erw [category.comp_id, category.id_comp], refl }) +lemma is_sheaf_iff_is_sheaf_equalizer_products : + F.is_sheaf ↔ F.is_sheaf_equalizer_products := +iff.intro (is_sheaf_equalizer_products_of_is_sheaf F) (is_sheaf_of_is_sheaf_equalizer_products F) end Top.presheaf @@ -524,13 +486,13 @@ namespace Top.sheaf open category_theory topological_space Top opposite -variables {C : Type u} [category.{v} C] [limits.has_products.{v} C] +variables {C : Type u} [category.{v} C] variables {X : Top.{v}} {ι : Type*} {B : ι → opens X} -variables (F : presheaf C X) (F' : sheaf C X) (h : opens.is_basis (set.range B)) +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 ∅)) := -((presheaf.Sheaf_spaces_to_sheaf_sites C X).obj F).is_terminal_of_bot_cover ∅ (by tidy) +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 = ∅) : @@ -544,7 +506,7 @@ by convert F.is_terminal_of_empty def restrict_hom_equiv_hom : ((induced_functor B).op ⋙ F ⟶ (induced_functor B).op ⋙ F'.1) ≃ (F ⟶ F'.1) := @cover_dense.restrict_hom_equiv_hom _ _ _ _ _ _ _ _ (opens.cover_dense_induced_functor h) - _ F ((presheaf.Sheaf_spaces_to_sheaf_sites C X).obj F') + _ F F' @[simp] lemma extend_hom_app (α : ((induced_functor B).op ⋙ F ⟶ (induced_functor B).op ⋙ F'.1)) (i : ι) : (restrict_hom_equiv_hom F F' h α).app (op (B i)) = α.app (op i) := diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index 379304cb7482a..7e9b68806d3b1 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -144,6 +144,7 @@ in terms of unique gluings. lemma is_sheaf_of_is_sheaf_unique_gluing_types (Fsh : F.is_sheaf_unique_gluing) : F.is_sheaf := begin + rw is_sheaf_iff_is_sheaf_equalizer_products, intros ι U, refine ⟨fork.is_limit.mk' _ _⟩, intro s, @@ -172,6 +173,7 @@ The sheaf condition in terms of unique gluings can be obtained from the usual lemma is_sheaf_unique_gluing_of_is_sheaf_types (Fsh : F.is_sheaf) : F.is_sheaf_unique_gluing := begin + rw is_sheaf_iff_is_sheaf_equalizer_products at Fsh, intros ι U sf hsf, let sf' := (pi_opens_iso_sections_family F U).inv sf, have hsf' : left_res F U sf' = right_res F U sf', @@ -241,7 +243,7 @@ A more convenient way of obtaining a unique gluing of sections for a sheaf. lemma exists_unique_gluing (sf : Π i : ι, F.1.obj (op (U i))) (h : is_compatible F.1 U sf ) : ∃! s : F.1.obj (op (supr U)), is_gluing F.1 U sf s := -(is_sheaf_iff_is_sheaf_unique_gluing F.1).mp F.property U sf h +(is_sheaf_iff_is_sheaf_unique_gluing F.1).mp F.cond U sf h /-- In this version of the lemma, the inclusion homs `iUV` can be specified directly by the user, diff --git a/src/topology/sheaves/sheafify.lean b/src/topology/sheaves/sheafify.lean index acc3b40344654..46368756c0bce 100644 --- a/src/topology/sheaves/sheafify.lean +++ b/src/topology/sheaves/sheafify.lean @@ -76,7 +76,7 @@ def to_sheafify : F ⟶ F.sheafify.1 := The natural morphism from the stalk of the sheafification to the original stalk. In `sheafify_stalk_iso` we show this is an isomorphism. -/ -def stalk_to_fiber (x : X) : F.sheafify.1.stalk x ⟶ F.stalk x := +def stalk_to_fiber (x : X) : F.sheafify.presheaf.stalk x ⟶ F.stalk x := stalk_to_fiber (sheafify.is_locally_germ F) x lemma stalk_to_fiber_surjective (x : X) : function.surjective (F.stalk_to_fiber x) := @@ -121,7 +121,7 @@ end /-- The isomorphism betweeen a stalk of the sheafification and the original stalk. -/ -def sheafify_stalk_iso (x : X) : F.sheafify.1.stalk x ≅ F.stalk x := +def sheafify_stalk_iso (x : X) : F.sheafify.presheaf.stalk x ≅ F.stalk x := (equiv.of_bijective _ ⟨stalk_to_fiber_injective _ _, stalk_to_fiber_surjective _ _⟩).to_iso -- PROJECT functoriality, and that sheafification is the left adjoint of the forgetful functor. diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 7199ba7708cd2..5ff3e104a9df6 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -388,12 +388,12 @@ Let `F` be a sheaf valued in a concrete category, whose forgetful functor reflec preserves limits and filtered colimits. Then two sections who agree on every stalk must be equal. -/ lemma section_ext (F : sheaf C X) (U : opens X) (s t : F.1.obj (op U)) - (h : ∀ x : U, F.1.germ x s = F.1.germ x t) : + (h : ∀ x : U, F.presheaf.germ x s = F.presheaf.germ x t) : s = t := begin -- We use `germ_eq` and the axiom of choice, to pick for every point `x` a neighbourhood -- `V x`, such that the restrictions of `s` and `t` to `V x` coincide. - choose V m i₁ i₂ heq using λ x : U, F.1.germ_eq x.1 x.2 x.2 s t (h x), + 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₁, @@ -427,10 +427,10 @@ We claim that it suffices to find preimages *locally*. That is, for each `x : U` a neighborhood `V ≤ U` and a section `s : F.obj (op V))` such that `f.app (op V) s` and `t` agree on `V`. -/ lemma app_surjective_of_injective_of_locally_surjective {F G : sheaf C X} (f : F ⟶ G) - (U : opens X) (hinj : ∀ x : U, function.injective ((stalk_functor C x.1).map f)) + (U : opens X) (hinj : ∀ x : U, function.injective ((stalk_functor C x.1).map f.1)) (hsurj : ∀ (t) (x : U), ∃ (V : opens X) (m : x.1 ∈ V) (iVU : V ⟶ U) (s : F.1.obj (op V)), - f.app (op V) s = G.1.map iVU.op t) : - function.surjective (f.app (op U)) := + f.1.app (op V) s = G.1.map iVU.op t) : + function.surjective (f.1.app (op U)) := begin intro t, -- We use the axiom of choice to pick around each point `x` an open neighborhood `V` and a @@ -446,7 +446,7 @@ begin { use s, apply G.eq_of_locally_eq' V U iVU V_cover, intro x, - rw [← comp_apply, ← f.naturality, comp_apply, s_spec, heq] }, + rw [← comp_apply, ← f.1.naturality, comp_apply, s_spec, heq] }, { intros x y, -- What's left to show here is that the secions `sf` are compatible, i.e. they agree on -- the intersections `V x ⊓ V y`. We prove this by showing that all germs are equal. @@ -456,48 +456,48 @@ begin apply (hinj ⟨z, (iVU x).le ((inf_le_left : V x ⊓ V y ≤ V x) z.2)⟩), dsimp only, erw [stalk_functor_map_germ_apply, stalk_functor_map_germ_apply], - simp_rw [← comp_apply, f.naturality, comp_apply, heq, ← comp_apply, ← G.1.map_comp], + simp_rw [← comp_apply, f.1.naturality, comp_apply, heq, ← comp_apply, ← G.1.map_comp], refl } end lemma app_surjective_of_stalk_functor_map_bijective {F G : sheaf C X} (f : F ⟶ G) - (U : opens X) (h : ∀ x : U, function.bijective ((stalk_functor C x.val).map f)) : - function.surjective (f.app (op U)) := + (U : opens X) (h : ∀ x : U, function.bijective ((stalk_functor C x.val).map f.1)) : + function.surjective (f.1.app (op U)) := begin refine app_surjective_of_injective_of_locally_surjective f U (λ x, (h x).1) (λ t x, _), -- Now we need to prove our initial claim: That we can find preimages of `t` locally. -- Since `f` is surjective on stalks, we can find a preimage `s₀` of the germ of `t` at `x` - obtain ⟨s₀,hs₀⟩ := (h x).2 (G.1.germ x t), + obtain ⟨s₀,hs₀⟩ := (h x).2 (G.presheaf.germ x t), -- ... and this preimage must come from some section `s₁` defined on some open neighborhood `V₁` - obtain ⟨V₁,hxV₁,s₁,hs₁⟩ := F.1.germ_exist x.1 s₀, + obtain ⟨V₁,hxV₁,s₁,hs₁⟩ := F.presheaf.germ_exist x.1 s₀, subst hs₁, rename hs₀ hs₁, - erw stalk_functor_map_germ_apply V₁ ⟨x.1,hxV₁⟩ f s₁ at hs₁, + erw stalk_functor_map_germ_apply V₁ ⟨x.1,hxV₁⟩ f.1 s₁ at hs₁, -- Now, the germ of `f.app (op V₁) s₁` equals the germ of `t`, hence they must coincide on -- some open neighborhood `V₂`. - obtain ⟨V₂, hxV₂, iV₂V₁, iV₂U, heq⟩ := G.1.germ_eq x.1 hxV₁ x.2 _ _ hs₁, + obtain ⟨V₂, hxV₂, iV₂V₁, iV₂U, heq⟩ := G.presheaf.germ_eq x.1 hxV₁ x.2 _ _ hs₁, -- The restriction of `s₁` to that neighborhood is our desired local preimage. use [V₂, hxV₂, iV₂U, F.1.map iV₂V₁.op s₁], - rw [← comp_apply, f.naturality, comp_apply, heq], + rw [← comp_apply, f.1.naturality, comp_apply, heq], end lemma app_bijective_of_stalk_functor_map_bijective {F G : sheaf C X} (f : F ⟶ G) - (U : opens X) (h : ∀ x : U, function.bijective ((stalk_functor C x.val).map f)) : - function.bijective (f.app (op U)) := -⟨app_injective_of_stalk_functor_map_injective f U (λ x, (h x).1), + (U : opens X) (h : ∀ x : U, function.bijective ((stalk_functor C x.val).map f.1)) : + function.bijective (f.1.app (op U)) := +⟨app_injective_of_stalk_functor_map_injective f.1 U (λ x, (h x).1), app_surjective_of_stalk_functor_map_bijective f U h⟩ lemma app_is_iso_of_stalk_functor_map_iso {F G : sheaf C X} (f : F ⟶ G) (U : opens X) - [∀ x : U, is_iso ((stalk_functor C x.val).map f)] : is_iso (f.app (op U)) := + [∀ x : U, is_iso ((stalk_functor C x.val).map f.1)] : is_iso (f.1.app (op U)) := begin -- Since the forgetful functor of `C` reflects isomorphisms, it suffices to see that the -- underlying map between types is an isomorphism, i.e. bijective. - suffices : is_iso ((forget C).map (f.app (op U))), - { exactI is_iso_of_reflects_iso (f.app (op U)) (forget C) }, + suffices : is_iso ((forget C).map (f.1.app (op U))), + { exactI is_iso_of_reflects_iso (f.1.app (op U)) (forget C) }, rw is_iso_iff_bijective, apply app_bijective_of_stalk_functor_map_bijective, intro x, apply (is_iso_iff_bijective _).mp, - exact functor.map_is_iso (forget C) ((stalk_functor C x.1).map f) + exact functor.map_is_iso (forget C) ((stalk_functor C x.1).map f.1) end /-- @@ -507,15 +507,15 @@ isomorphisms, preserves limits and filtered colimits. Then if the stalk maps of -/ -- Making this an instance would cause a loop in typeclass resolution with `functor.map_is_iso` lemma is_iso_of_stalk_functor_map_iso {F G : sheaf C X} (f : F ⟶ G) - [∀ x : X, is_iso ((stalk_functor C x).map f)] : is_iso f := + [∀ x : X, is_iso ((stalk_functor C x).map f.1)] : is_iso f := begin -- Since the inclusion functor from sheaves to presheaves is fully faithful, it suffices to -- show that `f`, as a morphism between _presheaves_, is an isomorphism. suffices : is_iso ((sheaf.forget C X).map f), { exactI is_iso_of_fully_faithful (sheaf.forget C X) f }, -- We show that all components of `f` are isomorphisms. - suffices : ∀ U : (opens X)ᵒᵖ, is_iso (f.app U), - { exact @nat_iso.is_iso_of_is_iso_app _ _ _ _ F.1 G.1 f this, }, + suffices : ∀ U : (opens X)ᵒᵖ, is_iso (f.1.app U), + { exact @nat_iso.is_iso_of_is_iso_app _ _ _ _ F.1 G.1 f.1 this, }, intro U, induction U using opposite.rec, apply app_is_iso_of_stalk_functor_map_iso end @@ -526,11 +526,11 @@ isomorphisms, preserves limits and filtered colimits. Then a morphism `f : F ⟶ isomorphism if and only if all of its stalk maps are isomorphisms. -/ lemma is_iso_iff_stalk_functor_map_iso {F G : sheaf C X} (f : F ⟶ G) : - is_iso f ↔ ∀ x : X, is_iso ((stalk_functor C x).map f) := + is_iso f ↔ ∀ x : X, is_iso ((stalk_functor C x).map f.1) := begin split, { intros h x, resetI, - exact @functor.map_is_iso _ _ _ _ _ _ (stalk_functor C x) f + exact @functor.map_is_iso _ _ _ _ _ _ (stalk_functor C x) f.1 ((sheaf.forget C X).map_is_iso f) }, { intro h, exactI is_iso_of_stalk_functor_map_iso f } From 3eb459b351c203db64d11e0faebf54a4ece98cc1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 12:53:04 +0000 Subject: [PATCH 090/173] feat(ring_theory/ring_hom/finite): Finite ring morphisms are stable under base change (#15427) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/finiteness.lean | 18 ++++++++++++ src/ring_theory/ring_hom/finite.lean | 44 ++++++++++++++++++++++++++++ 2 files changed, 62 insertions(+) create mode 100644 src/ring_theory/ring_hom/finite.lean diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index b90466397fd90..3e70cf5d43427 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -144,6 +144,24 @@ end finite end module +instance module.finite.base_change [comm_semiring R] [semiring A] [algebra R A] + [add_comm_monoid M] [module R M] [h : module.finite R M] : + module.finite A (tensor_product R A M) := +begin + classical, + obtain ⟨s, hs⟩ := h.out, + refine ⟨⟨s.image (tensor_product.mk R A M 1), eq_top_iff.mpr $ λ x _, _⟩⟩, + apply tensor_product.induction_on x, + { exact zero_mem _ }, + { intros x y, + rw [finset.coe_image, ← submodule.span_span_of_tower R, submodule.span_image, hs, + submodule.map_top, linear_map.range_coe], + change _ ∈ submodule.span A (set.range $ tensor_product.mk R A M 1), + rw [← mul_one x, ← smul_eq_mul, ← tensor_product.smul_tmul'], + exact submodule.smul_mem _ x (submodule.subset_span $ set.mem_range_self y) }, + { exact λ _ _, submodule.add_mem _ } +end + instance module.finite.tensor_product [comm_semiring R] [add_comm_monoid M] [module R M] [add_comm_monoid N] [module R N] [hM : module.finite R M] [hN : module.finite R N] : module.finite R (tensor_product R M N) := diff --git a/src/ring_theory/ring_hom/finite.lean b/src/ring_theory/ring_hom/finite.lean new file mode 100644 index 0000000000000..fb5a12a3c455d --- /dev/null +++ b/src/ring_theory/ring_hom/finite.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2021 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import ring_theory.local_properties + +/-! + +# The meta properties of finite ring homomorphisms. + +-/ + +namespace ring_hom + +open_locale tensor_product + +open tensor_product algebra.tensor_product + +lemma finite_stable_under_composition : + stable_under_composition @finite := +by { introv R hf hg, exactI hg.comp hf } + +lemma finite_respects_iso : + respects_iso @finite := +begin + apply finite_stable_under_composition.respects_iso, + introsI, + exact finite.of_surjective _ e.to_equiv.surjective, +end + +lemma finite_stable_under_base_change : + stable_under_base_change @finite := +begin + classical, + introv R h, + resetI, + replace h : module.finite R T := by { convert h, ext, rw algebra.smul_def, refl }, + suffices : module.finite S (S ⊗[R] T), + { change module.finite _ _, convert this, ext, rw algebra.smul_def, refl }, + exactI infer_instance +end + +end ring_hom From 5759bac0a1ae58af477bca99735d0310840e5654 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 12:53:06 +0000 Subject: [PATCH 091/173] feat(ring_theory/tensor_product): A predicate for being the tensor product. (#15512) --- src/ring_theory/is_tensor_product.lean | 191 +++++++++++++++++++++++++ 1 file changed, 191 insertions(+) create mode 100644 src/ring_theory/is_tensor_product.lean diff --git a/src/ring_theory/is_tensor_product.lean b/src/ring_theory/is_tensor_product.lean new file mode 100644 index 0000000000000..d79925faff746 --- /dev/null +++ b/src/ring_theory/is_tensor_product.lean @@ -0,0 +1,191 @@ +/- +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.tensor_product + +/-! +# The characteristice predicate of tensor product + +## Main definitions + +- `is_tensor_product`: A predicate on `f : M₁ →ₗ[R] M₂ →ₗ[R] M` expressing that `f` realizes `M` as + the tensor product of `M₁ ⊗[R] M₂`. This is defined by requiring the lift `M₁ ⊗[R] M₂ → M` to be + bijective. +- `is_base_change`: A predicate on an `R`-algebra `S` and a map `f : M →ₗ[R] N` with `N` being a + `S`-module, expressing that `f` realizes `N` as the base change of `M` along `R → S`. + +## Main results +- `tensor_product.is_base_change`: `S ⊗[R] M` is the base change of `M` along `R → S`. + +-/ + +universes u v₁ v₂ v₃ v₄ + +open_locale tensor_product + +open tensor_product + +section is_tensor_product + +variables {R : Type*} [comm_ring R] +variables {M₁ M₂ M M' : Type*} +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'] +variable (f : M₁ →ₗ[R] M₂ →ₗ[R] M) +variables {N₁ N₂ N : Type*} [add_comm_monoid N₁] [add_comm_monoid N₂] [add_comm_monoid N] +variables [module R N₁] [module R N₂] [module R N] +variable {g : N₁ →ₗ[R] N₂ →ₗ[R] N} + +/-- +Given a bilinear map `f : M₁ →ₗ[R] M₂ →ₗ[R] M`, `is_tensor_product f` means that +`M` is the tensor product of `M₁` and `M₂` via `f`. +This is defined by requiring the lift `M₁ ⊗[R] M₂ → M` to be bijective. +-/ +def is_tensor_product : Prop := function.bijective (tensor_product.lift f) + +variables (R M N) {f} + +lemma tensor_product.is_tensor_product : is_tensor_product (tensor_product.mk R M N) := +begin + delta is_tensor_product, + convert_to function.bijective linear_map.id using 2, + { apply tensor_product.ext', simp }, + { exact function.bijective_id } +end + +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 + +@[simp] lemma is_tensor_product.equiv_to_linear_map (h : is_tensor_product f) : + h.equiv.to_linear_map = tensor_product.lift f := rfl + +@[simp] lemma is_tensor_product.equiv_symm_apply (h : is_tensor_product f) (x₁ : M₁) (x₂ : M₂) : + h.equiv.symm (f x₁ x₂) = x₁ ⊗ₜ x₂ := +begin + apply h.equiv.injective, + refine (h.equiv.apply_symm_apply _).trans _, + simp +end + +/-- If `M` is the tensor product of `M₁` and `M₂`, we may lift a bilinear map `M₁ →ₗ[R] M₂ →ₗ[R] M'` +to a `M →ₗ[R] M'`. -/ +noncomputable +def is_tensor_product.lift (h : is_tensor_product f) (f' : M₁ →ₗ[R] M₂ →ₗ[R] M') : M →ₗ[R] M' := +(tensor_product.lift f').comp h.equiv.symm.to_linear_map + +lemma is_tensor_product.lift_eq (h : is_tensor_product f) (f' : M₁ →ₗ[R] M₂ →ₗ[R] M') + (x₁ : M₁) (x₂ : M₂) : h.lift f' (f x₁ x₂) = f' x₁ x₂ := +begin + delta is_tensor_product.lift, + simp, +end + +/-- The tensor product of a pair of linear maps between modules. -/ +noncomputable +def is_tensor_product.map (hf : is_tensor_product f) (hg : is_tensor_product g) + (i₁ : M₁ →ₗ[R] N₁) (i₂ : M₂ →ₗ[R] N₂) : M →ₗ[R] N := +hg.equiv.to_linear_map.comp ((tensor_product.map i₁ i₂).comp hf.equiv.symm.to_linear_map) + +lemma is_tensor_product.map_eq (hf : is_tensor_product f) (hg : is_tensor_product g) + (i₁ : M₁ →ₗ[R] N₁) (i₂ : M₂ →ₗ[R] N₂) (x₁ : M₁) (x₂ : M₂) : + hf.map hg i₁ i₂ (f x₁ x₂) = g (i₁ x₁) (i₂ x₂) := +begin + delta is_tensor_product.map, + simp +end + +lemma is_tensor_product.induction_on (h : is_tensor_product f) {C : M → Prop} + (m : M) (h0 : C 0) (htmul : ∀ x y, C (f x y)) (hadd : ∀ x y, C x → C y → C (x + y)) : C m := +begin + rw ← h.equiv.right_inv m, + generalize : h.equiv.inv_fun m = y, + change C (tensor_product.lift f y), + induction y using tensor_product.induction_on, + { rwa map_zero }, + { rw tensor_product.lift.tmul, apply htmul }, + { rw map_add, apply hadd; assumption } +end + +end is_tensor_product + +section is_base_change + +variables {R M N : Type*} (S : Type*) [add_comm_monoid M] [add_comm_monoid N] [comm_ring R] +variables [comm_ring S] [algebra R S] [module R M] [module R N] [module S N] [is_scalar_tower R S N] +variables (f : M →ₗ[R] N) + +include f + +/-- Given an `R`-algebra `S` and an `R`-module `M`, an `S`-module `N` together with a map +`f : M →ₗ[R] N` is the base change of `M` to `S` if the map `S × M → N, (s, m) ↦ s • f m` is the +tensor product. -/ +def is_base_change : Prop := is_tensor_product +(((algebra.of_id S $ module.End S (M →ₗ[R] N)).to_linear_map.flip f).restrict_scalars R) + +variables {S f} (h : is_base_change S f) +variables {P Q : Type*} [add_comm_monoid P] [module R P] +variables[add_comm_monoid Q] [module R Q] [module S Q] [is_scalar_tower R S Q] + +/-- Suppose `f : M →ₗ[R] N` is the base change of `M` along `R → S`. Then any `R`-linear map from +`M` to an `S`-module factors thorugh `f`. -/ +noncomputable +def is_base_change.lift (g : M →ₗ[R] Q) : N →ₗ[S] Q := +{ map_smul' := λ r x, begin + let F := ((algebra.of_id S $ module.End S (M →ₗ[R] Q)) + .to_linear_map.flip g).restrict_scalars R, + have hF : ∀ (s : S) (m : M), h.lift F (s • f m) = s • g m := h.lift_eq F, + change h.lift F (r • x) = r • h.lift F x, + apply h.induction_on x, + { rw [smul_zero, map_zero, smul_zero] }, + { intros s m, + change h.lift F (r • s • f m) = r • h.lift F (s • f m), + rw [← mul_smul, hF, hF, mul_smul] }, + { intros x₁ x₂ e₁ e₂, rw [map_add, smul_add, map_add, smul_add, e₁, e₂] } + end, + ..(h.lift (((algebra.of_id S $ module.End S (M →ₗ[R] Q)) + .to_linear_map.flip g).restrict_scalars R)) } + +lemma is_base_change.lift_eq (g : M →ₗ[R] Q) (x : M) : h.lift g (f x) = g x := +begin + have hF : ∀ (s : S) (m : M), h.lift g (s • f m) = s • g m := h.lift_eq _, + convert hF 1 x; rw one_smul, +end + +lemma is_base_change.lift_comp (g : M →ₗ[R] Q) : ((h.lift g).restrict_scalars R).comp f = g := +linear_map.ext (h.lift_eq g) + +variables (R M N S) + +omit f + +lemma tensor_product.is_base_change : is_base_change S (tensor_product.mk R S M 1) := +begin + delta is_base_change, + convert tensor_product.is_tensor_product R S M using 1, + ext s x, + change s • 1 ⊗ₜ x = s ⊗ₜ x, + rw tensor_product.smul_tmul', + congr' 1, + exact mul_one _, +end + +variables {R M N S} + +/-- The base change of `M` along `R → S` is linearly equivalent to `S ⊗[R] M`. -/ +noncomputable +def is_base_change.equiv : S ⊗[R] M ≃ₗ[R] N := h.equiv + +lemma is_base_change.equiv_tmul (s : S) (m : M) : h.equiv (s ⊗ₜ m) = s • (f m) := +tensor_product.lift.tmul s m + +lemma is_base_change.equiv_symm_apply (m : M) : h.equiv.symm (f m) = 1 ⊗ₜ m := +by rw [h.equiv.symm_apply_eq, h.equiv_tmul, one_smul] + +end is_base_change From 194b2fb73f04f94fea5e205015c03b6e47b255cf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hu=E1=BB=B3nh=20Tr=E1=BA=A7n=20Khanh?= Date: Wed, 27 Jul 2022 14:34:49 +0000 Subject: [PATCH 092/173] feat(data/sym/card): Prove stars and bars (#11162) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this file we prove (in `sym.card_sym_eq_multichoose`) that `multichoose (card α) k` counts the number of multisets of cardinality `k` over a finite alphabet `α`. In conjunction with `nat.multichoose_eq`, which shows that `multichoose n k = choose (n + k - 1) k`, we immediately derive `card (sym α k) = (card α + k - 1).choose k`, which is the essence of the "stars and bars" technique in combinatorics. Co-authored-by: Stuart Presnell Co-authored-by: Huỳnh Trần Khanh Co-authored-by: YaelDillies Co-authored-by: Kyle Miller --- src/data/nat/choose/basic.lean | 11 ++- src/data/sym/card.lean | 122 ++++++++++++++++++++++++++++----- 2 files changed, 112 insertions(+), 21 deletions(-) diff --git a/src/data/nat/choose/basic.lean b/src/data/nat/choose/basic.lean index f649a0375cbc6..00e1d4f2d4670 100644 --- a/src/data/nat/choose/basic.lean +++ b/src/data/nat/choose/basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes, Bhavik Mehta +Authors: Chris Hughes, Bhavik Mehta, Stuart Presnell -/ import data.nat.factorial.basic @@ -22,11 +22,16 @@ requiring more imports). factorial. This is used to prove `nat.choose_le_pow` and variants. We provide similar statements for the ascending factorial. * `nat.multichoose`: whereas `choose` counts combinations, `multichoose` counts multicombinations. -* `nat.multichoose_eq` : a proof that `multichoose n k = (n + k - 1).choose k` +The fact that this is indeed the correct counting function for multisets is proved in +`sym.card_sym_eq_multichoose` in `data/sym/card`. +* `nat.multichoose_eq` : a proof that `multichoose n k = (n + k - 1).choose k`. +This is central to the "stars and bars" technique in informal mathematics, where we switch between +counting multisets of size `k` over an alphabet of size `n` to counting strings of `k` elements +("stars") separated by `n-1` dividers ("bars"). See `data/sym/card` for more detail. ## Tags -binomial coefficient, combination, multicombination +binomial coefficient, combination, multicombination, stars and bars -/ open_locale nat diff --git a/src/data/sym/card.lean b/src/data/sym/card.lean index 9d6a18ed295dd..09bba6bf1c06b 100644 --- a/src/data/sym/card.lean +++ b/src/data/sym/card.lean @@ -1,7 +1,7 @@ /- 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 +Authors: Yaël Dillies, Bhavik Mehta, Huỳnh Trần Khanh, Stuart Presnell -/ import algebra.big_operators.basic import data.finset.sym @@ -9,37 +9,123 @@ import data.finset.sym /-! # Stars and bars -In this file, we prove the case `n = 2` of stars and bars. +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 +`data/nat/choose/basic`, which shows that `multichoose n k = choose (n + k - 1) k`, +this is central to the "stars and bars" technique in combinatorics, where we switch between +counting multisets of size `k` over an alphabet of size `n` to counting strings of `k` elements +("stars") separated by `n-1` dividers ("bars"). ## Informal statement -If we have `n` objects to put in `k` boxes, we can do so in exactly `(n + k - 1).choose n` ways. +Many problems in mathematics are of the form of (or can be reduced to) putting `k` indistinguishable +objects into `n` distinguishable boxes; for example, the problem of finding natural numbers +`x1, ..., xn` whose sum is `k`. This is equivalent to forming a multiset of cardinality `k` from +an alphabet of cardinality `n` -- for each box `i ∈ [1, n]` the multiset contains as many copies +of `i` as there are items in the `i`th box. -## Formal statement +The "stars and bars" technique arises from another way of presenting the same problem. Instead of +putting `k` items into `n` boxes, we take a row of `k` items (the "stars") and separate them by +inserting `n-1` dividers (the "bars"). For example, the pattern `*|||**|*|` exhibits 4 items +distributed into 6 boxes -- note that any box, including the first and last, may be empty. +Such arrangements of `k` stars and `n-1` bars are in 1-1 correspondence with multisets of size `k` +over an alphabet of size `n`, and are counted by `choose (n + k - 1) k`. + +Note that this problem is one component of Gian-Carlo Rota's "Twelvefold Way" +https://en.wikipedia.org/wiki/Twelvefold_way -We can identify the `k` boxes with the elements of a fintype `α` of card `k`. Then placing `n` -elements in those boxes corresponds to choosing how many of each element of `α` appear in a multiset -of card `n`. `sym α n` being the subtype of `multiset α` of multisets of card `n`, writing stars -and bars using types gives -```lean --- TODO: this lemma is not yet proven -lemma stars_and_bars {α : Type*} [fintype α] (n : ℕ) : - card (sym α n) = (card α + n - 1).choose (card α) := sorry -``` +## Formal statement -## TODO +Here we generalise the alphabet to an arbitrary fintype `α`, and we use `sym α k` as the type of +multisets of size `k` over `α`. Thus the statement that these are counted by `multichoose` is: +`sym.card_sym_eq_multichoose : card (sym α k) = multichoose (card α) k` +while the "stars and bars" technique gives +`sym.card_sym_eq_choose : card (sym α k) = choose (card α + k - 1) k` -Prove the general case of stars and bars. ## Tags -stars and bars +stars and bars, multichoose +-/ + +open finset fintype function sum nat + +variables {α β : Type*} + +namespace sym + +section sym +variables (α) (n : ℕ) + +/-- +Over `fin n+1`, the multisets of size `k+1` containing `0` are equivalent to those of size `k`, +as demonstrated by respectively erasing or appending `0`. -/ +protected def E1 {n k : ℕ} : + {s : sym (fin n.succ) k.succ // ↑0 ∈ s} ≃ sym (fin n.succ) k := +{ to_fun := λ s, s.1.erase 0 s.2, + inv_fun := λ s, ⟨cons 0 s, mem_cons_self 0 s⟩, + left_inv := λ s, by simp, + right_inv := λ s, by simp } + +/-- +The multisets of size `k` over `fin n+2` not containing `0` +are equivalent to those of size `k` over `fin n+1`, +as demonstrated by respectively decrementing or incrementing every element of the multiset. +-/ +protected def E2 {n k : ℕ} : + {s : sym (fin n.succ.succ) k // ↑0 ∉ s} ≃ sym (fin n.succ) k := +{ to_fun := λ s, map (fin.pred_above 0) s.1, + inv_fun := λ s, ⟨map (fin.succ_above 0) s, + (mt mem_map.1) (not_exists.2 (λ t, (not_and.2 (λ _, (fin.succ_above_ne _ t)))))⟩, + left_inv := λ s, by + { obtain ⟨s, hs⟩ := s, + simp only [fin.zero_succ_above, map_map, comp_app], + nth_rewrite_rhs 0 ←(map_id' s), + refine sym.map_congr (λ v hv, _), + simp [fin.pred_above_zero (ne_of_mem_of_not_mem hv hs)] }, + right_inv := λ s, by + { 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] } } + +lemma card_sym_fin_eq_multichoose (n k : ℕ) : card (sym (fin n) k) = multichoose n k := +begin + apply @pincer_recursion (λ n k, card (sym (fin n) k) = multichoose n k), + { simp }, + { intros b, + induction b with b IHb, { simp }, + rw [multichoose_zero_succ, card_eq_zero_iff], + apply_instance }, + { intros x y h1 h2, + rw [multichoose_succ_succ, ←h1, ←h2, add_comm], + cases x, + { simp only [card_eq_zero_iff, nat.nat_zero_eq_zero, card_unique, self_eq_add_right], + apply_instance }, + rw ←card_sum, + refine fintype.card_congr (equiv.symm _), + apply (equiv.sum_congr sym.E1.symm sym.E2.symm).trans, + apply equiv.sum_compl }, +end + +/-- For any fintype `α` of cardinality `n`, `card (sym α k) = multichoose (card α) k` -/ +lemma card_sym_eq_multichoose (α : Type*) (k : ℕ) [fintype α] [fintype (sym α k)] : + card (sym α k) = multichoose (card α) k := +by { rw ←card_sym_fin_eq_multichoose, exact card_congr (equiv_congr (equiv_fin α)) } + +/-- The *stars and bars* lemma: the cardinality of `sym α k` is equal to +`nat.choose (card α + k - 1) k`. -/ +lemma card_sym_eq_choose {α : Type*} [fintype α] (k : ℕ) [fintype (sym α k)] : + card (sym α k) = (card α + k - 1).choose k := +by rw [card_sym_eq_multichoose, nat.multichoose_eq] -open finset fintype +end sym +end sym namespace sym2 -variables {α : Type*} [decidable_eq α] +variables [decidable_eq α] /-- The `diag` of `s : finset α` is sent on a finset of `sym2 α` of card `s.card`. -/ lemma card_image_diag (s : finset α) : (s.diag.image quotient.mk).card = s.card := From f1bf211c42c86ff85dd74aaf79b56d04728a0c1d Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 27 Jul 2022 14:34:50 +0000 Subject: [PATCH 093/173] feat(group_theory/complement): transversal for the transfer homomorphism (#14182) In order to make use of the transfer homomorphism, there is a messy computation that must be done. This computation involves a specific choice of transversal, which is constructed in this PR. Co-authored-by: Eric Wieser --- src/group_theory/complement.lean | 82 +++++++++++++++++++++++++++++++- 1 file changed, 80 insertions(+), 2 deletions(-) diff --git a/src/group_theory/complement.lean b/src/group_theory/complement.lean index 36948ff976a43..a9f02a2192708 100644 --- a/src/group_theory/complement.lean +++ b/src/group_theory/complement.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import group_theory.group_action.quotient -import group_theory.order_of_element +import data.zmod.quotient /-! # Complements @@ -20,6 +19,8 @@ In this file we define the complement of a subgroup. i.e. the set of all `S : set G` that contain exactly one element of each left coset of `T`. - `right_transversals S` where `S` is a subset of `G` is the set of all right-complements of `S`, i.e. the set of all `T : set G` that contain exactly one element of each right coset of `S`. +- `transfer_transversal H g` is a specific `left_transversal` of `H` that is used in the + computation of the transfer homomorphism evaluated at an element `g : G`. ## Main results @@ -406,3 +407,80 @@ begin end end subgroup + +namespace subgroup + +open equiv function mem_left_transversals mul_action mul_action.quotient zmod + +universe u + +variables {G : Type u} [group G] (H : subgroup G) (g : G) + +/-- Partition `G ⧸ H` into orbits of the action of `g : G`. -/ +noncomputable def quotient_equiv_sigma_zmod : G ⧸ H ≃ + Σ (q : orbit_rel.quotient (zpowers g) (G ⧸ H)), zmod (minimal_period ((•) g) q.out') := +(self_equiv_sigma_orbits (zpowers g) (G ⧸ H)).trans + (sigma_congr_right (λ q, orbit_zpowers_equiv g q.out')) + +lemma quotient_equiv_sigma_zmod_symm_apply + (q : orbit_rel.quotient (zpowers g) (G ⧸ H)) (k : zmod (minimal_period ((•) g) q.out')) : + (quotient_equiv_sigma_zmod H g).symm ⟨q, k⟩ = g ^ (k : ℤ) • q.out' := +rfl + +lemma quotient_equiv_sigma_zmod_apply (q : orbit_rel.quotient (zpowers g) (G ⧸ H)) (k : ℤ) : + quotient_equiv_sigma_zmod H g (g ^ k • q.out') = ⟨q, k⟩ := +by rw [apply_eq_iff_eq_symm_apply, quotient_equiv_sigma_zmod_symm_apply, + zmod.coe_int_cast, zpow_smul_mod_minimal_period] + +/-- The transfer transversal as a function. Given a `⟨g⟩`-orbit `q₀, g • q₀, ..., g ^ (m - 1) • q₀` + in `G ⧸ H`, an element `g ^ k • q₀` is mapped to `g ^ k • g₀` for a fixed choice of + representative `g₀` of `q₀`. -/ +noncomputable def transfer_function : G ⧸ H → G := +λ q, g ^ ((quotient_equiv_sigma_zmod H g q).2 : ℤ) * (quotient_equiv_sigma_zmod H g q).1.out'.out' + +lemma transfer_function_apply (q : G ⧸ H) : transfer_function H g q = + g ^ ((quotient_equiv_sigma_zmod H g q).2 : ℤ) * (quotient_equiv_sigma_zmod H g q).1.out'.out' := +rfl + +lemma coe_transfer_function (q : G ⧸ H) : ↑(transfer_function H g q) = q := +by rw [transfer_function_apply, ←smul_eq_mul, coe_smul_out', + ←quotient_equiv_sigma_zmod_symm_apply, sigma.eta, symm_apply_apply] + +/-- The transfer transversal as a set. Contains elements of the form `g ^ k • g₀` for fixed choices + of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ +def transfer_set : set G := +set.range (transfer_function H g) + +lemma mem_transfer_set (q : G ⧸ H) : transfer_function H g q ∈ transfer_set H g := +⟨q, rfl⟩ + +/-- The transfer transversal. Contains elements of the form `g ^ k • g₀` for fixed choices + of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ +def transfer_transversal : left_transversals (H : set G) := +⟨transfer_set H g, range_mem_left_transversals (coe_transfer_function H g)⟩ + +lemma transfer_transversal_apply (q : G ⧸ H) : + ↑(to_equiv (transfer_transversal H g).2 q) = transfer_function H g q := +to_equiv_apply (coe_transfer_function H g) q + +lemma transfer_transversal_apply' + (q : orbit_rel.quotient (zpowers g) (G ⧸ H)) (k : zmod (minimal_period ((•) g) q.out')) : + ↑(to_equiv (transfer_transversal H g).2 (g ^ (k : ℤ) • q.out')) = g ^ (k : ℤ) * q.out'.out' := +by rw [transfer_transversal_apply, transfer_function_apply, + ←quotient_equiv_sigma_zmod_symm_apply, apply_symm_apply] + +lemma transfer_transversal_apply'' + (q : orbit_rel.quotient (zpowers g) (G ⧸ H)) (k : zmod (minimal_period ((•) g) q.out')) : + ↑(to_equiv (g • transfer_transversal H g).2 (g ^ (k : ℤ) • q.out')) = + if k = 0 then g ^ minimal_period ((•) g) q.out' * q.out'.out' else g ^ (k : ℤ) * q.out'.out' := +begin + rw [smul_apply_eq_smul_apply_inv_smul, transfer_transversal_apply, transfer_function_apply, + ←mul_smul, ←zpow_neg_one, ←zpow_add, quotient_equiv_sigma_zmod_apply, smul_eq_mul, + ←mul_assoc, ←zpow_one_add, int.cast_add, int.cast_neg, int.cast_one, int_cast_cast, + cast_id', id.def, ←sub_eq_neg_add, cast_sub_one, add_sub_cancel'_right], + by_cases hk : k = 0, + { rw [if_pos hk, if_pos hk, zpow_coe_nat] }, + { rw [if_neg hk, if_neg hk] }, +end + +end subgroup From 2616567fde9294983b36ff9881645e9388a28e1a Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Jul 2022 14:34:52 +0000 Subject: [PATCH 094/173] feat(category_theory/preadditive): constructing a semiadditive structure from binary biproducts (#14445) --- .../preadditive/of_biproducts.lean | 95 +++++++++++++++++++ 1 file changed, 95 insertions(+) create mode 100644 src/category_theory/preadditive/of_biproducts.lean diff --git a/src/category_theory/preadditive/of_biproducts.lean b/src/category_theory/preadditive/of_biproducts.lean new file mode 100644 index 0000000000000..d1f0cc0508307 --- /dev/null +++ b/src/category_theory/preadditive/of_biproducts.lean @@ -0,0 +1,95 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import category_theory.limits.shapes.biproducts +import group_theory.eckmann_hilton + +/-! +# Constructing a semiadditive structure from binary biproducts + +We show that any category with zero morphisms and binary biproducts is enriched over the category +of commutative monoids. + +-/ + +noncomputable theory + +universes v u + +open category_theory +open category_theory.limits + +namespace category_theory.semiadditive_of_binary_biproducts +variables {C : Type u} [category.{v} C] [has_zero_morphisms C] [has_binary_biproducts C] + +section +variables (X Y : C) + +/-- `f +ₗ g` is the composite `X ⟶ Y ⊞ Y ⟶ Y`, where the first map is `(f, g)` and the second map + is `(𝟙 𝟙)`. -/ +@[simp] def left_add (f g : X ⟶ Y) : X ⟶ Y := +biprod.lift f g ≫ biprod.desc (𝟙 Y) (𝟙 Y) + +/-- `f +ᵣ g` is the composite `X ⟶ X ⊞ X ⟶ Y`, where the first map is `(𝟙, 𝟙)` and the second map + is `(f g)`. -/ +@[simp] def right_add (f g : X ⟶ Y) : X ⟶ Y := +biprod.lift (𝟙 X) (𝟙 X) ≫ biprod.desc f g + +local infixr ` +ₗ `:65 := left_add X Y +local infixr ` +ᵣ `:65 := right_add X Y + +lemma is_unital_left_add : eckmann_hilton.is_unital (+ₗ) 0 := +⟨⟨λ f, by simp [show biprod.lift (0 : X ⟶ Y) f = f ≫ biprod.inr, by ext; simp]⟩, + ⟨λ f, by simp [show biprod.lift f (0 : X ⟶ Y) = f ≫ biprod.inl, by ext; simp]⟩⟩ + +lemma is_unital_right_add : eckmann_hilton.is_unital (+ᵣ) 0 := +⟨⟨λ f, by simp [show biprod.desc (0 : X ⟶ Y) f = biprod.snd ≫ f, by ext; simp]⟩, + ⟨λ f, by simp [show biprod.desc f (0 : X ⟶ Y) = biprod.fst ≫ f, by ext; simp]⟩⟩ + +lemma distrib (f g h k : X ⟶ Y) : (f +ᵣ g) +ₗ (h +ᵣ k) = (f +ₗ h) +ᵣ (g +ₗ k) := +begin + let diag : X ⊞ X ⟶ Y ⊞ Y := biprod.lift (biprod.desc f g) (biprod.desc h k), + have hd₁ : biprod.inl ≫ diag = biprod.lift f h := by { ext; simp }, + have hd₂ : biprod.inr ≫ diag = biprod.lift g k := by { ext; simp }, + have h₁ : biprod.lift (f +ᵣ g) (h +ᵣ k) = biprod.lift (𝟙 X) (𝟙 X) ≫ diag := by { ext; simp }, + have h₂ : diag ≫ biprod.desc (𝟙 Y) (𝟙 Y) = biprod.desc (f +ₗ h) (g +ₗ k), + { ext; simp [reassoc_of hd₁, reassoc_of hd₂] }, + rw [left_add, h₁, category.assoc, h₂, right_add] +end + +/-- In a category with binary biproducts, the morphisms form a commutative monoid. -/ +def add_comm_monoid_hom_of_has_binary_biproducts : add_comm_monoid (X ⟶ Y) := +{ add := (+ᵣ), + add_assoc := (eckmann_hilton.mul_assoc (is_unital_left_add X Y) + (is_unital_right_add X Y) (distrib X Y)).assoc, + zero := 0, + zero_add := (is_unital_right_add X Y).left_id, + add_zero := (is_unital_right_add X Y).right_id, + add_comm := (eckmann_hilton.mul_comm (is_unital_left_add X Y) + (is_unital_right_add X Y) (distrib X Y)).comm } + +end + +section +variables {X Y Z : C} + +local attribute [instance] add_comm_monoid_hom_of_has_binary_biproducts + +lemma add_eq_right_addition (f g : X ⟶ Y) : f + g = biprod.lift (𝟙 X) (𝟙 X) ≫ biprod.desc f g := +rfl + +lemma add_eq_left_addition (f g : X ⟶ Y) : f + g = biprod.lift f g ≫ biprod.desc (𝟙 Y) (𝟙 Y) := +congr_fun₂ + (eckmann_hilton.mul (is_unital_left_add X Y) (is_unital_right_add X Y) (distrib X Y)).symm f g + +lemma add_comp (f g : X ⟶ Y) (h : Y ⟶ Z) : (f + g) ≫ h = f ≫ h + g ≫ h := +by { simp only [add_eq_right_addition, category.assoc], congr, ext; simp } + +lemma comp_add (f : X ⟶ Y) (g h : Y ⟶ Z) : f ≫ (g + h) = f ≫ g + f ≫ h := +by { simp only [add_eq_left_addition, ← category.assoc], congr, ext; simp } + +end + +end category_theory.semiadditive_of_binary_biproducts From 9a82b72a03e4be75e3873f8c3837e7cf327eefc8 Mon Sep 17 00:00:00 2001 From: MichaelBlyth <39394803+Michael14071407@users.noreply.github.com> Date: Wed, 27 Jul 2022 14:34:54 +0000 Subject: [PATCH 095/173] feat(linear_algebra/basis): shows the lattice of submodules of a module over a division ring is atomistic (#14883) This PR shows that the lattice of submodules of a module over a division ring is atomistic, and that the atoms of this lattice are the spans of nonzero elements. --- src/linear_algebra/basis.lean | 48 +++++++++++++++++++++++++++++++++++ src/linear_algebra/span.lean | 12 +++++++++ 2 files changed, 60 insertions(+) diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 030f264424770..7305b4da13c7c 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -1225,6 +1225,54 @@ 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)⟩ +section atoms_of_submodule_lattice + +variables {K V} + +/-- For a module over a division ring, the span of a nonzero element is an atom of the +lattice of submodules. -/ +lemma nonzero_span_atom (v : V) (hv : v ≠ 0) : is_atom (span K {v} : submodule K V) := +begin + split, + { rw submodule.ne_bot_iff, exact ⟨v, ⟨mem_span_singleton_self v, hv⟩⟩ }, + { intros T hT, by_contra, apply hT.2, + change (span K {v}) ≤ T, + simp_rw [span_singleton_le_iff_mem, ← ne.def, submodule.ne_bot_iff] at *, + rcases h with ⟨s, ⟨hs, hz⟩⟩, + cases (mem_span_singleton.1 (hT.1 hs)) with a ha, + have h : a ≠ 0, by { intro h, rw [h, zero_smul] at ha, exact hz ha.symm }, + apply_fun (λ x, a⁻¹ • x) at ha, + simp_rw [← mul_smul, inv_mul_cancel h, one_smul, ha] at *, exact smul_mem T _ hs}, +end + +/-- The atoms of the lattice of submodules of a module over a division ring are the +submodules equal to the span of a nonzero element of the module. -/ +lemma atom_iff_nonzero_span (W : submodule K V) : + is_atom W ↔ ∃ (v : V) (hv : v ≠ 0), W = span K {v} := +begin + refine ⟨λ h, _, λ h, _ ⟩, + { cases h with hbot h, + rcases ((submodule.ne_bot_iff W).1 hbot) with ⟨v, ⟨hW, hv⟩⟩, + refine ⟨v, ⟨hv, _⟩⟩, + by_contra heq, + specialize h (span K {v}), + rw [span_singleton_eq_bot, lt_iff_le_and_ne] at h, + exact hv (h ⟨(span_singleton_le_iff_mem v W).2 hW, ne.symm heq⟩) }, + { rcases h with ⟨v, ⟨hv, rfl⟩⟩, exact nonzero_span_atom v hv }, +end + +/-- The lattice of submodules of a module over a division ring is atomistic. -/ +instance : is_atomistic (submodule K V) := +{ eq_Sup_atoms := + begin + intro W, + use {T : submodule K V | ∃ (v : V) (hv : v ∈ W) (hz : v ≠ 0), T = span K {v}}, + refine ⟨submodule_eq_Sup_le_nonzero_spans W, _⟩, + rintros _ ⟨w, ⟨_, ⟨hw, rfl⟩⟩⟩, exact nonzero_span_atom w hw + end } + +end atoms_of_submodule_lattice + end division_ring section field diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index ab875fafaa65e..fd7d3f49fe80b 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -596,6 +596,18 @@ instance : is_compactly_generated (submodule R M) := apply singleton_span_is_compact_element, end, by rw [Sup_eq_supr, supr_image, ←span_eq_supr_of_singleton_spans, span_eq]⟩⟩⟩ +/-- A submodule is equal to the supremum of the spans of the submodule's nonzero elements. -/ +lemma submodule_eq_Sup_le_nonzero_spans (p : submodule R M) : + p = Sup {T : submodule R M | ∃ (m : M) (hm : m ∈ p) (hz : m ≠ 0), T = span R {m}} := +begin + let S := {T : submodule R M | ∃ (m : M) (hm : m ∈ p) (hz : m ≠ 0), T = span R {m}}, + apply le_antisymm, + { intros m hm, by_cases h : m = 0, + { rw h, simp }, + { exact @le_Sup _ _ S _ ⟨m, ⟨hm, ⟨h, rfl⟩⟩⟩ m (mem_span_singleton_self m) } }, + { rw Sup_le_iff, rintros S ⟨_, ⟨_, ⟨_, rfl⟩⟩⟩, rwa span_singleton_le_iff_mem } +end + lemma lt_sup_iff_not_mem {I : submodule R M} {a : M} : I < I ⊔ (R ∙ a) ↔ a ∉ I := begin split, From 41844de039e97e950db29760a6db90419fcec0f1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 14:34:56 +0000 Subject: [PATCH 096/173] feat(category_theory/limits/construction): Construct finite limits from terminal objects and pullbacks (#14948) Also provides the dual version, and also in terms of `preserves_limit`. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- .../limits/constructions/binary_products.lean | 160 +++++++++++------- .../limits/constructions/equalizers.lean | 155 ++++++++++++++++- .../limits_of_products_and_equalizers.lean | 49 ++++++ .../limits/preserves/shapes/pullbacks.lean | 8 + 4 files changed, 304 insertions(+), 68 deletions(-) diff --git a/src/category_theory/limits/constructions/binary_products.lean b/src/category_theory/limits/constructions/binary_products.lean index 665b3856b9392..bd282853b05af 100644 --- a/src/category_theory/limits/constructions/binary_products.lean +++ b/src/category_theory/limits/constructions/binary_products.lean @@ -6,6 +6,8 @@ Authors: Bhavik Mehta, Andrew Yang import category_theory.limits.shapes.terminal import category_theory.limits.shapes.pullbacks import category_theory.limits.shapes.binary_products +import category_theory.limits.preserves.shapes.pullbacks +import category_theory.limits.preserves.shapes.terminal /-! # Constructing binary product from pullbacks and terminal object. @@ -16,34 +18,44 @@ has pullbacks and a terminal object, then it has binary products. We also provide the dual. -/ -universes v u +universes v v' u u' open category_theory category_theory.category category_theory.limits -variables {C : Type u} [category.{v} C] +variables {C : Type u} [category.{v} C] {D : Type u'} [category.{v'} D] (F : C ⥤ D) + +/-- If a span is the pullback span over the terminal object, then it is a binary product. -/ +def is_binary_product_of_is_terminal_is_pullback (F : discrete walking_pair ⥤ C) (c : cone F) + {X : C} (hX : is_terminal X) + (f : F.obj ⟨walking_pair.left⟩ ⟶ X) (g : F.obj ⟨walking_pair.right⟩ ⟶ X) + (hc : is_limit (pullback_cone.mk (c.π.app ⟨walking_pair.left⟩) (c.π.app ⟨walking_pair.right⟩ : _) + $ hX.hom_ext (_ ≫ f) (_ ≫ g))) : is_limit c := +{ lift := λ s, hc.lift + (pullback_cone.mk (s.π.app ⟨walking_pair.left⟩) (s.π.app ⟨walking_pair.right⟩) + (hX.hom_ext _ _)), + fac' := λ s j, discrete.cases_on j + (λ j, walking_pair.cases_on j (hc.fac _ walking_cospan.left) (hc.fac _ walking_cospan.right)), + uniq' := λ s m J, + begin + let c' := pullback_cone.mk + (m ≫ c.π.app ⟨walking_pair.left⟩) (m ≫ c.π.app ⟨walking_pair.right⟩ : _) + (hX.hom_ext (_ ≫ f) (_ ≫ g)), + rw [←J, ←J], + apply hc.hom_ext, + rintro (_|(_|_)); simp only [pullback_cone.mk_π_app_one, pullback_cone.mk_π_app], + exacts [(category.assoc _ _ _).symm.trans (hc.fac_assoc c' walking_cospan.left f).symm, + (hc.fac c' walking_cospan.left).symm, (hc.fac c' walking_cospan.right).symm] + end } /-- The pullback over the terminal object is the product -/ def is_product_of_is_terminal_is_pullback {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (h : W ⟶ X) (k : W ⟶ Y) (H₁ : is_terminal Z) (H₂ : is_limit (pullback_cone.mk _ _ (show h ≫ f = k ≫ g, from H₁.hom_ext _ _))) : is_limit (binary_fan.mk h k) := -{ lift := λ c, H₂.lift (pullback_cone.mk - (c.π.app ⟨walking_pair.left⟩) (c.π.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)), - fac' := λ c j, - begin - cases j, - convert H₂.fac (pullback_cone.mk (c.π.app ⟨walking_pair.left⟩) - (c.π.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)) (some j) using 1, - rcases j; refl, - end, - uniq' := λ c m hm, - begin - apply pullback_cone.is_limit.hom_ext H₂, - { exact (hm ⟨walking_pair.left⟩).trans (H₂.fac (pullback_cone.mk (c.π.app ⟨walking_pair.left⟩) - (c.π.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)) walking_cospan.left).symm }, - { exact (hm ⟨walking_pair.right⟩).trans (H₂.fac (pullback_cone.mk (c.π.app ⟨walking_pair.left⟩) - (c.π.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)) walking_cospan.right).symm }, - end } +begin + apply is_binary_product_of_is_terminal_is_pullback _ _ H₁, + exact H₂ +end /-- The product is the pullback over the terminal object. -/ def is_pullback_of_is_terminal_is_product {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (h : W ⟶ X) @@ -71,19 +83,8 @@ noncomputable def limit_cone_of_terminal_and_pullbacks [has_terminal C] [has_pul (terminal.from (F.obj ⟨walking_pair.right⟩)), π := discrete.nat_trans (λ x, discrete.cases_on x (λ x, walking_pair.cases_on x pullback.fst pullback.snd)) }, - is_limit := - { lift := λ c, pullback.lift ((c.π).app ⟨walking_pair.left⟩) - ((c.π).app ⟨walking_pair.right⟩) - (subsingleton.elim _ _), - fac' := λ s c, discrete.cases_on c - (λ c, walking_pair.cases_on c (limit.lift_π _ _) (limit.lift_π _ _)), - uniq' := λ s m J, - begin - rw [←J, ←J], - ext; - rw limit.lift_π; - refl - end } } + is_limit := is_binary_product_of_is_terminal_is_pullback + F _ terminal_is_terminal _ _ (pullback_is_pullback _ _) } variable (C) @@ -94,6 +95,22 @@ lemma has_binary_products_of_terminal_and_pullbacks has_binary_products C := { has_limit := λ F, has_limit.mk (limit_cone_of_terminal_and_pullbacks F) } +variable {C} + +/-- A functor that preserves terminal objects and pullbacks preserves binary products. -/ +noncomputable +def preserves_binary_products_of_preserves_terminal_and_pullbacks + [has_terminal C] [has_pullbacks C] + [preserves_limits_of_shape (discrete.{0} pempty) F] + [preserves_limits_of_shape walking_cospan F] : + preserves_limits_of_shape (discrete walking_pair) F := +⟨λ K, preserves_limit_of_preserves_limit_cone (limit_cone_of_terminal_and_pullbacks K).2 +begin + apply is_binary_product_of_is_terminal_is_pullback _ _ + (is_limit_of_has_terminal_of_preserves_limit F), + apply is_limit_of_has_pullback_of_preserves_limit, +end⟩ + /-- In a category with a terminal object and pullbacks, a product of objects `X` and `Y` is isomorphic to a pullback. -/ noncomputable @@ -101,30 +118,40 @@ def prod_iso_pullback [has_terminal C] [has_pullbacks C] (X Y : C) [has_binary_p X ⨯ Y ≅ pullback (terminal.from X) (terminal.from Y) := limit.iso_limit_cone (limit_cone_of_terminal_and_pullbacks _) -variable {C} +/-- If a cospan is the pushout cospan under the initial object, then it is a binary coproduct. -/ +def is_binary_coproduct_of_is_initial_is_pushout (F : discrete walking_pair ⥤ C) (c : cocone F) + {X : C} (hX : is_initial X) + (f : X ⟶ F.obj ⟨walking_pair.left⟩) (g : X ⟶ F.obj ⟨walking_pair.right⟩) + (hc : is_colimit (pushout_cocone.mk (c.ι.app ⟨walking_pair.left⟩) + (c.ι.app ⟨walking_pair.right⟩ : _) $ hX.hom_ext (f ≫ _) (g ≫ _))) : is_colimit c := +{ desc := λ s, hc.desc + (pushout_cocone.mk (s.ι.app ⟨walking_pair.left⟩) (s.ι.app ⟨walking_pair.right⟩) + (hX.hom_ext _ _)), + fac' := λ s j, discrete.cases_on j + (λ j, walking_pair.cases_on j (hc.fac _ walking_span.left) (hc.fac _ walking_span.right)), + uniq' := λ s m J, + begin + let c' := pushout_cocone.mk + (c.ι.app ⟨walking_pair.left⟩ ≫ m) (c.ι.app ⟨walking_pair.right⟩ ≫ m) + (hX.hom_ext (f ≫ _) (g ≫ _)), + rw [←J, ←J], + apply hc.hom_ext, + rintro (_|(_|_)); + simp only [pushout_cocone.mk_ι_app_zero, pushout_cocone.mk_ι_app, category.assoc], + congr' 1, + exacts [(hc.fac c' walking_span.left).symm, + (hc.fac c' walking_span.left).symm, (hc.fac c' walking_span.right).symm] + end } /-- The pushout under the initial object is the coproduct -/ def is_coproduct_of_is_initial_is_pushout {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (h : W ⟶ X) (k : W ⟶ Y) (H₁ : is_initial W) (H₂ : is_colimit (pushout_cocone.mk _ _ (show h ≫ f = k ≫ g, from H₁.hom_ext _ _))) : is_colimit (binary_cofan.mk f g) := -{ desc := λ c, H₂.desc (pushout_cocone.mk - (c.ι.app ⟨walking_pair.left⟩) (c.ι.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)), - fac' := λ c j, - begin - cases j, - convert H₂.fac (pushout_cocone.mk (c.ι.app ⟨walking_pair.left⟩) (c.ι.app ⟨walking_pair.right⟩) - (H₁.hom_ext _ _)) (some j) using 1, - cases j; refl - end, - uniq' := λ c m hm, - begin - apply pushout_cocone.is_colimit.hom_ext H₂, - { exact (hm ⟨walking_pair.left⟩).trans (H₂.fac (pushout_cocone.mk (c.ι.app ⟨walking_pair.left⟩) - (c.ι.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)) walking_cospan.left).symm }, - { exact (hm ⟨walking_pair.right⟩).trans (H₂.fac (pushout_cocone.mk (c.ι.app ⟨walking_pair.left⟩) - (c.ι.app ⟨walking_pair.right⟩) (H₁.hom_ext _ _)) walking_cospan.right).symm }, - end } +begin + apply is_binary_coproduct_of_is_initial_is_pushout _ _ H₁, + exact H₂ +end /-- The coproduct is the pushout under the initial object. -/ def is_pushout_of_is_initial_is_coproduct {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (h : W ⟶ X) @@ -152,19 +179,9 @@ noncomputable def colimit_cocone_of_initial_and_pushouts [has_initial C] [has_pu (initial.to (F.obj ⟨walking_pair.right⟩)), ι := discrete.nat_trans (λ x, discrete.cases_on x (λ x, walking_pair.cases_on x pushout.inl pushout.inr)) }, - is_colimit := - { desc := λ c, pushout.desc (c.ι.app ⟨walking_pair.left⟩) - (c.ι.app ⟨walking_pair.right⟩) - (subsingleton.elim _ _), - fac' := λ s c, discrete.cases_on c - (λ c, walking_pair.cases_on c (colimit.ι_desc _ _) (colimit.ι_desc _ _)), - uniq' := λ s m J, - begin - rw [←J, ←J], - ext; - rw colimit.ι_desc; - refl - end } } + is_colimit := is_binary_coproduct_of_is_initial_is_pushout + F _ initial_is_initial _ _ (pushout_is_pushout _ _) } + variable (C) @@ -175,6 +192,23 @@ lemma has_binary_coproducts_of_initial_and_pushouts has_binary_coproducts C := { has_colimit := λ F, has_colimit.mk (colimit_cocone_of_initial_and_pushouts F) } +variable {C} + +/-- A functor that preserves initial objects and pushouts preserves binary coproducts. -/ +noncomputable +def preserves_binary_coproducts_of_preserves_initial_and_pushouts + [has_initial C] [has_pushouts C] + [preserves_colimits_of_shape (discrete.{0} pempty) F] + [preserves_colimits_of_shape walking_span F] : + preserves_colimits_of_shape (discrete walking_pair) F := +⟨λ K, preserves_colimit_of_preserves_colimit_cocone (colimit_cocone_of_initial_and_pushouts K).2 +begin + apply is_binary_coproduct_of_is_initial_is_pushout _ _ + (is_colimit_of_has_initial_of_preserves_colimit F), + apply is_colimit_of_has_pushout_of_preserves_colimit, +end⟩ + + /-- In a category with an initial object and pushouts, a coproduct of objects `X` and `Y` is isomorphic to a pushout. -/ noncomputable diff --git a/src/category_theory/limits/constructions/equalizers.lean b/src/category_theory/limits/constructions/equalizers.lean index 7cc0d326d6694..eb32cf33463d3 100644 --- a/src/category_theory/limits/constructions/equalizers.lean +++ b/src/category_theory/limits/constructions/equalizers.lean @@ -1,33 +1,38 @@ /- Copyright (c) 2020 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Bhavik Mehta +Authors: Bhavik Mehta, Andrew Yang -/ import category_theory.limits.shapes.equalizers import category_theory.limits.shapes.binary_products import category_theory.limits.shapes.pullbacks +import category_theory.limits.preserves.shapes.pullbacks +import category_theory.limits.preserves.shapes.binary_products /-! # Constructing equalizers from pullbacks and binary products. If a category has pullbacks and binary products, then it has equalizers. -TODO: provide the dual result. +TODO: generalize universe -/ noncomputable theory -universes v u +universes v v' u u' open category_theory category_theory.category namespace category_theory.limits -variables {C : Type u} [category.{v} C] [has_binary_products C] [has_pullbacks C] +variables {C : Type u} [category.{v} C] +variables {D : Type u'} [category.{v'} D] (G : C ⥤ D) -- We hide the "implementation details" inside a namespace namespace has_equalizers_of_pullbacks_and_binary_products +variables [has_binary_products C] [has_pullbacks C] + /-- Define the equalizing object -/ @[reducible] def construct_equalizer (F : walking_parallel_pair ⥤ C) : C := @@ -76,10 +81,150 @@ end has_equalizers_of_pullbacks_and_binary_products open has_equalizers_of_pullbacks_and_binary_products /-- Any category with pullbacks and binary products, has equalizers. -/ -- This is not an instance, as it is not always how one wants to construct equalizers! -lemma has_equalizers_of_pullbacks_and_binary_products : +lemma has_equalizers_of_pullbacks_and_binary_products [has_binary_products C] [has_pullbacks C] : has_equalizers C := { has_limit := λ F, has_limit.mk { cone := equalizer_cone F, is_limit := equalizer_cone_is_limit F } } +local attribute[instance] has_pullback_of_preserves_pullback + +/-- A functor that preserves pullbacks and binary products also presrves equalizers. -/ +def preserves_equalizers_of_pullbacks_and_binary_products + [has_binary_products C] [has_pullbacks C] + [preserves_limits_of_shape (discrete walking_pair) G] + [preserves_limits_of_shape walking_cospan G] : + preserves_limits_of_shape walking_parallel_pair G := +⟨λ K, preserves_limit_of_preserves_limit_cone (equalizer_cone_is_limit K) $ +{ lift := λ c, begin + refine pullback.lift _ _ _ ≫ (@@preserves_pullback.iso _ _ _ _ _ _ _ _).inv, + { exact c.π.app walking_parallel_pair.zero }, + { exact c.π.app walking_parallel_pair.zero }, + apply (map_is_limit_of_preserves_of_is_limit G _ _ (prod_is_prod _ _)).hom_ext, + swap, apply_instance, + rintro (_|_), + { simp only [category.assoc, ← G.map_comp, prod.lift_fst, + binary_fan.π_app_left, binary_fan.mk_fst], }, + { simp only [binary_fan.π_app_right, binary_fan.mk_snd, + category.assoc, ← G.map_comp, prod.lift_snd], + exact (c.π.naturality (walking_parallel_pair_hom.left)).symm.trans + (c.π.naturality (walking_parallel_pair_hom.right)) }, + end, + fac' := λ c j, begin + rcases j with (_|_); + simp only [category.comp_id, preserves_pullback.iso_inv_fst, cone.of_fork_π, G.map_comp, + preserves_pullback.iso_inv_fst_assoc, functor.map_cone_π_app, eq_to_hom_refl, + category.assoc, fork.of_ι_π_app, pullback.lift_fst, pullback.lift_fst_assoc], + exact (c.π.naturality (walking_parallel_pair_hom.left)).symm.trans (category.id_comp _) + end, + uniq' := λ s m h, begin + rw iso.eq_comp_inv, + have := h walking_parallel_pair.zero, + dsimp [equalizer_cone] at this, + ext; simp only [preserves_pullback.iso_hom_snd, + category.assoc, preserves_pullback.iso_hom_fst, pullback.lift_fst, pullback.lift_snd, + category.comp_id, ← pullback_fst_eq_pullback_snd, ← this], + end }⟩ + + +-- We hide the "implementation details" inside a namespace +namespace has_coequalizers_of_pushouts_and_binary_coproducts + +variables [has_binary_coproducts C] [has_pushouts C] + +/-- Define the equalizing object -/ +@[reducible] +def construct_coequalizer (F : walking_parallel_pair ⥤ C) : C := +pushout (coprod.desc (𝟙 _) (F.map walking_parallel_pair_hom.left)) + (coprod.desc (𝟙 _) (F.map walking_parallel_pair_hom.right)) + +/-- Define the equalizing morphism -/ +abbreviation pushout_inl (F : walking_parallel_pair ⥤ C) : + F.obj walking_parallel_pair.one ⟶ construct_coequalizer F := +pushout.inl + +lemma pushout_inl_eq_pushout_inr (F : walking_parallel_pair ⥤ C) : + pushout_inl F = pushout.inr := +by convert limits.coprod.inl ≫= pushout.condition; simp + +/-- Define the equalizing cocone -/ +@[reducible] +def coequalizer_cocone (F : walking_parallel_pair ⥤ C) : cocone F := +cocone.of_cofork + (cofork.of_π (pushout_inl F) + (begin + conv_rhs { rw pushout_inl_eq_pushout_inr, }, + convert limits.coprod.inr ≫= pushout.condition using 1; simp + end)) + +/-- Show the equalizing cocone is a colimit -/ +def coequalizer_cocone_is_colimit (F : walking_parallel_pair ⥤ C) : + is_colimit (coequalizer_cocone F) := +{ desc := + begin + intro c, apply pushout.desc (c.ι.app _) (c.ι.app _), + apply colimit.hom_ext, + rintro (_ | _); simp + end, + fac' := by rintros c (_ | _); simp, + uniq' := + begin + intros c _ J, + have J1 : pushout_inl F ≫ m = c.ι.app walking_parallel_pair.one := + by simpa using J walking_parallel_pair.one, + apply pushout.hom_ext, + { rw colimit.ι_desc, exact J1 }, + { rw [colimit.ι_desc, ← pushout_inl_eq_pushout_inr], exact J1 } + end } + +end has_coequalizers_of_pushouts_and_binary_coproducts + +open has_coequalizers_of_pushouts_and_binary_coproducts +/-- Any category with pullbacks and binary products, has equalizers. -/ +-- This is not an instance, as it is not always how one wants to construct equalizers! +lemma has_coequalizers_of_pushouts_and_binary_coproducts + [has_binary_coproducts C] [has_pushouts C] : has_coequalizers C := +{ has_colimit := λ F, has_colimit.mk + { cocone := coequalizer_cocone F, + is_colimit := coequalizer_cocone_is_colimit F } } + +local attribute[instance] has_pushout_of_preserves_pushout + +/-- A functor that preserves pushouts and binary coproducts also presrves coequalizers. -/ +def preserves_coequalizers_of_pushouts_and_binary_coproducts + [has_binary_coproducts C] [has_pushouts C] + [preserves_colimits_of_shape (discrete walking_pair) G] + [preserves_colimits_of_shape walking_span G] : + preserves_colimits_of_shape walking_parallel_pair G := +⟨λ K, preserves_colimit_of_preserves_colimit_cocone (coequalizer_cocone_is_colimit K) $ +{ desc := λ c, begin + refine (@@preserves_pushout.iso _ _ _ _ _ _ _ _).inv ≫ pushout.desc _ _ _, + { exact c.ι.app walking_parallel_pair.one }, + { exact c.ι.app walking_parallel_pair.one }, + apply (map_is_colimit_of_preserves_of_is_colimit G _ _ (coprod_is_coprod _ _)).hom_ext, + swap, apply_instance, + rintro (_|_), + { simp only [binary_cofan.ι_app_left, binary_cofan.mk_inl, category.assoc, ← G.map_comp_assoc, + coprod.inl_desc] }, + { simp only [binary_cofan.ι_app_right, binary_cofan.mk_inr, category.assoc, ← G.map_comp_assoc, + coprod.inr_desc], + exact (c.ι.naturality walking_parallel_pair_hom.left).trans + (c.ι.naturality walking_parallel_pair_hom.right).symm, }, + end, + fac' := λ c j, begin + rcases j with (_|_); simp only [functor.map_cocone_ι_app, cocone.of_cofork_ι, category.id_comp, + eq_to_hom_refl, category.assoc, functor.map_comp, cofork.of_π_ι_app, pushout.inl_desc, + preserves_pushout.inl_iso_inv_assoc], + exact (c.ι.naturality (walking_parallel_pair_hom.left)).trans (category.comp_id _) + end, + uniq' := λ s m h, begin + rw iso.eq_inv_comp, + have := h walking_parallel_pair.one, + dsimp [coequalizer_cocone] at this, + ext; simp only [preserves_pushout.inl_iso_hom_assoc, category.id_comp, pushout.inl_desc, + pushout.inr_desc, preserves_pushout.inr_iso_hom_assoc, + ← pushout_inl_eq_pushout_inr, ← this], + end }⟩ + + end category_theory.limits 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 36ff81958f62f..c97bd217a8437 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 @@ -8,6 +8,9 @@ import category_theory.limits.shapes.finite_products import category_theory.limits.preserves.shapes.products import category_theory.limits.preserves.shapes.equalizers import category_theory.limits.preserves.finite +import category_theory.limits.constructions.finite_products_of_binary_products +import category_theory.limits.constructions.equalizers +import category_theory.limits.constructions.binary_products /-! # Constructing limits from products and equalizers. @@ -214,6 +217,29 @@ preserves_limits_of_size.{w w} G := { preserves_limits_of_shape := λ J 𝒥, by exactI preserves_limit_of_preserves_equalizers_and_product G } +lemma has_finite_limits_of_has_terminal_and_pullbacks [has_terminal C] [has_pullbacks C] : + has_finite_limits C := +@@finite_limits_from_equalizers_and_finite_products _ + (@@has_finite_products_of_has_binary_and_terminal _ + (has_binary_products_of_terminal_and_pullbacks C) infer_instance) + (@@has_equalizers_of_pullbacks_and_binary_products _ + (has_binary_products_of_terminal_and_pullbacks C) infer_instance) + +/-- If G preserves terminal objects and pullbacks, it preserves all finite limits. -/ +def preserves_finite_limits_of_preserves_terminal_and_pullbacks + [has_terminal C] [has_pullbacks C] (G : C ⥤ D) + [preserves_limits_of_shape (discrete.{0} pempty) G] + [preserves_limits_of_shape walking_cospan G] : +preserves_finite_limits G := +begin + haveI : has_finite_limits C := has_finite_limits_of_has_terminal_and_pullbacks, + haveI : preserves_limits_of_shape (discrete walking_pair) G := + preserves_binary_products_of_preserves_terminal_and_pullbacks G, + exact @@preserves_finite_limits_of_preserves_equalizers_and_finite_products _ _ _ _ G + (preserves_equalizers_of_pullbacks_and_binary_products G) + (preserves_finite_products_of_preserves_binary_and_terminal G), +end + /-! We now dualize the above constructions, resorting to copy-paste. -/ @@ -398,4 +424,27 @@ preserves_colimits_of_size.{w} G := { preserves_colimits_of_shape := λ J 𝒥, by exactI preserves_colimit_of_preserves_coequalizers_and_coproduct G } +lemma has_finite_colimits_of_has_initial_and_pushouts [has_initial C] [has_pushouts C] : + has_finite_colimits C := +@@finite_colimits_from_coequalizers_and_finite_coproducts _ + (@@has_finite_coproducts_of_has_binary_and_terminal _ + (has_binary_coproducts_of_initial_and_pushouts C) infer_instance) + (@@has_coequalizers_of_pushouts_and_binary_coproducts _ + (has_binary_coproducts_of_initial_and_pushouts C) infer_instance) + +/-- If G preserves initial objects and pushouts, it preserves all finite colimits. -/ +def preserves_finite_colimits_of_preserves_initial_and_pushouts + [has_initial C] [has_pushouts C] (G : C ⥤ D) + [preserves_colimits_of_shape (discrete.{0} pempty) G] + [preserves_colimits_of_shape walking_span G] : +preserves_finite_colimits G := +begin + haveI : has_finite_colimits C := has_finite_colimits_of_has_initial_and_pushouts, + haveI : preserves_colimits_of_shape (discrete walking_pair) G := + preserves_binary_coproducts_of_preserves_initial_and_pushouts G, + exact @@preserves_finite_colimits_of_preserves_coequalizers_and_finite_coproducts _ _ _ _ G + (preserves_coequalizers_of_pushouts_and_binary_coproducts G) + (preserves_finite_coproducts_of_preserves_binary_and_initial G), +end + end category_theory.limits diff --git a/src/category_theory/limits/preserves/shapes/pullbacks.lean b/src/category_theory/limits/preserves/shapes/pullbacks.lean index e791e9be8db90..b6478e24bab23 100644 --- a/src/category_theory/limits/preserves/shapes/pullbacks.lean +++ b/src/category_theory/limits/preserves/shapes/pullbacks.lean @@ -85,6 +85,10 @@ def preserves_pullback_symmetry : preserves_limit (cospan g f) G := (c.π.naturality walking_cospan.hom.inl : _) } end } +lemma has_pullback_of_preserves_pullback [has_pullback f g] : + has_pullback (G.map f) (G.map g) := +⟨⟨⟨_, is_limit_pullback_cone_map_of_is_limit G _ (pullback_is_pullback _ _)⟩⟩⟩ + variables [has_pullback f g] [has_pullback (G.map f) (G.map g)] /-- If `G` preserves the pullback of `(f,g)`, then the pullback comparison map for `G` at `(f,g)` is @@ -167,6 +171,10 @@ def preserves_pushout_symmetry : preserves_colimit (span g f) G := (c.ι.naturality walking_span.hom.fst).symm } end } +lemma has_pushout_of_preserves_pushout [has_pushout f g] : + has_pushout (G.map f) (G.map g) := +⟨⟨⟨_, is_colimit_pushout_cocone_map_of_is_colimit G _ (pushout_is_pushout _ _)⟩⟩⟩ + variables [has_pushout f g] [has_pushout (G.map f) (G.map g)] /-- If `G` preserves the pushout of `(f,g)`, then the pushout comparison map for `G` at `(f,g)` is From e47961faf80c971a1b4b714aba3a6462de4c43b2 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Wed, 27 Jul 2022 14:35:02 +0000 Subject: [PATCH 097/173] chore(ring_theory/dedekind_domain/ideal): generalize a lemma (#15104) --- src/ring_theory/dedekind_domain/ideal.lean | 10 ++-------- src/ring_theory/unique_factorization_domain.lean | 5 +++++ 2 files changed, 7 insertions(+), 8 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index fd824db58ecb4..f75b732dd923a 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -753,12 +753,6 @@ open multiset unique_factorization_monoid ideal lemma prod_normalized_factors_eq_self (hI : I ≠ ⊥) : (normalized_factors I).prod = I := associated_iff_eq.1 (normalized_factors_prod hI) -lemma normalized_factors_prod {α : multiset (ideal T)} - (h : ∀ p ∈ α, prime p) : normalized_factors α.prod = α := -by { simp_rw [← multiset.rel_eq, ← associated_eq_eq], - exact prime_factors_unique (prime_of_normalized_factor) h - (normalized_factors_prod (α.prod_ne_zero_of_prime h)) } - lemma count_le_of_ideal_ge {I J : ideal T} (h : I ≤ J) (hI : I ≠ ⊥) (K : ideal T) : count K (normalized_factors J) ≤ count K (normalized_factors I) := le_iff_count.1 ((dvd_iff_normalized_factors_le_normalized_factors (ne_bot_of_le_ne_bot hI h) hI).1 @@ -769,7 +763,7 @@ lemma sup_eq_prod_inf_factors (hI : I ≠ ⊥) (hJ : J ≠ ⊥) : begin have H : normalized_factors (normalized_factors I ∩ normalized_factors J).prod = normalized_factors I ∩ normalized_factors J, - { apply _root_.normalized_factors_prod, + { apply normalized_factors_prod_of_prime, intros p hp, rw mem_inter at hp, exact prime_of_normalized_factor p hp.left }, @@ -783,7 +777,7 @@ begin { rw [dvd_iff_normalized_factors_le_normalized_factors this hJ, H], exact inf_le_right } }, { rw [← dvd_iff_le, dvd_iff_normalized_factors_le_normalized_factors, - _root_.normalized_factors_prod, le_iff_count], + normalized_factors_prod_of_prime, le_iff_count], { intro a, rw multiset.count_inter, exact le_min (count_le_of_ideal_ge le_sup_left hI a) diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index 752b80f6fb625..7ee4910bab9b8 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -627,6 +627,11 @@ begin rwa [multiset.eq_repeat_of_mem (λ b, h), multiset.prod_repeat] at this end +lemma normalized_factors_prod_of_prime [nontrivial α] [unique αˣ] {m : multiset α} + (h : ∀ p ∈ m, prime p) : (normalized_factors m.prod) = m := +by simpa only [←multiset.rel_eq, ←associated_eq_eq] using prime_factors_unique + (prime_of_normalized_factor) h (normalized_factors_prod (m.prod_ne_zero_of_prime h)) + end unique_factorization_monoid namespace unique_factorization_monoid From bfa735359b95956d9bc3094da605d0800f429f22 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 14:35:04 +0000 Subject: [PATCH 098/173] =?UTF-8?q?feat(ring=5Ftheory):=20Formally=20?= =?UTF-8?q?=C3=A9tale/smooth/unramified=20morphisms=20are=20stable=20under?= =?UTF-8?q?=20base=20change.=20(#15243)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/ring_theory/etale.lean | 56 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index afe7e6e9c29c0..901886dd0cec5 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -244,4 +244,60 @@ formally_etale.iff_unramified_and_smooth.mpr end comp +section base_change + +open_locale tensor_product + +variables {R : Type u} [comm_semiring R] +variables {A : Type u} [semiring A] [algebra R A] +variables (B : Type u) [comm_semiring B] [algebra R B] + + +instance formally_unramified.base_change [formally_unramified R A] : + formally_unramified B (B ⊗[R] A) := +begin + constructor, + introsI C _ _ I hI f₁ f₂ e, + letI := ((algebra_map B C).comp (algebra_map R B)).to_algebra, + haveI : is_scalar_tower R B C := is_scalar_tower.of_algebra_map_eq' rfl, + apply alg_hom.restrict_scalars_injective R, + apply tensor_product.ext, + any_goals { apply_instance }, + intros b a, + have : b ⊗ₜ[R] a = b • (1 ⊗ₜ a), by { rw [tensor_product.smul_tmul', smul_eq_mul, mul_one] }, + rw [this, alg_hom.restrict_scalars_apply, alg_hom.restrict_scalars_apply, map_smul, map_smul], + congr' 1, + change ((f₁.restrict_scalars R).comp tensor_product.include_right) a = + ((f₂.restrict_scalars R).comp tensor_product.include_right) a, + congr' 1, + refine formally_unramified.ext I ⟨2, hI⟩ _, + intro x, + exact alg_hom.congr_fun e (1 ⊗ₜ x) +end + +instance formally_smooth.base_change [formally_smooth R A] : + formally_smooth B (B ⊗[R] A) := +begin + constructor, + introsI C _ _ I hI f, + letI := ((algebra_map B C).comp (algebra_map R B)).to_algebra, + haveI : is_scalar_tower R B C := is_scalar_tower.of_algebra_map_eq' rfl, + refine ⟨tensor_product.product_left_alg_hom (algebra.of_id B C) _, _⟩, + { exact formally_smooth.lift I ⟨2, hI⟩ + ((f.restrict_scalars R).comp tensor_product.include_right) }, + { apply alg_hom.restrict_scalars_injective R, + apply tensor_product.ext, + any_goals { apply_instance }, + intros b a, + suffices : algebra_map B _ b * f (1 ⊗ₜ[R] a) = f (b ⊗ₜ[R] a), + { simpa [algebra.of_id_apply] }, + rw [← algebra.smul_def, ← map_smul, tensor_product.smul_tmul', smul_eq_mul, mul_one] }, +end + +instance formally_etale.base_change [formally_etale R A] : + formally_etale B (B ⊗[R] A) := +formally_etale.iff_unramified_and_smooth.mpr ⟨infer_instance, infer_instance⟩ + +end base_change + end algebra From e90cb2d9f0a03e6607fe4e55778c195881925a35 Mon Sep 17 00:00:00 2001 From: samvang <59202064+samvang@users.noreply.github.com> Date: Wed, 27 Jul 2022 14:35:07 +0000 Subject: [PATCH 099/173] feat(algebraic_geometry): comap of surjective homomorphism is closed embedding (#15291) The comap of a surjective homomorphism is a closed embedding between the Zariski spectra. Co-authored-by: Jake Levinson @jakelev Co-authored-by: Jake Levinson --- .../prime_spectrum/basic.lean | 57 +++++++++++++++++++ 1 file changed, 57 insertions(+) diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index da0a51b0b9481..f93eb904aa8c6 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -585,6 +585,63 @@ begin exact is_localization.comap_map_of_is_prime_disjoint M S _ x.2 h } end +section spec_of_surjective +/-! The comap of a surjective ring homomorphism is a closed embedding between the prime spectra. -/ + +open function ring_hom + +lemma comap_inducing_of_surjective (hf : surjective f) : inducing (comap f) := +{ induced := begin + simp_rw [topological_space_eq_iff, ←is_closed_compl_iff, is_closed_induced_iff, + is_closed_iff_zero_locus], + refine λ s, ⟨λ ⟨F, hF⟩, ⟨zero_locus (f ⁻¹' F), ⟨f ⁻¹' F, rfl⟩, + by rw [preimage_comap_zero_locus, surjective.image_preimage hf, hF]⟩, _⟩, + rintros ⟨-, ⟨F, rfl⟩, hF⟩, + exact ⟨f '' F, hF.symm.trans (preimage_comap_zero_locus f F)⟩, + end } + +lemma image_comap_zero_locus_eq_zero_locus_comap (hf : surjective f) (I : ideal S) : + comap f '' zero_locus I = zero_locus (I.comap f) := +begin + simp only [set.ext_iff, set.mem_image, mem_zero_locus, set_like.coe_subset_coe], + refine λ p, ⟨_, λ h_I_p, _⟩, + { rintro ⟨p, hp, rfl⟩ a ha, + exact hp ha }, + { have hp : ker f ≤ p.as_ideal := (ideal.comap_mono bot_le).trans h_I_p, + refine ⟨⟨p.as_ideal.map f, ideal.map_is_prime_of_surjective hf hp⟩, λ x hx, _, _⟩, + { obtain ⟨x', rfl⟩ := hf x, + exact ideal.mem_map_of_mem f (h_I_p hx) }, + { ext x, + change f x ∈ p.as_ideal.map f ↔ _, + rw ideal.mem_map_iff_of_surjective f hf, + refine ⟨_, λ hx, ⟨x, hx, rfl⟩⟩, + rintros ⟨x', hx', heq⟩, + rw ← sub_sub_cancel x' x, + refine p.as_ideal.sub_mem hx' (hp _), + rwa [mem_ker, map_sub, sub_eq_zero] } }, +end + +lemma range_comap_of_surjective (hf : surjective f) : + set.range (comap f) = zero_locus (ker f) := +begin + rw ← set.image_univ, + convert image_comap_zero_locus_eq_zero_locus_comap _ _ hf _, + rw zero_locus_bot, +end + +lemma is_closed_range_comap_of_surjective (hf : surjective f) : is_closed (set.range (comap f)) := +begin + rw range_comap_of_surjective _ f hf, + exact is_closed_zero_locus ↑(ker f), +end + +lemma closed_embedding_comap_of_surjective (hf : surjective f) : closed_embedding (comap f) := +{ induced := (comap_inducing_of_surjective S f hf).induced, + inj := comap_injective_of_surjective f hf, + closed_range := is_closed_range_comap_of_surjective S f hf } + +end spec_of_surjective + end comap section basic_open From e1a6884726f51786807ab2dc5b9fc659713fbf17 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 14:35:09 +0000 Subject: [PATCH 100/173] feat(linear_algebra/finite_dimensional): One-dimensional iff principal (#15558) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/linear_algebra/dimension.lean | 18 ++++++++++++++++++ src/linear_algebra/finite_dimensional.lean | 10 ++++++++++ src/ring_theory/principal_ideal_domain.lean | 1 + 3 files changed, 29 insertions(+) diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 4679c2385b197..551ad63201b6a 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1281,6 +1281,24 @@ begin simp [hw] } } 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, + le_span_singleton_iff, span_singleton_le_iff_mem], + split, + { rintro ⟨⟨m, hm⟩, hm'⟩, + choose f hf using hm', + exact ⟨m, ⟨λ v hv, ⟨f ⟨v, hv⟩, congr_arg coe (hf ⟨v, hv⟩)⟩, hm⟩⟩ }, + { rintro ⟨a, ⟨h, ha⟩⟩, + choose f hf using h, + exact ⟨⟨a, ha⟩, λ v, ⟨f v.1 v.2, subtype.ext (hf v.1 v.2)⟩⟩ } +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] + end division_ring section field diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index a44a8ae6263cb..79fe00f572d85 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -1646,6 +1646,16 @@ begin exact finrank_le_one v p, } end +lemma submodule.finrank_le_one_iff_is_principal (W : submodule K V) [finite_dimensional K W] : + finite_dimensional.finrank K W ≤ 1 ↔ W.is_principal := +by rw [← W.rank_le_one_iff_is_principal, ← finite_dimensional.finrank_eq_dim, + ← cardinal.nat_cast_le, nat.cast_one] + +lemma module.finrank_le_one_iff_top_is_principal [finite_dimensional K V] : + finite_dimensional.finrank K V ≤ 1 ↔ (⊤ : submodule K V).is_principal := +by rw [← module.rank_le_one_iff_top_is_principal, ← finite_dimensional.finrank_eq_dim, + ← cardinal.nat_cast_le, nat.cast_one] + -- We use the `linear_map.compatible_smul` typeclass here, to encompass two situations: -- * `A = K` -- * `[field K] [algebra K A] [is_scalar_tower K A V] [is_scalar_tower K A W]` diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index 4aeb34fc49dc6..8283f90ce82a8 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -38,6 +38,7 @@ section variables [ring R] [add_comm_group M] [module R M] /-- An `R`-submodule of `M` is principal if it is generated by one element. -/ +@[mk_iff] class submodule.is_principal (S : submodule R M) : Prop := (principal [] : ∃ a, S = span R {a}) From 09001169fee8963447df034d7b3e1132d0689b69 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 27 Jul 2022 14:35:11 +0000 Subject: [PATCH 101/173] feat(ring_theory/algebraic): `is_algebraic_algebra_map_iff` (#15586) This PR adds `is_algebraic_algebra_map_iff`, which is the `is_algebraic` analog of `is_integral_algebra_map_iff`. I also added the special case of an `intermediate_field`, which has been useful. --- src/field_theory/intermediate_field.lean | 3 +++ src/ring_theory/algebraic.lean | 9 ++++++++- 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index b79744d223f6d..afa1a1ac2bf71 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -465,6 +465,9 @@ eq_of_le_of_finrank_le' h_le h_finrank.le end finite_dimensional +lemma algebraic_iff {x : S} : is_algebraic K x ↔ is_algebraic K (x : L) := +(is_algebraic_algebra_map_iff (algebra_map S L).injective).symm + end intermediate_field /-- If `L/K` is algebraic, the `K`-subalgebras of `L` are all fields. -/ diff --git a/src/ring_theory/algebraic.lean b/src/ring_theory/algebraic.lean index 3c25db81256af..dc992c142e6d4 100644 --- a/src/ring_theory/algebraic.lean +++ b/src/ring_theory/algebraic.lean @@ -104,9 +104,16 @@ lemma is_algebraic_rat (R : Type u) {A : Type v} [division_ring A] [field R] [ch [algebra R A] (n : ℚ) : is_algebraic R (n : A) := by { rw ←map_rat_cast (algebra_map R A), exact is_algebraic_algebra_map n } +open is_scalar_tower + lemma is_algebraic_algebra_map_of_is_algebraic {a : S} : is_algebraic R a → is_algebraic R (algebra_map S A a) := -λ ⟨f, hf₁, hf₂⟩, ⟨f, hf₁, by rw [← is_scalar_tower.algebra_map_aeval R S A, hf₂, ring_hom.map_zero]⟩ +λ ⟨f, hf₁, hf₂⟩, ⟨f, hf₁, by rw [←algebra_map_aeval, hf₂, map_zero]⟩ + +lemma is_algebraic_algebra_map_iff {a : S} (h : function.injective (algebra_map S A)) : + is_algebraic R (algebra_map S A a) ↔ is_algebraic R a := +⟨λ ⟨p, hp0, hp⟩, ⟨p, hp0, h (by rwa [map_zero, algebra_map_aeval])⟩, + is_algebraic_algebra_map_of_is_algebraic⟩ end zero_ne_one From 235c31d2a02fa3f119d8cc61ea4ef2aaf921bfdd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Jul 2022 14:35:14 +0000 Subject: [PATCH 102/173] feat(algebraic_topology/dold_kan): construction of idempotent endomorphisms (#15616) This PR introduces a sequence of idempotent endomorphisms of the alternating face map complex of a simplicial object in a preadditive category. In a future PR, by taking the "limit" of this sequence, we shall get the projection on the normalized subcomplex, parallel to the degenerate sucomplex. --- src/algebraic_topology/dold_kan/faces.lean | 12 +- .../dold_kan/homotopies.lean | 2 +- .../dold_kan/projections.lean | 174 ++++++++++++++++++ 3 files changed, 181 insertions(+), 7 deletions(-) create mode 100644 src/algebraic_topology/dold_kan/projections.lean diff --git a/src/algebraic_topology/dold_kan/faces.lean b/src/algebraic_topology/dold_kan/faces.lean index f661668c0f7bb..17795196e3833 100644 --- a/src/algebraic_topology/dold_kan/faces.lean +++ b/src/algebraic_topology/dold_kan/faces.lean @@ -61,8 +61,8 @@ lemma of_comp {Y Z : C} {q n : ℕ} {φ : Y ⟶ X _[n+1]} higher_faces_vanish q (f ≫ φ) := λ j hj, by rw [assoc, v j hj, comp_zero] -lemma comp_Hσ_eq {Y : C} {n a q : ℕ} (hnaq : n=a+q) {φ : Y ⟶ X _[n+1]} - (v : higher_faces_vanish q φ) : φ ≫ (Hσ q).f (n+1) = +lemma comp_Hσ_eq {Y : C} {n a q : ℕ} {φ : Y ⟶ X _[n+1]} + (v : higher_faces_vanish q φ) (hnaq : n=a+q) : φ ≫ (Hσ q).f (n+1) = - φ ≫ X.δ ⟨a+1, nat.succ_lt_succ (nat.lt_succ_iff.mpr (nat.le.intro hnaq.symm))⟩ ≫ X.σ ⟨a, nat.lt_succ_iff.mpr (nat.le.intro hnaq.symm)⟩ := begin @@ -134,8 +134,8 @@ begin ring_exp, }, end -lemma comp_Hσ_eq_zero {Y : C} {n q : ℕ} (hqn : n Date: Wed, 27 Jul 2022 14:35:16 +0000 Subject: [PATCH 103/173] feat(analysis/convex/cone): add inner_dual_cone_eq_Inter_inner_dual_cone_singleton (#15639) Proof that a dual cone equals the intersection of dual cones of singleton sets. Part of #15637 --- src/analysis/convex/cone.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/src/analysis/convex/cone.lean b/src/analysis/convex/cone.lean index 795a445305b49..1f6aad7a947a2 100644 --- a/src/analysis/convex/cone.lean +++ b/src/analysis/convex/cone.lean @@ -616,4 +616,17 @@ lemma inner_dual_cone_le_inner_dual_cone (h : t ⊆ s) : lemma pointed_inner_dual_cone : s.inner_dual_cone.pointed := λ x hx, by rw inner_zero_right +/-- 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) := +begin + simp_rw [set.Inter_coe_set, subtype.coe_mk], + refine set.ext (λ x, iff.intro (λ hx, _) _), + { refine set.mem_Inter.2 (λ i, set.mem_Inter.2 (λ hi _, _)), + rintro ⟨ ⟩, + exact hx i hi }, + { simp only [set.mem_Inter, convex_cone.mem_coe, mem_inner_dual_cone, + set.mem_singleton_iff, forall_eq, imp_self] } +end + end dual From 31bab032f340bf635f20bafdf4dd2599956c4200 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Wed, 27 Jul 2022 14:35:18 +0000 Subject: [PATCH 104/173] feat(data/nat/totient): lemmas `totient_mul_of_prime_of_{not_}dvd` (#15645) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add `totient_mul_of_prime_of_not_dvd (hp : p.prime) (h : ¬ p ∣ n) : (p * n).totient = (p - 1) * n.totient` and golf `totient_mul_of_prime_of_dvd`. --- src/data/nat/totient.lean | 30 +++++++++++++++--------------- 1 file changed, 15 insertions(+), 15 deletions(-) diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index b80672dd76787..8526eaab5b8be 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -208,21 +208,6 @@ by rcases exists_eq_succ_of_ne_zero (pos_iff_ne_zero.1 hn) with ⟨m, rfl⟩; lemma totient_prime {p : ℕ} (hp : p.prime) : φ p = p - 1 := by rw [← pow_one p, totient_prime_pow hp]; simp -lemma totient_mul_of_prime_of_dvd {p n : ℕ} (hp : p.prime) (h : p ∣ n) : - (p * n).totient = p * n.totient := -begin - by_cases hzero : n = 0, - { simp [hzero] }, - { have hfin := (multiplicity.finite_nat_iff.2 ⟨hp.ne_one, zero_lt_iff.2 hzero⟩), - have h0 : 0 < (multiplicity p n).get hfin := multiplicity.pos_of_dvd hfin h, - obtain ⟨m, hm, hndiv⟩ := multiplicity.exists_eq_pow_mul_and_not_dvd hfin, - rw [hm, ← mul_assoc, ← pow_succ, nat.totient_mul (coprime_comm.mp (hp.coprime_pow_of_not_dvd - hndiv)), nat.totient_mul (coprime_comm.mp (hp.coprime_pow_of_not_dvd hndiv)), ← mul_assoc], - congr, - rw [ ← succ_pred_eq_of_pos h0, totient_prime_pow_succ hp, totient_prime_pow_succ hp, - succ_pred_eq_of_pos h0, ← mul_assoc p, ← pow_succ, ← succ_pred_eq_of_pos h0, nat.pred_succ] } -end - lemma totient_eq_iff_prime {p : ℕ} (hp : 0 < p) : p.totient = p - 1 ↔ p.prime := begin refine ⟨λ h, _, totient_prime⟩, @@ -351,4 +336,19 @@ begin apply mul_le_mul_left' (nat.totient_le d), end +lemma totient_mul_of_prime_of_dvd {p n : ℕ} (hp : p.prime) (h : p ∣ n) : + (p * n).totient = p * n.totient := +begin + have h1 := totient_gcd_mul_totient_mul p n, + rw [(gcd_eq_left h), mul_assoc] at h1, + simpa [(totient_pos hp.pos).ne', mul_comm] using h1, +end + +lemma totient_mul_of_prime_of_not_dvd {p n : ℕ} (hp : p.prime) (h : ¬ p ∣ n) : + (p * n).totient = (p - 1) * n.totient := +begin + rw [totient_mul _, totient_prime hp], + simpa [h] using coprime_or_dvd_of_prime hp n, +end + end nat From 366f262909db72436d122c06507f532d93efdb90 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 27 Jul 2022 17:03:23 +0000 Subject: [PATCH 105/173] feat(order/well_founded): typeclasses for well-founded `<` and `>` (#15399) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit # Well-founded typeclasses We introduce a new unbundled typeclass `is_well_founded` for a general well-founded relation, and reducible defs `well_founded_lt` and `well_founded_gt` for well-founded `<` and `>` specifically, to be used as mixins. For now, all we do is specialize only the most basic API on well-founded relations to these new typeclasses. This is just an initial development. If this is merged, a subsequent PR will redefine `is_well_order` in terms of `is_well_founded` (and remove the redundant `irrefl` field). Further PRs will focus on providing all the possible instances for this new typeclass, and actually using it throughout mathlib. Moved from #15023. ## Why do we want this? Here's the part where I justify why this is a good thing we want in mathlib. ### Well-ordered `<` There is a need to talk about well-ordered `<` relations, as a quick search for `is_well_order α (<)` or `is_well_order ι (<)` reveals. The most obvious way to spell this condition out, namely `[has_lt α] [is_well_order α (<)]` is actually the most inconvenient, since none of the results for linear orders are available. You actually need to write `[linear_order α] [is_well_order α (<)]` instead. With this new typeclass, the obvious spelling is now the good one: `[linear_order α] [well_founded_lt α]` ### Redundant typeclasses We actually already at least four typeclasses for the well-foundedness of a concrete relation: `category.noetherian_object`, `category.artinian_object`, `topological_space.noetherian_space`, and `wf_dvd_monoid`. One currently needs to specialize the characterizing properties of a well-founded relation to these separate typeclasses in order to use them. Redefining these typeclasses in terms of these new ones, and making them `reducible def`s, would alleviate this problem. ### Typeclass inference Well-founded relations are actually a nice candidate for typeclass inference, since there's many common constructions that automatically preserve well-foundedness. These include taking inverse images (including `measure` and `order.preimage`), maps through relation embeddings (including subtypes), lexicographic sums or products, and adding a top or bottom element. Many of these are currently provided as instances, but only for well-orders. ### Misleading theorem names Perhaps the most useful result on well-founded relations is that every nonempty set has a minimum - when you view the relation as `<`, that is. This leads to misleading theorem names when you're using them on a well-founded `>` relation. With this refactor, we could simply specialize `well_founded.min` into `well_founded_gt.max` and end up with clearer theorem statements. On a related note, given a linear order and a well-founded `<` relation, we can rephrase `well_founded.not_lt_min` into the much more convenient `well_founded_lt.min_le`. --- src/order/rel_classes.lean | 126 ++++++++++++++++++++++++++++++++++++ src/order/well_founded.lean | 8 +-- 2 files changed, 127 insertions(+), 7 deletions(-) diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index a9a297a7d63bf..322771f8582c3 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -204,6 +204,67 @@ instance is_strict_total_order_of_is_strict_total_order' /-! ### Well-order -/ +/-- A well-founded relation. Not to be confused with `is_well_order`. -/ +@[algebra, mk_iff] class is_well_founded (α : Type u) (r : α → α → Prop) : Prop := +(wf : well_founded r) + +instance has_well_founded.is_well_founded [h : has_well_founded α] : + is_well_founded α has_well_founded.r := { ..h } + +namespace is_well_founded +variables (r) [is_well_founded α r] + +/-- Induction on a well-founded relation. -/ +theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, r y x → C y) → C x) → C a := +wf.induction + +/-- All values are accessible under the well-founded relation. -/ +theorem apply : ∀ a, acc r a := wf.apply + +/-- Creates data, given a way to generate a value from all that compare as less under a well-founded +relation. See also `is_well_founded.fix_eq`. -/ +def fix {C : α → Sort*} : (Π (x : α), (Π (y : α), r y x → C y) → C x) → Π (x : α), C x := wf.fix + +/-- The value from `is_well_founded.fix` is built from the previous ones as specified. -/ +theorem fix_eq {C : α → Sort*} (F : Π (x : α), (Π (y : α), r y x → C y) → C x) : + ∀ x, fix r F x = F x (λ y h, fix r F y) := +wf.fix_eq F + +/-- Derive a `has_well_founded` instance from an `is_well_founded` instance. -/ +def to_has_well_founded : has_well_founded α := ⟨r, is_well_founded.wf⟩ + +end is_well_founded + +theorem well_founded.asymmetric {α : Sort*} {r : α → α → Prop} (h : well_founded r) : + ∀ ⦃a b⦄, r a b → ¬ r b a +| a := λ b hab hba, well_founded.asymmetric hba hab +using_well_founded { rel_tac := λ _ _, `[exact ⟨_, h⟩], + dec_tac := tactic.assumption } + +@[priority 100] -- see Note [lower instance priority] +instance is_well_founded.is_asymm (r : α → α → Prop) [is_well_founded α r] : is_asymm α r := +⟨is_well_founded.wf.asymmetric⟩ + +@[priority 100] -- see Note [lower instance priority] +instance is_well_founded.is_irrefl (r : α → α → Prop) [is_well_founded α r] : is_irrefl α r := +is_asymm.is_irrefl + +/-- A class for a well founded relation `<`. -/ +@[reducible] def well_founded_lt (α : Type*) [has_lt α] : Prop := is_well_founded α (<) + +/-- A class for a well founded relation `>`. -/ +@[reducible] def well_founded_gt (α : Type*) [has_lt α] : Prop := is_well_founded α (>) + +@[priority 100] -- See note [lower instance priority] +instance (α : Type*) [has_lt α] [h : well_founded_lt α] : well_founded_gt αᵒᵈ := h +@[priority 100] -- See note [lower instance priority] +instance (α : Type*) [has_lt α] [h : well_founded_gt α] : well_founded_lt αᵒᵈ := h + +theorem well_founded_gt_dual_iff (α : Type*) [has_lt α] : well_founded_gt αᵒᵈ ↔ well_founded_lt α := +⟨λ h, ⟨h.wf⟩, λ h, ⟨h.wf⟩⟩ +theorem well_founded_lt_dual_iff (α : Type*) [has_lt α] : well_founded_lt αᵒᵈ ↔ well_founded_gt α := +⟨λ h, ⟨h.wf⟩, λ h, ⟨h.wf⟩⟩ + /-- A well order is a well-founded linear order. -/ @[algebra] class is_well_order (α : Type u) (r : α → α → Prop) extends is_strict_total_order' α r : Prop := @@ -225,6 +286,56 @@ instance is_well_order.is_irrefl {α} (r : α → α → Prop) [is_well_order α instance is_well_order.is_asymm {α} (r : α → α → Prop) [is_well_order α r] : is_asymm α r := by apply_instance +namespace well_founded_lt +variables [has_lt α] [well_founded_lt α] + +/-- Inducts on a well-founded `<` relation. -/ +theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, y < x → C y) → C x) → C a := +is_well_founded.induction _ + +/-- All values are accessible under the well-founded `<`. -/ +theorem apply : ∀ a : α, acc (<) a := is_well_founded.apply _ + +/-- Creates data, given a way to generate a value from all that compare as lesser. See also +`well_founded_lt.fix_eq`. -/ +def fix {C : α → Sort*} : (Π (x : α), (Π (y : α), y < x → C y) → C x) → Π (x : α), C x := +is_well_founded.fix (<) + +/-- The value from `well_founded_lt.fix` is built from the previous ones as specified. -/ +theorem fix_eq {C : α → Sort*} (F : Π (x : α), (Π (y : α), y < x → C y) → C x) : + ∀ x, fix F x = F x (λ y h, fix F y) := +is_well_founded.fix_eq _ F + +/-- Derive a `has_well_founded` instance from a `well_founded_lt` instance. -/ +def to_has_well_founded : has_well_founded α := is_well_founded.to_has_well_founded (<) + +end well_founded_lt + +namespace well_founded_gt +variables [has_lt α] [well_founded_gt α] + +/-- Inducts on a well-founded `>` relation. -/ +theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, x < y → C y) → C x) → C a := +is_well_founded.induction _ + +/-- All values are accessible under the well-founded `>`. -/ +theorem apply : ∀ a : α, acc (>) a := is_well_founded.apply _ + +/-- Creates data, given a way to generate a value from all that compare as greater. See also +`well_founded_gt.fix_eq`. -/ +def fix {C : α → Sort*} : (Π (x : α), (Π (y : α), x < y → C y) → C x) → Π (x : α), C x := +is_well_founded.fix (>) + +/-- The value from `well_founded_gt.fix` is built from the successive ones as specified. -/ +theorem fix_eq {C : α → Sort*} (F : Π (x : α), (Π (y : α), x < y → C y) → C x) : + ∀ x, fix F x = F x (λ y h, fix F y) := +is_well_founded.fix_eq _ F + +/-- Derive a `has_well_founded` instance from a `well_founded_gt` instance. -/ +def to_has_well_founded : has_well_founded α := is_well_founded.to_has_well_founded (>) + +end well_founded_gt + /-- Construct a decidable linear order from a well-founded linear order. -/ noncomputable def is_well_order.linear_order (r : α → α → Prop) [is_well_order α r] : linear_order α := @@ -252,6 +363,10 @@ instance is_empty.is_well_order [is_empty α] (r : α → α → Prop) : is_well trans := is_empty_elim, wf := well_founded_of_empty r } +instance prod.lex.is_well_founded [is_well_founded α r] [is_well_founded β s] : + is_well_founded (α × β) (prod.lex r s) := +⟨prod.lex_wf is_well_founded.wf is_well_founded.wf⟩ + instance prod.lex.is_well_order [is_well_order α r] [is_well_order β s] : is_well_order (α × β) (prod.lex r s) := { trichotomous := λ ⟨a₁, a₂⟩ ⟨b₁, b₂⟩, @@ -276,6 +391,16 @@ instance prod.lex.is_well_order [is_well_order α r] [is_well_order β s] : end, wf := prod.lex_wf is_well_order.wf is_well_order.wf } +instance inv_image.is_well_founded (r : α → α → Prop) [is_well_founded α r] (f : β → α) : + is_well_founded _ (inv_image r f) := +⟨inv_image.wf f is_well_founded.wf⟩ + +instance measure.is_well_founded (f : α → ℕ) : is_well_founded _ (measure f) := ⟨measure_wf f⟩ + +theorem subrelation.is_well_founded (r : α → α → Prop) [is_well_founded α r] {s : α → α → Prop} + (h : subrelation s r) : is_well_founded α s := +⟨h.wf is_well_founded.wf⟩ + namespace set /-- An unbounded or cofinal set. -/ @@ -483,6 +608,7 @@ lemma transitive_gt [preorder α] : transitive (@gt α _) := transitive_of_trans instance order_dual.is_total_le [has_le α] [is_total α (≤)] : is_total αᵒᵈ (≤) := @is_total.swap α _ _ +instance : well_founded_lt ℕ := ⟨nat.lt_wf⟩ instance nat.lt.is_well_order : is_well_order ℕ (<) := ⟨nat.lt_wf⟩ instance [linear_order α] [h : is_well_order α (<)] : is_well_order αᵒᵈ (>) := h diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 11da41aa9abff..6ebd7fa9314f7 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -21,14 +21,8 @@ variables {α : Type*} namespace well_founded -theorem not_gt_of_lt {α : Sort*} {r : α → α → Prop} (h : well_founded r) : - ∀ ⦃a b⦄, r a b → ¬ r b a -| a := λ b hab hba, not_gt_of_lt hba hab -using_well_founded { rel_tac := λ _ _, `[exact ⟨_, h⟩], - dec_tac := tactic.assumption } - protected theorem is_asymm {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_asymm α r := -⟨h.not_gt_of_lt⟩ +⟨h.asymmetric⟩ instance {α : Sort*} [has_well_founded α] : is_asymm α has_well_founded.r := has_well_founded.wf.is_asymm From 8339dff4c5220e7eba52bd78fc260b94c7390fb1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 27 Jul 2022 17:03:24 +0000 Subject: [PATCH 106/173] feat(computability/ackermann): the Ackermann function isn't primitive recursive (#15505) See module docs for a thorough explanation of the proof. --- src/computability/ackermann.lean | 330 +++++++++++++++++++++++++++++++ 1 file changed, 330 insertions(+) create mode 100644 src/computability/ackermann.lean diff --git a/src/computability/ackermann.lean b/src/computability/ackermann.lean new file mode 100644 index 0000000000000..779c3b679b7ca --- /dev/null +++ b/src/computability/ackermann.lean @@ -0,0 +1,330 @@ +/- +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 computability.primrec +import tactic.linarith + +/-! +# Ackermann function + +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. + +## Main results + +- `exists_lt_ack_of_primrec`: any primitive recursive function is pointwise bounded above by `ack m` + for some `m`. +- `not_primrec₂_ack`: the two-argument Ackermann function is not primitive recursive. + +## Proof approach + +We very broadly adapt the proof idea from +https://www.planetmath.org/ackermannfunctionisnotprimitiverecursive. Namely, we prove that for any +primitive recursive `f : ℕ → ℕ`, there exists `m` such that `f n < ack m n` for all `n`. This then +implies that `λ n, ack n n` can't be primitive recursive, and so neither can `ack`. We aren't able +to use the same bounds as in that proof though, since our approach of using pairing functions +differs from their approach of using multivariate functions. + +The important bounds we show during the main inductive proof (`exists_lt_ack_of_primrec`) are the +following. Assuming `∀ n, f n < ack a n` and `∀ n, g n < ack b n`, we have: + +- `∀ n, nat.mkpair (f n) (g n) < ack (max a b + 3) n`. +- `∀ n, g (f n) < ack (max a b + 2) n`. +- `∀ n, nat.elim (f n.unpair.1) (λ (y IH : ℕ), g (nat.mkpair n.unpair.1 (nat.mkpair y IH))) + n.unpair.2 < ack (max a b + 9) n`. + +The last one is evidently the hardest. Using `nat.unpair_add_le`, we reduce it to the more +manageable + +- `∀ m n, elim (f m) (λ (y IH : ℕ), g (nat.mkpair m (nat.mkpair y IH))) n < + ack (max a b + 9) (m + n)`. + +We then prove this by induction on `n`. Our proof crucially depends on `ack_mkpair_lt`, which is +applied twice, giving us a constant of `4 + 4`. The rest of the proof consists of simpler bounds +which bump up our constant to `9`. +-/ + +open nat + +/-- The two-argument Ackermann function, defined so that + +- `ack 0 n = n + 1` +- `ack (m + 1) 0 = ack m 1` +- `ack (m + 1) (n + 1) = ack m (ack (m + 1) n)`. + +This is of interest as both a fast-growing function, and as an example of a recursive function that +isn't primitive recursive. -/ +def ack : ℕ → ℕ → ℕ +| 0 n := n + 1 +| (m + 1) 0 := ack m 1 +| (m + 1) (n + 1) := ack m (ack (m + 1) n) + +@[simp] theorem ack_zero (n : ℕ) : ack 0 n = n + 1 := by rw ack +@[simp] theorem ack_succ_zero (m : ℕ) : ack (m + 1) 0 = ack m 1 := by rw ack +@[simp] theorem ack_succ_succ (m n : ℕ) : ack (m + 1) (n + 1) = ack m (ack (m + 1) n) := by rw ack + +@[simp] theorem ack_one (n : ℕ) : ack 1 n = n + 2 := +begin + induction n with n IH, + { simp }, + { simp [IH] } +end + +@[simp] theorem ack_two (n : ℕ) : ack 2 n = 2 * n + 3 := +begin + induction n with n IH, + { simp }, + { simp [IH, mul_succ] } +end + +private theorem ack_three_aux (n : ℕ) : (ack 3 n : ℤ) = 2 ^ (n + 3) - 3 := +begin + induction n with n IH, + { simp, norm_num }, + { simp [IH, pow_succ], + rw [mul_sub, sub_add], + norm_num } +end + +@[simp] theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := +begin + zify, + rw cast_sub, + { exact_mod_cast ack_three_aux n }, + { have H : 3 ≤ 2 ^ 3 := by norm_num, + exact H.trans (pow_mono one_le_two $ le_add_left le_rfl) } +end + +theorem ack_pos : ∀ m n, 0 < ack m n +| 0 n := by simp +| (m + 1) 0 := by { rw ack_succ_zero, apply ack_pos } +| (m + 1) (n + 1) := by { rw ack_succ_succ, apply ack_pos } + +theorem one_lt_ack_succ_left : ∀ m n, 1 < ack (m + 1) n +| 0 n := by simp +| (m + 1) 0 := by { rw ack_succ_zero, apply one_lt_ack_succ_left } +| (m + 1) (n + 1) := by { rw ack_succ_succ, apply one_lt_ack_succ_left } + +theorem one_lt_ack_succ_right : ∀ m n, 1 < ack m (n + 1) +| 0 n := by simp +| (m + 1) n := begin + rw ack_succ_succ, + cases exists_eq_succ_of_ne_zero (ack_pos (m + 1) n).ne', + rw h, + apply one_lt_ack_succ_right +end + +theorem ack_strict_mono_right : ∀ m, strict_mono (ack m) +| 0 n₁ n₂ h := by simpa using h +| (m + 1) 0 (n + 1) h := begin + rw [ack_succ_zero, ack_succ_succ], + exact ack_strict_mono_right _ (one_lt_ack_succ_left m n) +end +| (m + 1) (n₁ + 1) (n₂ + 1) h := begin + rw [ack_succ_succ, ack_succ_succ], + apply ack_strict_mono_right _ (ack_strict_mono_right _ _), + rwa add_lt_add_iff_right at h +end + +theorem ack_mono_right (m : ℕ) : monotone (ack m) := (ack_strict_mono_right m).monotone + +theorem ack_injective_right (m : ℕ) : function.injective (ack m) := +(ack_strict_mono_right m).injective + +@[simp] theorem ack_lt_iff_right {m n₁ n₂ : ℕ} : ack m n₁ < ack m n₂ ↔ n₁ < n₂ := +(ack_strict_mono_right m).lt_iff_lt + +@[simp] theorem ack_le_iff_right {m n₁ n₂ : ℕ} : ack m n₁ ≤ ack m n₂ ↔ n₁ ≤ n₂ := +(ack_strict_mono_right m).le_iff_le + +@[simp] theorem ack_inj_right {m n₁ n₂ : ℕ} : ack m n₁ = ack m n₂ ↔ n₁ = n₂ := +(ack_injective_right m).eq_iff + +theorem max_ack_right (m n₁ n₂ : ℕ) : ack m (max n₁ n₂) = max (ack m n₁) (ack m n₂) := +(ack_mono_right m).map_max + +theorem add_lt_ack : ∀ m n, m + n < ack m n +| 0 n := by simp +| (m + 1) 0 := by simpa using add_lt_ack m 1 +| (m + 1) (n + 1) := +calc (m + 1) + n + 1 + ≤ m + (m + n + 2) : by linarith + ... < ack m (m + n + 2) : add_lt_ack _ _ + ... ≤ ack m (ack (m + 1) n) : ack_mono_right m $ + le_of_eq_of_le (by ring_nf) $ succ_le_of_lt $ add_lt_ack (m + 1) n + ... = ack (m + 1) (n + 1) : (ack_succ_succ m n).symm + +theorem add_add_one_le_ack (m n : ℕ) : m + n + 1 ≤ ack m n := succ_le_of_lt (add_lt_ack m n) + +theorem lt_ack_left (m n : ℕ) : m < ack m n := (self_le_add_right m n).trans_lt $ add_lt_ack m n +theorem lt_ack_right (m n : ℕ) : n < ack m n := (self_le_add_left n m).trans_lt $ add_lt_ack m n + +-- we reorder the arguments to appease the equation compiler +private theorem ack_strict_mono_left' : ∀ {m₁ m₂} n, m₁ < m₂ → ack m₁ n < ack m₂ n +| m 0 n := λ h, (nat.not_lt_zero m h).elim +| 0 (m + 1) 0 := λ h, by simpa using one_lt_ack_succ_right m 0 +| 0 (m + 1) (n + 1) := λ h, begin + rw [ack_zero, ack_succ_succ], + apply lt_of_le_of_lt (le_trans _ $ add_le_add_left (add_add_one_le_ack _ _) m) (add_lt_ack _ _), + linarith +end +| (m₁ + 1) (m₂ + 1) 0 := λ h, by simpa using ack_strict_mono_left' 1 ((add_lt_add_iff_right 1).1 h) +| (m₁ + 1) (m₂ + 1) (n + 1) := λ h, begin + rw [ack_succ_succ, ack_succ_succ], + exact (ack_strict_mono_left' _ $ (add_lt_add_iff_right 1).1 h).trans + (ack_strict_mono_right _ $ ack_strict_mono_left' n h) +end + +theorem ack_strict_mono_left (n : ℕ) : strict_mono (λ m, ack m n) := +λ m₁ m₂, ack_strict_mono_left' n + +theorem ack_mono_left (n : ℕ) : monotone (λ m, ack m n) := (ack_strict_mono_left n).monotone + +theorem ack_injective_left (n : ℕ) : function.injective (λ m, ack m n) := +(ack_strict_mono_left n).injective + +@[simp] theorem ack_lt_iff_left {m₁ m₂ n : ℕ} : ack m₁ n < ack m₂ n ↔ m₁ < m₂ := +(ack_strict_mono_left n).lt_iff_lt + +@[simp] theorem ack_le_iff_left {m₁ m₂ n : ℕ} : ack m₁ n ≤ ack m₂ n ↔ m₁ ≤ m₂ := +(ack_strict_mono_left n).le_iff_le + +@[simp] theorem ack_inj_left {m₁ m₂ n : ℕ} : ack m₁ n = ack m₂ n ↔ m₁ = m₂ := +(ack_injective_left n).eq_iff + +theorem max_ack_left (m₁ m₂ n : ℕ) : ack (max m₁ m₂) n = max (ack m₁ n) (ack m₂ n) := +(ack_mono_left n).map_max + +theorem ack_le_ack {m₁ m₂ n₁ n₂ : ℕ} (hm : m₁ ≤ m₂) (hn : n₁ ≤ n₂) : ack m₁ n₁ ≤ ack m₂ n₂ := +(ack_mono_left n₁ hm).trans $ ack_mono_right m₂ hn + +theorem ack_succ_right_le_ack_succ_left (m n : ℕ) : ack m (n + 1) ≤ ack (m + 1) n := +begin + cases n, + { simp }, + { rw [ack_succ_succ, succ_eq_add_one], + apply ack_mono_right m (le_trans _ $ add_add_one_le_ack _ n), + linarith } +end + +-- All the inequalities from this point onwards are specific to the main proof. + +private theorem sq_le_two_pow_add_one_minus_three (n : ℕ) : n ^ 2 ≤ 2 ^ (n + 1) - 3 := +begin + induction n with k hk, + { norm_num }, + { cases k, + { norm_num }, + { rw [succ_eq_add_one, add_sq, pow_succ 2, two_mul (2 ^ _), add_tsub_assoc_of_le, + add_comm (2 ^ _), add_assoc], + { apply add_le_add hk, + norm_num, + apply succ_le_of_lt, + rw [pow_succ, mul_lt_mul_left (@zero_lt_two ℕ _ _)], + apply lt_two_pow }, + { rw [pow_succ, pow_succ], + linarith [one_le_pow k 2 zero_lt_two] } } } +end + +theorem ack_add_one_sq_lt_ack_add_three : ∀ m n, (ack m n + 1) ^ 2 ≤ ack (m + 3) n +| 0 n := by simpa using sq_le_two_pow_add_one_minus_three (n + 2) +| (m + 1) 0 := by { rw [ack_succ_zero, ack_succ_zero], apply ack_add_one_sq_lt_ack_add_three } +| (m + 1) (n + 1) := begin + rw [ack_succ_succ, ack_succ_succ], + apply (ack_add_one_sq_lt_ack_add_three _ _).trans (ack_mono_right _ $ ack_mono_left _ _), + linarith +end + +theorem ack_ack_lt_ack_max_add_two (m n k : ℕ) : ack m (ack n k) < ack (max m n + 2) k := +calc ack m (ack n k) + ≤ ack (max m n) (ack n k) : ack_mono_left _ (le_max_left _ _) + ... < ack (max m n) (ack (max m n + 1) k) : ack_strict_mono_right _ $ ack_strict_mono_left k $ + lt_succ_of_le $ le_max_right m n + ... = ack (max m n + 1) (k + 1) : (ack_succ_succ _ _).symm + ... ≤ ack (max m n + 2) k : ack_succ_right_le_ack_succ_left _ _ + +theorem ack_add_one_sq_lt_ack_add_four (m n : ℕ) : ack m ((n + 1) ^ 2) < ack (m + 4) n := +calc ack m ((n + 1) ^ 2) + < ack m ((ack m n + 1) ^ 2) : ack_strict_mono_right m $ + pow_lt_pow_of_lt_left (succ_lt_succ $ lt_ack_right m n) zero_lt_two + ... ≤ ack m (ack (m + 3) n) : ack_mono_right m $ ack_add_one_sq_lt_ack_add_three m n + ... ≤ ack (m + 2) (ack (m + 3) n) : ack_mono_left _ $ by linarith + ... = ack (m + 3) (n + 1) : (ack_succ_succ _ n).symm + ... ≤ ack (m + 4) n : ack_succ_right_le_ack_succ_left _ n + +theorem ack_mkpair_lt (m n k : ℕ) : ack m (mkpair n k) < ack (m + 4) (max n k) := +(ack_strict_mono_right m $ mkpair_lt_max_add_one_sq n k).trans $ ack_add_one_sq_lt_ack_add_four _ _ + +/-- If `f` is primitive recursive, there exists `m` such that `f n < ack m n` for all `n`. -/ +theorem exists_lt_ack_of_nat_primrec {f : ℕ → ℕ} (hf : nat.primrec f) : ∃ m, ∀ n, f n < ack m n := +begin + induction hf with f g hf hg IHf IHg f g hf hg IHf IHg f g hf hg IHf IHg, + -- Zero function: + { exact ⟨0, ack_pos 0⟩ }, + -- Successor function: + { refine ⟨1, λ n, _⟩, + rw succ_eq_one_add, + apply add_lt_ack }, + -- Left projection: + { refine ⟨0, λ n, _⟩, + rw [ack_zero, lt_succ_iff], + exact unpair_left_le n }, + -- Right projection: + { refine ⟨0, λ n, _⟩, + rw [ack_zero, lt_succ_iff], + exact unpair_right_le n }, + all_goals { cases IHf with a ha, cases IHg with b hb }, + -- Pairing: + { refine ⟨max a b + 3, λ n, (mkpair_lt_max_add_one_sq _ _).trans_le $ + (nat.pow_le_pow_of_le_left (add_le_add_right _ _) 2).trans $ + ack_add_one_sq_lt_ack_add_three _ _⟩, + rw max_ack_left, + exact max_le_max (ha n).le (hb n).le }, + -- Composition: + { exact ⟨max a b + 2, λ n, + (ha _).trans $ (ack_strict_mono_right a $ hb n).trans $ ack_ack_lt_ack_max_add_two a b n⟩ }, + -- Primitive recursion operator: + { -- We prove this simpler inequality first. + have : ∀ {m n}, elim (f m) (λ y IH, g $ mkpair m $ mkpair y IH) n < ack (max a b + 9) (m + n), + { intros m n, + -- We induct on n. + induction n with n IH, + -- The base case is easy. + { apply (ha m).trans (ack_strict_mono_left m $ (le_max_left a b).trans_lt _), + linarith }, + { -- We get rid of the first `mkpair`. + rw elim_succ, + apply (hb _).trans ((ack_mkpair_lt _ _ _).trans_le _), + -- If m is the maximum, we get a very weak inequality. + cases lt_or_le _ m with h₁ h₁, + { rw max_eq_left h₁.le, + exact ack_le_ack (add_le_add (le_max_right a b) $ by norm_num) (self_le_add_right m _) }, + rw max_eq_right h₁, + -- We get rid of the second `mkpair`. + apply (ack_mkpair_lt _ _ _).le.trans, + -- If n is the maximum, we get a very weak inequality. + cases lt_or_le _ n with h₂ h₂, + { rw [max_eq_left h₂.le, add_assoc], + exact ack_le_ack (add_le_add (le_max_right a b) $ by norm_num) + ((le_succ n).trans $ self_le_add_left _ _) }, + rw max_eq_right h₂, + -- We now use the inductive hypothesis, and some simple algebraic manipulation. + apply (ack_strict_mono_right _ IH).le.trans, + rw [add_succ m, add_succ _ 8, ack_succ_succ (_ + 8), add_assoc], + exact ack_mono_left _ (add_le_add (le_max_right a b) le_rfl) } }, + -- The proof is now simple. + exact ⟨max a b + 9, λ n, this.trans_le $ ack_mono_right _ $ unpair_add_le n⟩ } +end + +theorem not_nat_primrec_ack_self : ¬ nat.primrec (λ n, ack n n) := +λ h, by { cases exists_lt_ack_of_nat_primrec h with m hm, exact (hm m).false } + +theorem not_primrec_ack_self : ¬ _root_.primrec (λ n, ack n n) := +by { rw primrec.nat_iff, exact not_nat_primrec_ack_self } + +/-- The Ackermann function is not primitive recursive. -/ +theorem not_primrec₂_ack : ¬ primrec₂ ack := +λ h, not_primrec_ack_self $ h.comp primrec.id primrec.id From c694f40634d010238143ab3ccdeae2d91ab45b56 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 27 Jul 2022 17:03:25 +0000 Subject: [PATCH 107/173] feat(linear_algebra): lemmas on associatedness of determinants of linear maps (#15587) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `det f` equals `det f'` up to units if `f'` is `f` composed with a linear equiv. First a homogeneous form where `f f' : M → M`, then a heterogeneous form where `f f' : M → N` and the equivs `e e' : N → M`. Co-authored-by: Anne Baanen --- src/linear_algebra/determinant.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index a8d473a41635e..f09292b5cfca6 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -409,6 +409,25 @@ have is_unit (linear_map.to_matrix (finite_dimensional.fin_basis 𝕜 M) by simp only [linear_map.det_to_matrix, is_unit_iff_ne_zero.2 hf], linear_equiv.of_is_unit_det this +lemma linear_map.associated_det_of_eq_comp (e : M ≃ₗ[R] M) (f f' : M →ₗ[R] M) + (h : ∀ x, f x = f' (e x)) : associated f.det f'.det := +begin + suffices : associated (f' ∘ₗ ↑e).det f'.det, + { convert this using 2, ext x, exact h x }, + rw [← mul_one f'.det, linear_map.det_comp], + exact associated.mul_left _ (associated_one_iff_is_unit.mpr e.is_unit_det') +end + +lemma linear_map.associated_det_comp_equiv {N : Type*} [add_comm_group N] [module R N] + (f : N →ₗ[R] M) (e e' : M ≃ₗ[R] N) : + associated (f ∘ₗ ↑e).det (f ∘ₗ ↑e').det := +begin + refine linear_map.associated_det_of_eq_comp (e.trans e'.symm) _ _ _, + intro x, + simp only [linear_map.comp_apply, linear_equiv.coe_coe, linear_equiv.trans_apply, + linear_equiv.apply_symm_apply], +end + /-- The determinant of a family of vectors with respect to some basis, as an alternating multilinear map. -/ def basis.det : alternating_map R M R ι := From a9eaeee9a78c59d7b85865d15cc9402832d1b3f7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 27 Jul 2022 17:03:26 +0000 Subject: [PATCH 108/173] doc(field_theory/splitting_field): update code example to reflect changes to diamond issues (#15620) We now have infrastructure in place to avoid diamonds in nat/int/rat algebras, so we can't appeal to a lack of it when explaining diamond issues. --- src/field_theory/splitting_field.lean | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index 0d14e4f887629..e6f2b4f51bc6d 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -630,17 +630,20 @@ splitting_field_aux.field _ _ instance inhabited : inhabited (splitting_field f) := ⟨37⟩ -/-- This should be an instance globally, but it creates diamonds with the `ℕ` and `ℤ` actions: +/-- This should be an instance globally, but it creates diamonds with the `ℕ`, `ℤ`, and `ℚ` algebras +(via their `smul` and `to_fun` fields): ```lean example : - (add_comm_monoid.nat_module : module ℕ (splitting_field f)) = - @algebra.to_module _ _ _ _ (splitting_field.algebra' f) := + (algebra_nat : algebra ℕ (splitting_field f)) = splitting_field.algebra' f := rfl -- fails example : - (add_comm_group.int_module _ : module ℤ (splitting_field f)) = - @algebra.to_module _ _ _ _ (splitting_field.algebra' f) := + (algebra_int _ : algebra ℤ (splitting_field f)) = splitting_field.algebra' f := +rfl -- fails + +example [char_zero K] [char_zero (splitting_field f)] : + (algebra_rat : algebra ℚ (splitting_field f)) = splitting_field.algebra' f := rfl -- fails ``` @@ -649,8 +652,8 @@ Until we resolve these diamonds, it's more convenient to only turn this instance 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, but this is a problem throughout the -library and not unique to this `algebra` instance. +`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 _ _ _ From 6db51145115a9148e1a9588272590642b76d4bf4 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 27 Jul 2022 17:03:27 +0000 Subject: [PATCH 109/173] feat(analysis/special_functions/trigonometric/angle): equality of `cos` or `sin` (#15651) `analysis.special_functions.trigonometric.angle` has results relating equality of `real.cos` of two reals, or `real.sin` of two reals, to relations between those reals converted to `angle`. Add variants of those results where one or both of the arguments are passed as `angle` instead of as reals, with `real.angle.cos` and `real.angle.sin` used on those `angle` arguments. The version for `cos` with one `angle` and one real argument, in particular, is what I want for proving that the oriented angle between two nonzero vectors is plus or minus the unoriented angle. --- .../trigonometric/angle.lean | 36 ++++++++++++++++--- 1 file changed, 31 insertions(+), 5 deletions(-) diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index 6cad43780acac..1830a5ad901cc 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -107,7 +107,8 @@ by convert two_nsmul_eq_iff; simp lemma two_zsmul_eq_zero_iff {θ : angle} : (2 : ℤ) • θ = 0 ↔ (θ = 0 ∨ θ = π) := by simp_rw [two_zsmul, ←two_nsmul, two_nsmul_eq_zero_iff] -theorem cos_eq_iff_eq_or_eq_neg {θ ψ : ℝ} : cos θ = cos ψ ↔ (θ : angle) = ψ ∨ (θ : angle) = -ψ := +theorem cos_eq_iff_coe_eq_or_eq_neg {θ ψ : ℝ} : + cos θ = cos ψ ↔ (θ : angle) = ψ ∨ (θ : angle) = -ψ := begin split, { intro Hcos, @@ -132,12 +133,12 @@ begin zero_mul] } end -theorem sin_eq_iff_eq_or_add_eq_pi {θ ψ : ℝ} : +theorem sin_eq_iff_coe_eq_or_add_eq_pi {θ ψ : ℝ} : sin θ = sin ψ ↔ (θ : angle) = ψ ∨ (θ : angle) + ψ = π := begin split, { intro Hsin, rw [← cos_pi_div_two_sub, ← cos_pi_div_two_sub] at Hsin, - cases cos_eq_iff_eq_or_eq_neg.mp Hsin with h h, + cases cos_eq_iff_coe_eq_or_eq_neg.mp Hsin with h h, { left, rw [coe_sub, coe_sub] at h, exact sub_right_inj.1 h }, right, rw [coe_sub, coe_sub, eq_neg_iff_add_eq_zero, add_sub, sub_add_eq_add_sub, ← coe_add, add_halves, sub_sub, sub_eq_zero] at h, @@ -156,8 +157,8 @@ end theorem cos_sin_inj {θ ψ : ℝ} (Hcos : cos θ = cos ψ) (Hsin : sin θ = sin ψ) : (θ : angle) = ψ := begin - cases cos_eq_iff_eq_or_eq_neg.mp Hcos with hc hc, { exact hc }, - cases sin_eq_iff_eq_or_add_eq_pi.mp Hsin with hs hs, { exact hs }, + cases cos_eq_iff_coe_eq_or_eq_neg.mp Hcos with hc hc, { exact hc }, + cases sin_eq_iff_coe_eq_or_add_eq_pi.mp Hsin with hs hs, { exact hs }, rw [eq_neg_iff_add_eq_zero, hs] at hc, obtain ⟨n, hn⟩ : ∃ n, n • _ = _ := quotient_add_group.left_rel_apply.mp (quotient.exact' hc), rw [← neg_one_mul, add_zero, ← sub_eq_zero, zsmul_eq_mul, ← mul_assoc, ← sub_mul, @@ -187,6 +188,31 @@ rfl @[continuity] lemma continuous_cos : continuous cos := continuous_quotient_lift_on' _ real.continuous_cos +lemma cos_eq_real_cos_iff_eq_or_eq_neg {θ : angle} {ψ : ℝ} : cos θ = real.cos ψ ↔ θ = ψ ∨ θ = -ψ := +begin + induction θ using real.angle.induction_on, + exact cos_eq_iff_coe_eq_or_eq_neg +end + +lemma cos_eq_iff_eq_or_eq_neg {θ ψ : angle} : cos θ = cos ψ ↔ θ = ψ ∨ θ = -ψ := +begin + induction ψ using real.angle.induction_on, + exact cos_eq_real_cos_iff_eq_or_eq_neg +end + +lemma sin_eq_real_sin_iff_eq_or_add_eq_pi {θ : angle} {ψ : ℝ} : + sin θ = real.sin ψ ↔ θ = ψ ∨ θ + ψ = π := +begin + induction θ using real.angle.induction_on, + exact sin_eq_iff_coe_eq_or_add_eq_pi +end + +lemma sin_eq_iff_eq_or_add_eq_pi {θ ψ : angle} : sin θ = sin ψ ↔ θ = ψ ∨ θ + ψ = π := +begin + induction ψ using real.angle.induction_on, + exact sin_eq_real_sin_iff_eq_or_add_eq_pi +end + end angle end real From b39ba3bcfd3414d92cba5e98663a124ee5a9502f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 27 Jul 2022 17:03:28 +0000 Subject: [PATCH 110/173] chore(analysis/convex/cone): add trivial coercion lemmas (#15713) This also changes `le` to unfold in terms of `coe` rather than `carrier`, since that's the form we have lemmas about. --- src/analysis/convex/cone.lean | 20 +++++++++++++++++--- 1 file changed, 17 insertions(+), 3 deletions(-) diff --git a/src/analysis/convex/cone.lean b/src/analysis/convex/cone.lean index 1f6aad7a947a2..66c018ce84c98 100644 --- a/src/analysis/convex/cone.lean +++ b/src/analysis/convex/cone.lean @@ -78,12 +78,14 @@ instance : has_coe (convex_cone 𝕜 E) (set E) := ⟨convex_cone.carrier⟩ instance : has_mem E (convex_cone 𝕜 E) := ⟨λ m S, m ∈ S.carrier⟩ -instance : has_le (convex_cone 𝕜 E) := ⟨λ S T, S.carrier ⊆ T.carrier⟩ +instance : has_le (convex_cone 𝕜 E) := ⟨λ S T, (S : set E) ⊆ T⟩ -instance : has_lt (convex_cone 𝕜 E) := ⟨λ S T, S.carrier ⊂ T.carrier⟩ +instance : has_lt (convex_cone 𝕜 E) := ⟨λ S T, (S : set E) ⊂ T⟩ @[simp, norm_cast] lemma mem_coe {x : E} : x ∈ (S : set E) ↔ x ∈ S := iff.rfl +@[simp] lemma coe_mk {s : set E} {h₁ h₂} : ↑(@mk 𝕜 _ _ _ _ s h₁ h₂) = s := rfl + @[simp] lemma mem_mk {s : set E} {h₁ h₂ x} : x ∈ @mk 𝕜 _ _ _ _ s h₁ h₂ ↔ x ∈ s := iff.rfl /-- Two `convex_cone`s are equal if the underlying sets are equal. -/ @@ -105,7 +107,7 @@ instance : has_inf (convex_cone 𝕜 E) := ⟨λ S T, ⟨S ∩ T, λ c hc x hx, ⟨S.smul_mem hc hx.1, T.smul_mem hc hx.2⟩, λ x hx y hy, ⟨S.add_mem hx.1 hy.1, T.add_mem hx.2 hy.2⟩⟩⟩ -lemma coe_inf : ((S ⊓ T : convex_cone 𝕜 E) : set E) = ↑S ∩ ↑T := rfl +@[simp] lemma coe_inf : ((S ⊓ T : convex_cone 𝕜 E) : set E) = ↑S ∩ ↑T := rfl lemma mem_inf {x} : x ∈ S ⊓ T ↔ x ∈ S ∧ x ∈ T := iff.rfl @@ -114,18 +116,30 @@ instance : has_Inf (convex_cone 𝕜 E) := λ c hc x hx, mem_bInter $ λ s hs, s.smul_mem hc $ mem_Inter₂.1 hx s hs, λ x hx y hy, mem_bInter $ λ s hs, s.add_mem (mem_Inter₂.1 hx s hs) (mem_Inter₂.1 hy s hs)⟩⟩ +@[simp] lemma coe_Inf (S : set (convex_cone 𝕜 E)) : ↑(Inf S) = ⋂ s ∈ S, (s : set E) := rfl + lemma mem_Inf {x : E} {S : set (convex_cone 𝕜 E)} : x ∈ Inf S ↔ ∀ s ∈ S, x ∈ s := mem_Inter₂ +@[simp] lemma coe_infi {ι : Sort*} (f : ι → convex_cone 𝕜 E) : ↑(infi f) = ⋂ i, (f i : set E) := +by simp [infi] + +lemma mem_infi {ι : Sort*} {x : E} {f : ι → convex_cone 𝕜 E} : x ∈ infi f ↔ ∀ i, x ∈ f i := +mem_Inter₂.trans $ by simp + variables (𝕜) instance : has_bot (convex_cone 𝕜 E) := ⟨⟨∅, λ c hc x, false.elim, λ x, false.elim⟩⟩ lemma mem_bot (x : E) : x ∈ (⊥ : convex_cone 𝕜 E) = false := rfl +@[simp] lemma coe_bot : ↑(⊥ : convex_cone 𝕜 E) = (∅ : set E) := rfl + instance : has_top (convex_cone 𝕜 E) := ⟨⟨univ, λ c hc x hx, mem_univ _, λ x hx y hy, mem_univ _⟩⟩ lemma mem_top (x : E) : x ∈ (⊤ : convex_cone 𝕜 E) := mem_univ x +@[simp] lemma coe_top : ↑(⊤ : convex_cone 𝕜 E) = (univ : set E) := rfl + instance : complete_lattice (convex_cone 𝕜 E) := { le := (≤), lt := (<), From 46a1e8025475be16cdd0140034fa598d7281ae9d Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 27 Jul 2022 19:44:47 +0000 Subject: [PATCH 111/173] chore(data/polynomial/cancel_leads): golf using `compute_degree_le` (#14776) This PR is the companion to #14762. It serves to show how to use the tactic in a few cases in what is already in mathlib. I generalized lemma `nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree` to assume `ring` instead of `comm_ring + is_domain`. The lemma with the weaker assumptions is `nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree_of_comm`. I left the old lemma as well,with proof a simple application of the newer lemma. --- src/data/polynomial/cancel_leads.lean | 58 +++++++++++++-------------- 1 file changed, 27 insertions(+), 31 deletions(-) diff --git a/src/data/polynomial/cancel_leads.lean b/src/data/polynomial/cancel_leads.lean index 377db80acaa79..a75bac5dacfab 100644 --- a/src/data/polynomial/cancel_leads.lean +++ b/src/data/polynomial/cancel_leads.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import data.polynomial.degree.definitions - +import tactic.compute_degree /-! # Cancel the leading terms of two polynomials @@ -23,7 +23,7 @@ open_locale polynomial variables {R : Type*} -section comm_ring +section ring variables [ring R] (p q : R[X]) /-- `cancel_leads p q` is formed by multiplying `p` and `q` by monomials so that they @@ -36,19 +36,9 @@ variables {p q} @[simp] lemma neg_cancel_leads : - p.cancel_leads q = q.cancel_leads p := neg_sub _ _ -end comm_ring - -section comm_ring -variables [comm_ring R] {p q : R[X]} - -lemma dvd_cancel_leads_of_dvd_of_dvd {r : R[X]} (pq : p ∣ q) (pr : p ∣ r) : - p ∣ q.cancel_leads r := -dvd_sub (pr.trans (dvd.intro_left _ rfl)) (pq.trans (dvd.intro_left _ rfl)) - -end comm_ring - -lemma nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree [comm_ring R] [is_domain R] - {p q : R[X]} (h : p.nat_degree ≤ q.nat_degree) (hq : 0 < q.nat_degree) : +lemma nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree_of_comm + (comm : p.leading_coeff * q.leading_coeff = q.leading_coeff * p.leading_coeff) + (h : p.nat_degree ≤ q.nat_degree) (hq : 0 < q.nat_degree) : (p.cancel_leads q).nat_degree < q.nat_degree := begin by_cases hp : p = 0, @@ -57,26 +47,32 @@ begin rw [cancel_leads, sub_eq_add_neg, tsub_eq_zero_iff_le.mpr h, pow_zero, mul_one], by_cases h0 : C p.leading_coeff * q + -(C q.leading_coeff * X ^ (q.nat_degree - p.nat_degree) * p) = 0, - { convert hq, - simp only [h0, nat_degree_zero], }, - have hq0 : ¬ q = 0, - { contrapose! hq, - simp [hq] }, + { exact (le_of_eq (by simp only [h0, nat_degree_zero])).trans_lt hq }, apply lt_of_le_of_ne, - { rw [← with_bot.coe_le_coe, ← degree_eq_nat_degree h0, ← degree_eq_nat_degree hq0], - apply le_trans (degree_add_le _ _), - rw ← leading_coeff_eq_zero at hp hq0, - simp only [max_le_iff, degree_C hp, degree_C hq0, le_refl q.degree, true_and, nat.cast_with_bot, - nsmul_one, degree_neg, degree_mul, zero_add, degree_X, degree_pow], - rw leading_coeff_eq_zero at hp hq0, - rw [degree_eq_nat_degree hp, degree_eq_nat_degree hq0, ← with_bot.coe_add, with_bot.coe_le_coe, - tsub_add_cancel_of_le h], }, + { compute_degree_le, + repeat { rwa nat.sub_add_cancel } }, { contrapose! h0, - rw [← leading_coeff_eq_zero, leading_coeff, h0, mul_assoc, mul_comm _ p, + rw [← leading_coeff_eq_zero, leading_coeff, h0, mul_assoc, X_pow_mul, ← tsub_add_cancel_of_le h, add_comm _ p.nat_degree], simp only [coeff_mul_X_pow, coeff_neg, coeff_C_mul, add_tsub_cancel_left, coeff_add], - rw [add_comm p.nat_degree, tsub_add_cancel_of_le h, ← leading_coeff, ← leading_coeff, - mul_comm _ q.leading_coeff, ← sub_eq_add_neg, ← mul_sub, sub_self, mul_zero] } + rw [add_comm p.nat_degree, tsub_add_cancel_of_le h, ← leading_coeff, ← leading_coeff, comm, + add_right_neg] } end +end ring + +section comm_ring +variables [comm_ring R] {p q : R[X]} + +lemma dvd_cancel_leads_of_dvd_of_dvd {r : R[X]} (pq : p ∣ q) (pr : p ∣ r) : + p ∣ q.cancel_leads r := +dvd_sub (pr.trans (dvd.intro_left _ rfl)) (pq.trans (dvd.intro_left _ rfl)) + +lemma nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree + (h : p.nat_degree ≤ q.nat_degree) (hq : 0 < q.nat_degree) : + (p.cancel_leads q).nat_degree < q.nat_degree := +nat_degree_cancel_leads_lt_of_nat_degree_le_nat_degree_of_comm (mul_comm _ _) h hq + +end comm_ring + end polynomial From 50391eec5e6c10dee809f9ed1880fc18dc7b4364 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Jul 2022 19:44:48 +0000 Subject: [PATCH 112/173] feat(category_theory): a characterization of separating objects (#14838) --- src/category_theory/generator.lean | 44 ++++++++++++++++++++++++++++++ 1 file changed, 44 insertions(+) diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index 23f20b7069223..7861e931a2a7c 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -239,6 +239,28 @@ lemma is_codetecting_empty_of_groupoid [∀ {X Y : C} (f : X ⟶ Y), is_iso f] : end empty +lemma is_separating_iff_epi (𝒢 : set C) + [Π (A : C), has_coproduct (λ f : Σ G : 𝒢, (G : C) ⟶ A, (f.1 : C))] : + is_separating 𝒢 ↔ ∀ A : C, epi (sigma.desc (@sigma.snd 𝒢 (λ G, (G : C) ⟶ A))) := +begin + refine ⟨λ h A, ⟨λ Z u v huv, h _ _ (λ G hG f, _)⟩, λ h X Y f g hh, _⟩, + { simpa using (sigma.ι (λ f : Σ G : 𝒢, (G : C) ⟶ A, (f.1 : C)) ⟨⟨G, hG⟩, f⟩) ≫= huv }, + { haveI := h X, + refine (cancel_epi (sigma.desc (@sigma.snd 𝒢 (λ G, (G : C) ⟶ X)))).1 (colimit.hom_ext (λ j, _)), + simpa using hh j.as.1.1 j.as.1.2 j.as.2 } +end + +lemma is_coseparating_iff_mono (𝒢 : set C) + [Π (A : C), has_product (λ f : Σ G : 𝒢, A ⟶ (G : C), (f.1 : C))] : + is_coseparating 𝒢 ↔ ∀ A : C, mono (pi.lift (@sigma.snd 𝒢 (λ G, A ⟶ (G : C)))) := +begin + refine ⟨λ h A, ⟨λ Z u v huv, h _ _ (λ G hG f, _)⟩, λ h X Y f g hh, _⟩, + { simpa using huv =≫ (pi.π (λ f : Σ G : 𝒢, A ⟶ (G : C), (f.1 : C)) ⟨⟨G, hG⟩, f⟩) }, + { haveI := h Y, + refine (cancel_mono (pi.lift (@sigma.snd 𝒢 (λ G, Y ⟶ (G : C))))).1 (limit.hom_ext (λ j, _)), + simpa using hh j.as.1.1 j.as.1.2 j.as.2 } +end + section well_powered namespace subobject @@ -381,6 +403,28 @@ lemma is_coseparator_iff_faithful_yoneda_obj (G : C) : λ h, (is_coseparator_def _).2 $ λ X Y f g hfg, quiver.hom.op_inj $ by exactI (yoneda.obj G).map_injective (funext hfg)⟩ +lemma is_separator_iff_epi (G : C) [Π A : C, has_coproduct (λ (f : G ⟶ A), G)] : + is_separator G ↔ ∀ (A : C), epi (sigma.desc (λ (f : G ⟶ A), f)) := +begin + rw is_separator_def, + refine ⟨λ h A, ⟨λ Z u v huv, h _ _ (λ i, _)⟩, λ h X Y f g hh, _⟩, + { simpa using (sigma.ι _ i) ≫= huv }, + { haveI := h X, + refine (cancel_epi (sigma.desc (λ (f : G ⟶ X), f))).1 (colimit.hom_ext (λ j, _)), + simpa using hh j.as } +end + +lemma is_coseparator_iff_mono (G : C) [Π A : C, has_product (λ (f : A ⟶ G), G)] : + is_coseparator G ↔ ∀ (A : C), mono (pi.lift (λ (f : A ⟶ G), f)) := +begin + rw is_coseparator_def, + refine ⟨λ h A, ⟨λ Z u v huv, h _ _ (λ i, _)⟩, λ h X Y f g hh, _⟩, + { simpa using huv =≫ (pi.π _ i) }, + { haveI := h Y, + refine (cancel_mono (pi.lift (λ (f : Y ⟶ G), f))).1 (limit.hom_ext (λ j, _)), + simpa using hh j.as } +end + section zero_morphisms variables [has_zero_morphisms C] From 6f8bcf82974b8667bd091a6e5cee22fc180b5c46 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Wed, 27 Jul 2022 19:44:49 +0000 Subject: [PATCH 113/173] feat(probability/martingale): the discrete stochastic integral of a submartingale is a submartingale (#14909) This PR proves that the discrete stochastic integral of a predictable process with a submartingale is a submartingale. Co-authored-by: RemyDegenne --- src/probability/martingale.lean | 37 +++++++++++++++++++++++++++++++++ src/probability/stopping.lean | 22 ++++++++++++++------ 2 files changed, 53 insertions(+), 6 deletions(-) diff --git a/src/probability/martingale.lean b/src/probability/martingale.lean index 119e8184fedbd..a793f225030b4 100644 --- a/src/probability/martingale.lean +++ b/src/probability/martingale.lean @@ -604,6 +604,43 @@ end end maximal +lemma submartingale.sum_mul_sub [is_finite_measure μ] {R : ℝ} {ξ f : ℕ → α → ℝ} + (hf : submartingale f 𝒢 μ) (hξ : adapted 𝒢 ξ) + (hbdd : ∀ n x, ξ n x ≤ R) (hnonneg : ∀ n x, 0 ≤ ξ n x) : + submartingale (λ n : ℕ, ∑ k in finset.range n, ξ k * (f (k + 1) - f k)) 𝒢 μ := +begin + have hξbdd : ∀ i, ∃ (C : ℝ), ∀ (x : α), |ξ i x| ≤ C := + λ i, ⟨R, λ x, (abs_of_nonneg (hnonneg i x)).trans_le (hbdd i x)⟩, + have hint : ∀ m, integrable (∑ k in finset.range m, ξ k * (f (k + 1) - f k)) μ := + λ m, integrable_finset_sum' _ + (λ i hi, integrable.bdd_mul ((hf.integrable _).sub (hf.integrable _)) + hξ.strongly_measurable.ae_strongly_measurable (hξbdd _)), + have hadp : adapted 𝒢 (λ (n : ℕ), ∑ (k : ℕ) in finset.range n, ξ k * (f (k + 1) - f k)), + { intro m, + refine finset.strongly_measurable_sum' _ (λ i hi, _), + rw finset.mem_range at hi, + exact (hξ.strongly_measurable_le hi.le).mul + ((hf.adapted.strongly_measurable_le (nat.succ_le_of_lt hi)).sub + (hf.adapted.strongly_measurable_le hi.le)) }, + refine submartingale_of_condexp_sub_nonneg_nat hadp hint (λ i, _), + simp only [← finset.sum_Ico_eq_sub _ (nat.le_succ _), finset.sum_apply, pi.mul_apply, + pi.sub_apply, nat.Ico_succ_singleton, finset.sum_singleton], + exact eventually_le.trans (eventually_le.mul_nonneg (eventually_of_forall (hnonneg _)) + (hf.condexp_sub_nonneg (nat.le_succ _))) (condexp_strongly_measurable_mul (hξ _) + (((hf.integrable _).sub (hf.integrable _)).bdd_mul + hξ.strongly_measurable.ae_strongly_measurable (hξbdd _)) + ((hf.integrable _).sub (hf.integrable _))).symm.le, +end + +/-- Given a discrete submartingale `f` and a predictable process `ξ` (i.e. `ξ (n + 1)` is adapted) +the process defined by `λ n, ∑ k in finset.range n, ξ (k + 1) * (f (k + 1) - f k)` is also a +submartingale. -/ +lemma submartingale.sum_mul_sub' [is_finite_measure μ] {R : ℝ} {ξ f : ℕ → α → ℝ} + (hf : submartingale f 𝒢 μ) (hξ : adapted 𝒢 (λ n, ξ (n + 1))) + (hbdd : ∀ n x, ξ n x ≤ R) (hnonneg : ∀ n x, 0 ≤ ξ n x) : + submartingale (λ n : ℕ, ∑ k in finset.range n, ξ (k + 1) * (f (k + 1) - f k)) 𝒢 μ := +hf.sum_mul_sub hξ (λ n, hbdd _) (λ n, hnonneg _) + end nat end measure_theory diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index ce4f010b0544c..cba03e7f8f1e9 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -238,17 +238,27 @@ def adapted (f : filtration ι m) (u : ι → α → β) : Prop := namespace adapted -lemma add [has_add β] [has_continuous_add β] (hu : adapted f u) (hv : adapted f v) : - adapted f (u + v) := -λ i, (hu i).add (hv i) +@[protected, to_additive] lemma mul [has_mul β] [has_continuous_mul β] + (hu : adapted f u) (hv : adapted f v) : + adapted f (u * v) := +λ i, (hu i).mul (hv i) -lemma neg [add_group β] [topological_add_group β] (hu : adapted f u) : adapted f (-u) := -λ i, (hu i).neg +@[protected, to_additive] lemma inv [group β] [topological_group β] (hu : adapted f u) : + adapted f u⁻¹ := +λ i, (hu i).inv -lemma smul [has_smul ℝ β] [has_continuous_smul ℝ β] (c : ℝ) (hu : adapted f u) : +@[protected] lemma smul [has_smul ℝ β] [has_continuous_smul ℝ β] (c : ℝ) (hu : adapted f u) : adapted f (c • u) := λ i, (hu i).const_smul c +@[protected] lemma strongly_measurable {i : ι} (hf : adapted f u) : + strongly_measurable[m] (u i) := +(hf i).mono (f.le i) + +lemma strongly_measurable_le {i j : ι} (hf : adapted f u) (hij : i ≤ j) : + strongly_measurable[f j] (u i) := +(hf i).mono (f.mono hij) + end adapted variable (β) From ae48e8cc3e24237396d11bdeb8ad925b4896dd22 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 27 Jul 2022 19:44:50 +0000 Subject: [PATCH 114/173] feat(analysis/special_functions/gaussian): formula for gaussian integrals (#15106) --- src/analysis/special_functions/gaussian.lean | 171 +++++++++++++++++++ 1 file changed, 171 insertions(+) create mode 100644 src/analysis/special_functions/gaussian.lean diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean new file mode 100644 index 0000000000000..f270da52503f3 --- /dev/null +++ b/src/analysis/special_functions/gaussian.lean @@ -0,0 +1,171 @@ +/- +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 + +/-! +# Gaussian integral + +We prove the formula `∫ x, exp (-b * x^2) = sqrt (π / b)`, in `integral_gaussian`. +-/ + +noncomputable theory + +open real set measure_theory filter asymptotics +open_locale real topological_space + +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 + have A : (λ (x : ℝ), -x - -b * x ^ 2) = (λ x, x * (b * x + (- 1))), by { ext x, ring }, + rw [is_o_exp_comp_exp_comp, A], + apply tendsto.at_top_mul_at_top tendsto_id, + apply tendsto_at_top_add_const_right at_top (-1 : ℝ), + exact tendsto.const_mul_at_top hb tendsto_id, +end + +lemma rpow_mul_exp_neg_mul_sq_is_o_exp_neg {b : ℝ} (hb : 0 < b) (s : ℝ) : + (λ x:ℝ, x ^ s * exp (-b * x^2)) =o[at_top] (λ x:ℝ, exp (-(1/2) * x)) := +begin + apply ((is_O_refl (λ x:ℝ, x ^ s) at_top).mul_is_o (exp_neg_mul_sq_is_o_exp_neg hb)).trans, + convert Gamma_integrand_is_o s, + simp_rw [mul_comm], +end + +lemma integrable_on_rpow_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) {s : ℝ} (hs : -1 < s) : + integrable_on (λ x:ℝ, x ^ s * exp (-b * x^2)) (Ioi 0) := +begin + rw [← Ioc_union_Ioi_eq_Ioi (zero_le_one : (0 : ℝ) ≤ 1), integrable_on_union], + split, + { rw [←integrable_on_Icc_iff_integrable_on_Ioc], + refine integrable_on.mul_continuous_on _ _ is_compact_Icc, + { refine (interval_integrable_iff_integrable_Icc_of_le zero_le_one).mp _, + exact interval_integral.interval_integrable_rpow' hs }, + { exact (continuous_exp.comp (continuous_const.mul (continuous_pow 2))).continuous_on } }, + { have B : (0 : ℝ) < 1/2, by norm_num, + apply integrable_of_is_O_exp_neg B _ (is_o.is_O (rpow_mul_exp_neg_mul_sq_is_o_exp_neg hb _)), + assume x hx, + have N : x ≠ 0, { refine (zero_lt_one.trans_le _).ne', exact hx }, + apply ((continuous_at_rpow_const _ _ (or.inl N)).mul _).continuous_within_at, + exact (continuous_exp.comp (continuous_const.mul (continuous_pow 2))).continuous_at }, +end + +lemma integrable_rpow_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) {s : ℝ} (hs : -1 < s) : + integrable (λ x:ℝ, x ^ s * exp (-b * x^2)) := +begin + rw [← integrable_on_univ, ← @Iio_union_Ici _ _ (0 : ℝ), integrable_on_union, + integrable_on_Ici_iff_integrable_on_Ioi], + refine ⟨_, integrable_on_rpow_mul_exp_neg_mul_sq hb hs⟩, + rw ← (measure.measure_preserving_neg (volume : measure ℝ)).integrable_on_comp_preimage + ((homeomorph.neg ℝ).to_measurable_equiv.measurable_embedding), + simp only [function.comp, neg_sq, neg_preimage, preimage_neg_Iio, neg_neg, neg_zero], + apply integrable.mono' (integrable_on_rpow_mul_exp_neg_mul_sq hb hs), + { apply measurable.ae_strongly_measurable, + exact (measurable_id'.neg.pow measurable_const).mul + ((measurable_id'.pow measurable_const).const_mul (-b)).exp }, + { have : measurable_set (Ioi (0 : ℝ)) := measurable_set_Ioi, + filter_upwards [ae_restrict_mem this] with x hx, + have h'x : 0 ≤ x := le_of_lt hx, + rw [real.norm_eq_abs, abs_mul, abs_of_nonneg (exp_pos _).le], + apply mul_le_mul_of_nonneg_right _ (exp_pos _).le, + simpa [abs_of_nonneg h'x] using abs_rpow_le_abs_rpow (-x) s } +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 + +lemma integrable_exp_neg_mul_sq_iff {b : ℝ} : + integrable (λ x:ℝ, exp (-b * x^2)) ↔ 0 < b := +begin + refine ⟨λ h, _, integrable_exp_neg_mul_sq⟩, + by_contra' hb, + have : ∫⁻ x:ℝ, 1 ≤ ∫⁻ x:ℝ, ∥exp (-b * x^2)∥₊, + { apply lintegral_mono (λ x, _), + simp only [neg_mul, ennreal.one_le_coe_iff, ← to_nnreal_one, to_nnreal_le_iff_le_coe, + real.norm_of_nonneg (exp_pos _).le, coe_nnnorm, one_le_exp_iff, right.nonneg_neg_iff], + exact mul_nonpos_of_nonpos_of_nonneg hb (sq_nonneg _) }, + simpa using this.trans_lt h.2, +end + +lemma integrable_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : + integrable (λ x:ℝ, x * exp (-b * x^2)) := +begin + have A : (-1 : ℝ) < 1, by norm_num, + convert integrable_rpow_mul_exp_neg_mul_sq hb A, + simp, +end + +lemma integral_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : + ∫ r in Ioi 0, r * exp (-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'], + 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 + (𝓝 ((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) }, + simpa using L, +end + +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], + { 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. -/ + calc + (∫ x, real.exp (-b * x^2)) ^ 2 + = ∫ p : ℝ × ℝ, exp (-b * p.1 ^ 2) * exp (-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) : + begin + rw ← set_integral_prod_mul, + congr' with p, + rw mul_one, + congr, + conv_rhs { rw [← one_mul (p.1^2), ← sin_sq_add_cos_sq p.2], }, + ring_exp, + end + ... = π / 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'], + ring, + end + ... = (sqrt (π / b)) ^ 2 : + by { rw sq_sqrt, exact div_nonneg pi_pos.le hb.le } +end From 39256a6a756d1a58baf42e9fa941117d926b1810 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Jul 2022 19:44:51 +0000 Subject: [PATCH 115/173] feat(combinatorics/additive/behrend): Behrend's lower bound on Roth numbers (#15327) This proves the explicit bound on Roth's numbers of the form `n / exp (O (sqrt (log n)))`. Co-authored-by: Bhavik Mehta --- src/algebra/order/floor.lean | 4 +- src/analysis/special_functions/pow.lean | 2 + src/combinatorics/additive/behrend.lean | 340 ++++++++++++++++++++++++ 3 files changed, 345 insertions(+), 1 deletion(-) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 78f15aa2d5227..25af371180a3d 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -199,6 +199,8 @@ eq_of_forall_ge_iff $ λ a, ceil_le.trans nat.cast_le @[simp] lemma ceil_eq_zero : ⌈a⌉₊ = 0 ↔ a ≤ 0 := by rw [← le_zero_iff, ceil_le, nat.cast_zero] +@[simp] lemma ceil_pos : 0 < ⌈a⌉₊ ↔ 0 < a := by rw [lt_ceil, cast_zero] + lemma lt_of_ceil_lt (h : ⌈a⌉₊ < n) : a < n := (le_ceil a).trans_lt (nat.cast_lt.2 h) lemma le_of_ceil_le (h : ⌈a⌉₊ ≤ n) : a ≤ n := (le_ceil a).trans (nat.cast_le.2 h) @@ -647,7 +649,7 @@ by rw [eq_sub_iff_add_eq, ← ceil_add_one, sub_add_cancel] lemma ceil_lt_add_one (a : α) : (⌈a⌉ : α) < a + 1 := by { rw [← lt_ceil, ← int.cast_one, ceil_add_int], apply lt_add_one } -lemma ceil_pos : 0 < ⌈a⌉ ↔ 0 < a := by rw [lt_ceil, int.cast_zero] +@[simp] lemma ceil_pos : 0 < ⌈a⌉ ↔ 0 < a := by rw [lt_ceil, cast_zero] @[simp] lemma ceil_zero : ⌈(0 : α)⌉ = 0 := by rw [← int.cast_zero, ceil_coe] diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 93a4772a168c1..5b6eacd69bc38 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -294,6 +294,8 @@ 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] } diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index 872e74b79b0ac..ee7dc823d7a72 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -5,6 +5,8 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import analysis.inner_product_space.pi_L2 import combinatorics.additive.salem_spencer +import combinatorics.pigeonhole +import data.complex.exponential_bounds /-! # Behrend's bound on Roth numbers @@ -26,6 +28,7 @@ integer points on that sphere and map them onto `ℕ` in a way that preserves ar digits in base `d`. * `behrend.card_sphere_le_roth_number_nat`: Implicit lower bound on Roth numbers in terms of `behrend.sphere`. +* `behrend.roth_lower_bound`: Behrend's explicit lower bound on Roth numbers. ## References @@ -196,4 +199,341 @@ begin exact λ h _ i, (h i).trans_le le_self_add, end +/-! +### Optimization + +Now that we know how to turn the integer points of any sphere into a Salem-Spencer set, we find a +sphere containing many integer points by the pigeonhole principle. This gives us an implicit bound +that we then optimize by tweaking the parameters. The (almost) optimal parameters are +`behrend.n_value` and `behrend.d_value`. +-/ + +lemma exists_large_sphere_aux (n d : ℕ) : + ∃ k ∈ range (n * (d - 1)^2 + 1), (↑(d ^ n) / (↑(n * (d - 1)^2) + 1) : ℝ) ≤ (sphere n d k).card := +begin + refine exists_le_card_fiber_of_nsmul_le_card_of_maps_to (λ x hx, _) nonempty_range_succ _, + { rw [mem_range, lt_succ_iff], + exact sum_sq_le_of_mem_box hx }, + { rw [card_range, _root_.nsmul_eq_mul, mul_div_assoc', cast_add_one, mul_div_cancel_left, + card_box], + exact (cast_add_one_pos _).ne' } +end + +lemma exists_large_sphere (n d : ℕ) : ∃ k, (d ^ n / ↑(n * d^2) : ℝ) ≤ (sphere n d k).card := +begin + obtain ⟨k, -, hk⟩ := exists_large_sphere_aux n d, + refine ⟨k, _⟩, + obtain rfl | hn := n.eq_zero_or_pos, + { simp }, + obtain rfl | hd := d.eq_zero_or_pos, + { simp }, + rw ←cast_pow, + refine (div_le_div_of_le_left _ _ _).trans hk, + { exact cast_nonneg _ }, + { exact cast_add_one_pos _ }, + simp only [←le_sub_iff_add_le', cast_mul, ←mul_sub, cast_pow, cast_sub hd, sub_sq, + one_pow, cast_one, mul_one, sub_add, sub_sub_self], + apply one_le_mul_of_one_le_of_one_le, + { rwa one_le_cast }, + rw le_sub_iff_add_le, + norm_num, + exact le_mul_of_one_le_right zero_le_two (one_le_cast.2 hd), +end + +lemma bound_aux' (n d : ℕ) : (d ^ n / ↑(n * d^2) : ℝ) ≤ roth_number_nat ((2 * d - 1)^n) := +let ⟨k, h⟩ := exists_large_sphere n d in h.trans $ cast_le.2 $ card_sphere_le_roth_number_nat _ _ _ + +lemma bound_aux (hd : d ≠ 0) (hn : 2 ≤ n) : + (d ^ (n - 2) / n : ℝ) ≤ roth_number_nat ((2 * d - 1)^n) := +begin + convert bound_aux' n d using 1, + rw [cast_mul, cast_pow, mul_comm, ←div_div, pow_sub₀ _ _ hn, ←div_eq_mul_inv], + rwa cast_ne_zero, +end + +open_locale filter topological_space +open real + +section numerical_bounds + +lemma log_two_mul_two_le_sqrt_log_eight : log 2 * 2 ≤ sqrt (log 8) := +begin + rw [show (8 : ℝ) = 2 ^ ((3 : ℕ) : ℝ), by norm_num1, log_rpow zero_lt_two (3:ℕ)], + apply le_sqrt_of_sq_le, + rw [mul_pow, sq (log 2), mul_assoc, mul_comm], + refine mul_le_mul_of_nonneg_right _ (log_nonneg one_le_two), + rw ←le_div_iff, + apply log_two_lt_d9.le.trans, + all_goals { norm_num1 } +end + +lemma two_div_one_sub_two_div_e_le_eight : 2 / (1 - 2 / exp 1) ≤ 8 := +begin + rw [div_le_iff, mul_sub, mul_one, mul_div_assoc', le_sub, div_le_iff (exp_pos _)], + { linarith [exp_one_gt_d9] }, + rw [sub_pos, div_lt_one]; + exact exp_one_gt_d9.trans' (by norm_num), +end + +lemma le_sqrt_log (hN : 4096 ≤ N) : log (2 / (1 - 2 / exp 1)) * (69 / 50) ≤ sqrt (log ↑N) := +begin + have : ((12 : ℕ) : ℝ) * log 2 ≤ log N, + { rw [←log_rpow zero_lt_two, log_le_log, rpow_nat_cast], + { norm_num1, + exact_mod_cast hN }, + { exact rpow_pos_of_pos zero_lt_two _ }, + rw cast_pos, + exact hN.trans_lt' (by norm_num1) }, + refine (mul_le_mul_of_nonneg_right ((log_le_log _ $ by norm_num1).2 + two_div_one_sub_two_div_e_le_eight) $ by norm_num1).trans (_), + { refine div_pos zero_lt_two _, + rw [sub_pos, div_lt_one (exp_pos _)], + exact exp_one_gt_d9.trans_le' (by norm_num1) }, + have l8 : log 8 = (3 : ℕ) * log 2, + { rw [←log_rpow zero_lt_two, rpow_nat_cast], + norm_num }, + rw [l8, cast_bit1, cast_one], + apply le_sqrt_of_sq_le (le_trans _ this), + simp only [cast_bit0, cast_bit1, cast_one], + rw [mul_right_comm, mul_pow, sq (log 2), ←mul_assoc], + apply mul_le_mul_of_nonneg_right _ (log_nonneg one_le_two), + rw ←le_div_iff' , + { exact log_two_lt_d9.le.trans (by norm_num1) }, + exact sq_pos_of_ne_zero _ (by norm_num1), +end + +lemma exp_neg_two_mul_le {x : ℝ} (hx : 0 < x) : exp (-2 * x) < exp (2 - ⌈x⌉₊) / ⌈x⌉₊ := +begin + have h₁ := ceil_lt_add_one hx.le, + have h₂ : 1 - x ≤ 2 - ⌈x⌉₊, + { rw le_sub_iff_add_le, + apply (add_le_add_left h₁.le _).trans_eq, + rw [←add_assoc, sub_add_cancel], + refl }, + have h₃ : exp (-(x+1)) ≤ 1 / (x + 1), + { rw [exp_neg, inv_eq_one_div], + refine one_div_le_one_div_of_le (add_pos hx zero_lt_one) _, + apply le_trans _ (add_one_le_exp_of_nonneg $ add_nonneg hx.le zero_le_one), + exact le_add_of_nonneg_right zero_le_one }, + refine lt_of_le_of_lt _ (div_lt_div_of_lt_left (exp_pos _) (cast_pos.2 $ ceil_pos.2 hx) h₁), + refine le_trans _ (div_le_div_of_le_of_nonneg (exp_le_exp.2 h₂) $ add_nonneg hx.le zero_le_one), + rw [le_div_iff (add_pos hx zero_lt_one), ←le_div_iff' (exp_pos _), ←exp_sub, neg_mul, + sub_neg_eq_add, two_mul, sub_add_add_cancel, add_comm _ x], + refine le_trans _ (add_one_le_exp_of_nonneg $ add_nonneg hx.le zero_le_one), + exact le_add_of_nonneg_right zero_le_one, +end + +lemma div_lt_floor {x : ℝ} (hx : 2 / (1 - 2 / exp 1) ≤ x) : x / exp 1 < (⌊x/2⌋₊ : ℝ) := +begin + apply lt_of_le_of_lt _ (sub_one_lt_floor _), + have : 0 < 1 - 2 / exp 1, + { rw [sub_pos, div_lt_one (exp_pos _)], + exact lt_of_le_of_lt (by norm_num) exp_one_gt_d9 }, + rwa [le_sub, div_eq_mul_one_div x, div_eq_mul_one_div x, ←mul_sub, div_sub', ←div_eq_mul_one_div, + mul_div_assoc', one_le_div, ←div_le_iff this], + { exact zero_lt_two }, + { exact two_ne_zero } +end + +lemma ceil_lt_mul {x : ℝ} (hx : 50/19 ≤ x) : (⌈x⌉₊ : ℝ) < 1.38 * x := +begin + refine (ceil_lt_add_one $ hx.trans' $ by norm_num).trans_le _, + rwa [←le_sub_iff_add_le', ←sub_one_mul, show (69/50 - 1 : ℝ) = (50/19)⁻¹, by norm_num1, + ←div_eq_inv_mul, one_le_div], + norm_num1, +end + +end numerical_bounds + +/-- The (almost) optimal value of `n` in `behrend.bound_aux`. -/ +noncomputable def n_value (N : ℕ) : ℕ := ⌈sqrt (log N)⌉₊ + +/-- The (almost) optimal value of `d` in `behrend.bound_aux`. -/ +noncomputable def d_value (N : ℕ) : ℕ := ⌊(N : ℝ)^(1 / n_value N : ℝ)/2⌋₊ + +lemma n_value_pos (hN : 2 ≤ N) : 0 < n_value N := +ceil_pos.2 $ real.sqrt_pos.2 $ log_pos $ one_lt_cast.2 $ hN + +lemma two_le_n_value (hN : 3 ≤ N) : 2 ≤ n_value N := +begin + refine succ_le_of_lt (lt_ceil.2 $ lt_sqrt_of_sq_lt _), + rw [cast_one, one_pow, lt_log_iff_exp_lt], + refine lt_of_lt_of_le _ (cast_le.2 hN), + { exact exp_one_lt_d9.trans_le (by norm_num) }, + rw cast_pos, + exact (zero_lt_succ _).trans_le hN, +end + +lemma three_le_n_value (hN : 64 ≤ N) : 3 ≤ n_value N := +begin + rw [n_value, ←lt_iff_add_one_le, lt_ceil, cast_two], + apply lt_sqrt_of_sq_lt, + have : (2 : ℝ)^((6 : ℕ) : ℝ) ≤ N, + { rw rpow_nat_cast, + exact (cast_le.2 hN).trans' (by norm_num1) }, + apply lt_of_lt_of_le _ ((log_le_log (rpow_pos_of_pos zero_lt_two _) _).2 this), + rw [log_rpow zero_lt_two, cast_bit0, cast_bit1, cast_one, ←div_lt_iff'], + { exact log_two_gt_d9.trans_le' (by norm_num1) }, + { norm_num1 }, + rw cast_pos, + exact hN.trans_lt' (by norm_num1), +end + +lemma d_value_pos (hN₃ : 8 ≤ N) : 0 < d_value N := +begin + have hN₀ : 0 < (N : ℝ) := cast_pos.2 (succ_pos'.trans_le hN₃), + rw [d_value, floor_pos, ←log_le_log zero_lt_one, log_one, log_div _ two_ne_zero, log_rpow hN₀, + div_mul_eq_mul_div, one_mul, sub_nonneg, le_div_iff], + { have : (n_value N : ℝ) ≤ 2 * sqrt (log N), + { apply (ceil_lt_add_one $ sqrt_nonneg _).le.trans, + rw [two_mul, add_le_add_iff_left], + apply le_sqrt_of_sq_le, + rw [one_pow, le_log_iff_exp_le hN₀], + exact (exp_one_lt_d9.le.trans $ by norm_num).trans (cast_le.2 hN₃) }, + apply (mul_le_mul_of_nonneg_left this $ log_nonneg one_le_two).trans _, + rw [←mul_assoc, ←le_div_iff (real.sqrt_pos.2 $ log_pos $ one_lt_cast.2 _), div_sqrt], + { apply log_two_mul_two_le_sqrt_log_eight.trans, + apply real.sqrt_le_sqrt, + rw log_le_log _ hN₀, + { exact_mod_cast hN₃ }, + { norm_num } }, + exact hN₃.trans_lt' (by norm_num) }, + { exact cast_pos.2 (n_value_pos $ hN₃.trans' $ by norm_num) }, + { exact (rpow_pos_of_pos hN₀ _).ne' }, + { exact div_pos (rpow_pos_of_pos hN₀ _) zero_lt_two }, +end + +lemma le_N (hN : 2 ≤ N) : (2 * (d_value N) - 1)^(n_value N) ≤ N := +begin + have : (2 * d_value N - 1)^(n_value N) ≤ (2 * d_value N)^(n_value N) := + nat.pow_le_pow_of_le_left (nat.sub_le _ _) _, + apply this.trans, + suffices : ((2 * d_value N)^n_value N : ℝ) ≤ N, by exact_mod_cast this, + rw ←rpow_nat_cast, + suffices i : (2 * d_value N : ℝ) ≤ (N : ℝ)^(1/n_value N : ℝ), + { apply (rpow_le_rpow (mul_nonneg zero_le_two (cast_nonneg _)) i (cast_nonneg _)).trans, + rw [←rpow_mul (cast_nonneg _), one_div_mul_cancel, rpow_one], + rw cast_ne_zero, + apply (n_value_pos hN).ne', }, + rw ←le_div_iff', + { exact floor_le (div_nonneg (rpow_nonneg_of_nonneg (cast_nonneg _) _) zero_le_two) }, + apply zero_lt_two +end + +lemma bound (hN : 4096 ≤ N) : (N : ℝ)^(1/n_value N : ℝ) / exp 1 < d_value N := +begin + apply div_lt_floor _, + rw [←log_le_log, log_rpow, mul_comm, ←div_eq_mul_one_div], + { apply le_trans _ (div_le_div_of_le_left _ _ (ceil_lt_mul _).le), + rw [mul_comm, ←div_div, div_sqrt, le_div_iff], + { exact le_sqrt_log hN }, + { norm_num1 }, + { apply log_nonneg, + rw one_le_cast, + exact hN.trans' (by norm_num1) }, + { rw [cast_pos, lt_ceil, cast_zero, real.sqrt_pos], + apply log_pos, + rw one_lt_cast, + exact hN.trans_lt' (by norm_num1) }, + apply le_sqrt_of_sq_le, + have : ((12 : ℕ) : ℝ) * log 2 ≤ log N, + { rw [←log_rpow zero_lt_two, log_le_log, rpow_nat_cast], + { norm_num1, + exact_mod_cast hN }, + { exact rpow_pos_of_pos zero_lt_two _ }, + rw cast_pos, + exact hN.trans_lt' (by norm_num1) }, + refine le_trans _ this, + simp only [cast_bit0, cast_bit1, cast_one], + rw ←div_le_iff', + { exact log_two_gt_d9.le.trans' (by norm_num1) }, + { norm_num1 } }, + { rw cast_pos, + exact hN.trans_lt' (by norm_num1) }, + { refine div_pos zero_lt_two _, + rw [sub_pos, div_lt_one (exp_pos _)], + exact lt_of_le_of_lt (by norm_num1) exp_one_gt_d9 }, + apply rpow_pos_of_pos, + rw cast_pos, + exact hN.trans_lt' (by norm_num1), +end + +lemma roth_lower_bound_explicit (hN : 4096 ≤ N) : + (N : ℝ) * exp (-4 * sqrt (log N)) < roth_number_nat N := +begin + let n := n_value N, + have hn : 0 < (n : ℝ) := cast_pos.2 (n_value_pos $ hN.trans' $ by norm_num1), + have hd : 0 < d_value N := d_value_pos (hN.trans' $ by norm_num1), + have hN₀ : 0 < (N : ℝ) := cast_pos.2 (hN.trans' $ by norm_num1), + have hn₂ : 2 ≤ n := two_le_n_value (hN.trans' $ by norm_num1), + have : (2 * d_value N - 1)^n ≤ N := le_N (hN.trans' $ by norm_num1), + refine ((bound_aux hd.ne' hn₂).trans $ cast_le.2 $ roth_number_nat.mono this).trans_lt' _, + refine (div_lt_div_of_lt hn $ pow_lt_pow_of_lt_left (bound hN) _ _).trans_le' _, + { exact div_nonneg (rpow_nonneg_of_nonneg (cast_nonneg _) _) (exp_pos _).le }, + { exact tsub_pos_of_lt (three_le_n_value $ hN.trans' $ by norm_num1) }, + rw [←rpow_nat_cast, div_rpow (rpow_nonneg_of_nonneg hN₀.le _) (exp_pos _).le, ←rpow_mul hN₀.le, + mul_comm (_ / _), mul_one_div, cast_sub hn₂, cast_two, same_sub_div hn.ne', exp_one_rpow, + div_div, rpow_sub hN₀, rpow_one, div_div, div_eq_mul_inv], + refine mul_le_mul_of_nonneg_left _ (cast_nonneg _), + rw [mul_inv, mul_inv, ←exp_neg, ←rpow_neg (cast_nonneg _), neg_sub, ←div_eq_mul_inv], + have : exp ((-4) * sqrt (log N)) = exp (-2 * sqrt (log N)) * exp (-2 * sqrt (log N)), + { rw [←exp_add, ←add_mul], + norm_num }, + rw this, + refine (mul_le_mul _ (exp_neg_two_mul_le $ real.sqrt_pos.2 $ log_pos _).le (exp_pos _).le $ + rpow_nonneg_of_nonneg (cast_nonneg _) _), + { rw [←le_log_iff_exp_le (rpow_pos_of_pos hN₀ _), log_rpow hN₀, ←le_div_iff, mul_div_assoc, + div_sqrt, neg_mul, neg_le_neg_iff, div_mul_eq_mul_div, div_le_iff hn], + { exact mul_le_mul_of_nonneg_left (le_ceil _) zero_le_two }, + refine real.sqrt_pos.2 (log_pos _), + rw one_lt_cast, + exact hN.trans_lt' (by norm_num1) }, + { rw one_lt_cast, + exact hN.trans_lt' (by norm_num1) } +end + +lemma exp_four_lt : exp 4 < 64 := +begin + rw [show (64 : ℝ) = 2 ^ ((6 : ℕ) : ℝ), by norm_num1, + ←lt_log_iff_exp_lt (rpow_pos_of_pos zero_lt_two _), log_rpow zero_lt_two, ←div_lt_iff'], + exact log_two_gt_d9.trans_le' (by norm_num1), + norm_num +end + +lemma four_zero_nine_six_lt_exp_sixteen : 4096 < exp 16 := +begin + rw [←log_lt_iff_lt_exp (show (0 : ℝ) < 4096, by norm_num), show (4096 : ℝ) = 2 ^ 12, by norm_num, + ←rpow_nat_cast, log_rpow zero_lt_two, cast_bit0, cast_bit0, cast_bit1, cast_one], + linarith [log_two_lt_d9], +end + +lemma lower_bound_le_one' (hN : 2 ≤ N) (hN' : N ≤ 4096) : (N : ℝ) * exp (-4 * sqrt (log N)) ≤ 1 := +begin + rw [←log_le_log (mul_pos (cast_pos.2 (zero_lt_two.trans_le hN)) (exp_pos _)) zero_lt_one, + log_one, log_mul (cast_pos.2 (zero_lt_two.trans_le hN)).ne' (exp_pos _).ne', log_exp, + neg_mul, ←sub_eq_add_neg, sub_nonpos, ←div_le_iff (real.sqrt_pos.2 $ log_pos $ + one_lt_cast.2 $ one_lt_two.trans_le hN), div_sqrt, sqrt_le_left + (zero_le_bit0.2 zero_le_two), log_le_iff_le_exp (cast_pos.2 (zero_lt_two.trans_le hN))], + norm_num1, + apply le_trans _ four_zero_nine_six_lt_exp_sixteen.le, + exact_mod_cast hN', +end + +lemma lower_bound_le_one (hN : 1 ≤ N) (hN' : N ≤ 4096) : (N : ℝ) * exp (-4 * sqrt (log N)) ≤ 1 := +begin + obtain rfl | hN := hN.eq_or_lt, + { norm_num }, + { exact lower_bound_le_one' hN hN' } +end + +lemma roth_lower_bound : (N : ℝ) * exp (-4 * sqrt (log N)) ≤ roth_number_nat N := +begin + obtain rfl | hN := nat.eq_zero_or_pos N, + { norm_num }, + obtain h₁ | h₁ := le_or_lt 4096 N, + { exact (roth_lower_bound_explicit h₁).le }, + { apply (lower_bound_le_one hN h₁.le).trans, + simpa using roth_number_nat.monotone hN } +end + end behrend From ee231101cd851919b8107cf515a1fa0119ab83ce Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 27 Jul 2022 19:44:52 +0000 Subject: [PATCH 116/173] feat(algebraic_geometry/morphisms): Quasi-compact morphisms of schemes (#15436) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/AffineScheme.lean | 10 +- .../morphisms/quasi_compact.lean | 117 ++++++++++++++++++ src/algebraic_geometry/open_immersion.lean | 23 ++++ src/topology/sets/opens.lean | 13 ++ src/topology/subset_properties.lean | 33 +++++ 5 files changed, 191 insertions(+), 5 deletions(-) create mode 100644 src/algebraic_geometry/morphisms/quasi_compact.lean diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 79a128d633e17..4444625b845fe 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -110,6 +110,10 @@ end AffineScheme def is_affine_open {X : Scheme} (U : opens X.carrier) : Prop := is_affine (X.restrict U.open_embedding) +/-- The set of affine opens as a subset of `opens X.carrier`. -/ +def Scheme.affine_opens (X : Scheme) : set (opens X.carrier) := +{ U : opens X.carrier | is_affine_open U } + lemma range_is_affine_open_of_open_immersion {X Y : Scheme} [is_affine X] (f : X ⟶ Y) [H : is_open_immersion f] : is_affine_open ⟨set.range f.1.base, H.base_open.open_range⟩ := begin @@ -130,7 +134,7 @@ instance Scheme.affine_basis_cover_is_affine (X : Scheme) (i : X.affine_basis_co algebraic_geometry.Spec_is_affine _ lemma is_basis_affine_open (X : Scheme) : - opens.is_basis { U : opens X.carrier | is_affine_open U } := + opens.is_basis X.affine_opens := begin rw opens.is_basis_iff_nbhd, rintros U x (hU : x ∈ (U : set X.carrier)), @@ -509,10 +513,6 @@ begin refl end -/-- The collection of affine open sets of `X`. -/ -def Scheme.affine_opens (X : Scheme) : set (opens X.carrier) := -{ U : opens X.carrier | is_affine_open U } - /-- The basic open set of a section `f` on an an affine open as an `X.affine_opens`. -/ abbreviation Scheme.affine_basic_open (X : Scheme) {U : X.affine_opens} (f : X.presheaf.obj $ op U) : X.affine_opens := diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean new file mode 100644 index 0000000000000..c9babc0c32634 --- /dev/null +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -0,0 +1,117 @@ +/- +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.AffineScheme +import topology.spectral.hom + +/-! +# Quasi-compact morphisms + +A morphism of schemes is quasi-compact if the preimages of quasi-compact open sets are +quasi-compact. + +It suffices to check that preimages of affine open sets are compact +(`quasi_compact_iff_forall_affine`). + +-/ + +noncomputable theory + +open category_theory category_theory.limits opposite topological_space + +universe u + +namespace algebraic_geometry + +variables {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- +A morphism is `quasi-compact` if the underlying map of topological spaces is, i.e. if the preimages +of quasi-compact open sets are quasi-compact. +-/ +@[mk_iff] +class quasi_compact (f : X ⟶ Y) : Prop := +(is_compact_preimage : ∀ U : set Y.carrier, is_open U → is_compact U → is_compact (f.1.base ⁻¹' U)) + +lemma quasi_compact_iff_spectral : quasi_compact f ↔ is_spectral_map f.1.base := +⟨λ ⟨h⟩, ⟨by continuity, h⟩, λ h, ⟨h.2⟩⟩ + +@[priority 900] +instance quasi_compact_of_is_iso {X Y : Scheme} (f : X ⟶ Y) [is_iso f] : quasi_compact f := +begin + constructor, + intros U hU hU', + convert hU'.image (inv f.1.base).continuous_to_fun using 1, + rw set.image_eq_preimage_of_inverse, + delta function.left_inverse, + exacts [is_iso.inv_hom_id_apply f.1.base, is_iso.hom_inv_id_apply f.1.base] +end + +instance quasi_compact_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [quasi_compact f] [quasi_compact g] : quasi_compact (f ≫ g) := +begin + constructor, + intros U hU hU', + rw [Scheme.comp_val_base, coe_comp, set.preimage_comp], + apply quasi_compact.is_compact_preimage, + { exact continuous.is_open_preimage (by continuity) _ hU }, + apply quasi_compact.is_compact_preimage; assumption +end + +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 + (coe : X.affine_opens → opens X.carrier), + { rw subtype.range_coe, exact is_basis_affine_open X }, + { intro i, exact 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 + +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, _⟩, + 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], + exact hS.compact_bUnion (λ i _, H i i.prop) +end + +@[elab_as_eliminator] +lemma compact_open_induction_on {P : opens X.carrier → Prop} (S : opens X.carrier) + (hS : is_compact S.1) + (h₁ : P ⊥) + (h₂ : ∀ (S : opens X.carrier) (hS : is_compact S.1) (U : X.affine_opens), P S → P (S ⊔ U)) : + P S := +begin + classical, + obtain ⟨s, hs, hs'⟩ := (is_compact_open_iff_eq_finset_affine_union S.1).mp ⟨hS, S.2⟩, + replace hs' : S = supr (λ i : s, (i : opens X.carrier)) := by { ext1, simpa using hs' }, + subst hs', + apply hs.induction_on, + { convert h₁, rw supr_eq_bot, rintro ⟨_, h⟩, exact h.elim }, + { intros x s h₃ hs h₄, + have : is_compact (⨆ i : s, (i : opens X.carrier)).1, + { refine ((is_compact_open_iff_eq_finset_affine_union _).mpr _).1, exact ⟨s, hs, by simp⟩ }, + convert h₂ _ this x h₄, + simp only [coe_coe], + rw [supr_subtype, sup_comm], + conv_rhs { rw supr_subtype }, + exact supr_insert } +end + +end algebraic_geometry diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index e6be7a87c13f1..ef8b042a2c6ef 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -1673,6 +1673,29 @@ def Scheme.open_cover.pullback_cover {X : Scheme} (𝒰 : X.open_cover) {W : Sch { rw ← Top.epi_iff_surjective, apply_instance } end } +lemma Scheme.open_cover.Union_range {X : Scheme} (𝒰 : X.open_cover) : + (⋃ i, set.range (𝒰.map i).1.base) = set.univ := +begin + rw set.eq_univ_iff_forall, + intros x, + rw set.mem_Union, + exact ⟨𝒰.f x, 𝒰.covers x⟩ +end + +lemma Scheme.open_cover.compact_space {X : Scheme} (𝒰 : X.open_cover) [finite 𝒰.J] + [H : ∀ i, compact_space (𝒰.obj i).carrier] : compact_space X.carrier := +begin + casesI nonempty_fintype 𝒰.J, + rw [← is_compact_univ_iff, ← 𝒰.Union_range], + apply compact_Union, + intro i, + rw is_compact_iff_compact_space, + exact @@homeomorph.compact_space _ _ (H i) + (Top.homeo_of_iso (as_iso (is_open_immersion.iso_of_range_eq (𝒰.map i) + (X.of_restrict (opens.open_embedding ⟨_, (𝒰.is_open i).base_open.open_range⟩)) + subtype.range_coe.symm).hom.1.base)) +end + /-- Given open covers `{ Uᵢ }` and `{ Uⱼ }`, we may form the open cover `{ Uᵢ ∩ Uⱼ }`. -/ def Scheme.open_cover.inter {X : Scheme.{u}} (𝒰₁ : Scheme.open_cover.{v₁} X) (𝒰₂ : Scheme.open_cover.{v₂} X) : X.open_cover := diff --git a/src/topology/sets/opens.lean b/src/topology/sets/opens.lean index 5a9eedf874e10..d1c609ee5179b 100644 --- a/src/topology/sets/opens.lean +++ b/src/topology/sets/opens.lean @@ -188,6 +188,19 @@ begin exact ⟨U, hUs Us, xU, le_Sup Us⟩ } 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)) + (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 + apply is_compact_open_iff_eq_finite_Union_of_is_topological_basis + (λ i : ι, (b i).1), + { convert hb, ext, simp }, + { exact hb' } +end + @[simp] lemma is_compact_element_iff (s : opens α) : complete_lattice.is_compact_element s ↔ is_compact (s : set α) := begin diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 981794e4b34c8..c66702eb6a70b 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -461,6 +461,39 @@ begin contradiction 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_topological_basis (b : ι → set α) + (hb : is_topological_basis (set.range b)) + (hb' : ∀ i, is_compact (b i)) (U : set α) : + is_compact U ∧ is_open U ↔ ∃ (s : set ι), s.finite ∧ U = ⋃ i ∈ s, b i := +begin + classical, + split, + { rintro ⟨h₁, h₂⟩, + obtain ⟨β, f, e, hf⟩ := hb.open_eq_Union h₂, + choose f' hf' using hf, + have : b ∘ f' = f := funext hf', subst this, + obtain ⟨t, ht⟩ := h₁.elim_finite_subcover (b ∘ f') + (λ i, hb.is_open (set.mem_range_self _)) (by rw e), + refine ⟨t.image f', set.finite.intro infer_instance, le_antisymm _ _⟩, + { refine set.subset.trans ht _, + simp only [set.Union_subset_iff, coe_coe], + intros i hi, + erw ← set.Union_subtype (λ x : ι, x ∈ t.image f') (λ i, b i.1), + exact set.subset_Union (λ i : t.image f', b i) ⟨_, finset.mem_image_of_mem _ hi⟩ }, + { apply set.Union₂_subset, + rintro i hi, + obtain ⟨j, hj, rfl⟩ := finset.mem_image.mp hi, + rw e, + exact set.subset_Union (b ∘ f') j } }, + { rintro ⟨s, hs, rfl⟩, + split, + { exact hs.compact_bUnion (λ i _, hb' i) }, + { apply is_open_bUnion, intros i hi, exact hb.is_open (set.mem_range_self _) } }, +end + + namespace filter /-- `filter.cocompact` is the filter generated by complements to compact sets. -/ From 868d1cf15046041b0687fe6dd750408fc4b0682a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Roberto=20=C3=81lvarez?= Date: Wed, 27 Jul 2022 22:36:27 +0000 Subject: [PATCH 117/173] =?UTF-8?q?feat(topology/subset=5Fproperties):=20`?= =?UTF-8?q?locally=5Fcompact=5Fspace`=20instance=20for=20`=CE=A0`=20types?= =?UTF-8?q?=20(#15707)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds - `locally_compact_space.pi` mirroring `locally_compact_space.prod` and - `locally_compact_space.pi_finite` for finite products Proof by: @alreadydone Co-authored-by: Junyan Xu --- src/topology/subset_properties.lean | 34 +++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index c66702eb6a70b..853e00436baea 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1051,6 +1051,40 @@ instance locally_compact_space.prod (α : Type*) (β : Type*) [topological_space have _ := λ x : α × β, (compact_basis_nhds x.1).prod_nhds' (compact_basis_nhds x.2), locally_compact_space_of_has_basis this $ λ x s ⟨⟨_, h₁⟩, _, h₂⟩, h₁.prod h₂ +section pi + +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) := +⟨λ t n hn, begin + rw [nhds_pi, filter.mem_pi] at hn, + obtain ⟨s, hs, n', hn', hsub⟩ := hn, + choose n'' hn'' hsub' hc using λ i, locally_compact_space.local_compact_nhds (t i) (n' i) (hn' i), + refine ⟨(set.univ : set ι).pi n'', _, subset_trans (λ _ h, _) hsub, is_compact_univ_pi hc⟩, + { exact (set_pi_mem_nhds_iff (@set.finite_univ ι _) _).mpr (λ i hi, hn'' i), }, + { exact λ i hi, hsub' i (h i trivial), }, +end⟩ + +/-- For spaces that are not Hausdorff. -/ +instance locally_compact_space.pi [∀ 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, + choose n'' hn'' hsub' hc using λ i, locally_compact_space.local_compact_nhds (t i) (n' i) (hn' i), + refine ⟨s.pi n'', _, subset_trans (λ _, _) hsub, _⟩, + { exact (set_pi_mem_nhds_iff hs _).mpr (λ i _, hn'' i), }, + { exact forall₂_imp (λ i hi hi', hsub' i hi'), }, + { rw ← set.univ_pi_ite, + refine is_compact_univ_pi (λ i, _), + by_cases i ∈ s, + { rw if_pos h, exact hc i, }, + { rw if_neg h, exact compact_space.compact_univ, } }, +end⟩ + +end pi + /-- A reformulation of the definition of locally compact space: In a locally compact space, every open set containing `x` has a compact subset containing `x` in its interior. -/ lemma exists_compact_subset [locally_compact_space α] {x : α} {U : set α} From 8044813bbf7beb9878f72487d43b78e8becc979c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 28 Jul 2022 00:18:08 +0000 Subject: [PATCH 118/173] feat(category_theory/limits/shapes/comm_sq): opposites of is_pullback and is_pushout (#15269) This PR shows that the dual of a pushout commutative square is a pullback square, and similar other results. The most basic `comm_sq` API is also moved to a separate file with fewer dependencies so as to prepare for a refactor of lifting properties which shall be based on `comm_sq`. --- src/category_theory/comm_sq.lean | 78 +++++++++++++++++++ .../limits/shapes/comm_sq.lean | 77 +++++++++++------- 2 files changed, 129 insertions(+), 26 deletions(-) create mode 100644 src/category_theory/comm_sq.lean diff --git a/src/category_theory/comm_sq.lean b/src/category_theory/comm_sq.lean new file mode 100644 index 0000000000000..95b3d528e095e --- /dev/null +++ b/src/category_theory/comm_sq.lean @@ -0,0 +1,78 @@ +/- +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.arrow + + +/-! +# Commutative squares + +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 +square is commutative. + +The structure `comm_sq` is extended in `category_theory/shapes/limits/comm_sq.lean` +as `is_pullback` and `is_pushout` in order to define pullback and pushout squares. + +## Future work + +Refactor `lift_struct` from `arrow.lean` and lifting properties using `comm_sq.lean`. + +-/ + +namespace category_theory + +variables {C : Type*} [category C] + +/-- The proposition that a square +``` + W ---f---> X + | | + g h + | | + v v + Y ---i---> Z + +``` +is a commuting square. +-/ +structure comm_sq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y ⟶ Z) : Prop := +(w : f ≫ h = g ≫ i) + +attribute [reassoc] comm_sq.w + +namespace comm_sq + +variables {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + +lemma flip (p : comm_sq f g h i) : comm_sq g f i h := ⟨p.w.symm⟩ + +lemma of_arrow {f g : arrow C} (h : f ⟶ g) : comm_sq f.hom h.left h.right g.hom := ⟨h.w.symm⟩ + +/-- The commutative square in the opposite category associated to a commutative square. -/ +lemma op (p : comm_sq f g h i) : comm_sq i.op h.op g.op f.op := +⟨by simp only [← op_comp, p.w]⟩ + +/-- The commutative square associated to a commutative square in the opposite category. -/ +lemma unop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + (p : comm_sq f g h i) : comm_sq i.unop h.unop g.unop f.unop := +⟨by simp only [← unop_comp, p.w]⟩ + +end comm_sq + +namespace functor + +variables {D : Type*} [category D] +variables (F : C ⥤ D) {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + +lemma map_comm_sq (s : comm_sq f g h i) : comm_sq (F.map f) (F.map g) (F.map h) (F.map i) := +⟨by simpa using congr_arg (λ k : W ⟶ Z, F.map k) s.w⟩ + +end functor + +alias functor.map_comm_sq ← comm_sq.map + +end category_theory diff --git a/src/category_theory/limits/shapes/comm_sq.lean b/src/category_theory/limits/shapes/comm_sq.lean index 30d7bc3df50ec..f3ff7bf35e394 100644 --- a/src/category_theory/limits/shapes/comm_sq.lean +++ b/src/category_theory/limits/shapes/comm_sq.lean @@ -1,11 +1,13 @@ /- Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Scott Morrison, Joël Riou -/ +import category_theory.comm_sq import category_theory.limits.preserves.shapes.pullbacks import category_theory.limits.shapes.zero_morphisms import category_theory.limits.constructions.binary_products +import category_theory.limits.opposites /-! # Pullback and pushout squares @@ -49,31 +51,10 @@ namespace category_theory variables {C : Type u₁} [category.{v₁} C] -/-- The proposition that a square -``` - W ---f---> X - | | - g h - | | - v v - Y ---i---> Z - -``` -is a commuting square. --/ -structure comm_sq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y ⟶ Z) : Prop := -(w : f ≫ h = g ≫ i) - -attribute [reassoc] comm_sq.w - namespace comm_sq variables {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} -lemma flip (p : comm_sq f g h i) : comm_sq g f i h := ⟨p.w.symm⟩ - -lemma of_arrow {f g : arrow C} (h : f ⟶ g) : comm_sq f.hom h.left h.right g.hom := ⟨h.w.symm⟩ - /-- The (not necessarily limiting) `pullback_cone h i` implicit in the statement that we have `comm_sq f g h i`. @@ -86,6 +67,32 @@ that we have `comm_sq f g h i`. -/ def cocone (s : comm_sq f g h i) : pushout_cocone f g := pushout_cocone.mk _ _ s.w +/-- The pushout cocone in the opposite category associated to the cone of +a commutative square identifies to the cocone of the flipped commutative square in +the opposite category -/ +def cone_op (p : comm_sq f g h i) : p.cone.op ≅ p.flip.op.cocone := +pushout_cocone.ext (iso.refl _) (by tidy) (by tidy) + +/-- The pullback cone in the opposite category associated to the cocone of +a commutative square identifies to the cone of the flipped commutative square in +the opposite category -/ +def cocone_op (p : comm_sq f g h i) : p.cocone.op ≅ p.flip.op.cone := +pullback_cone.ext (iso.refl _) (by tidy) (by tidy) + +/-- The pushout cocone obtained from the pullback cone associated to a +commutative square in the opposite category identifies to the cocone associated +to the flipped square. -/ +def cone_unop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + (p : comm_sq f g h i) : p.cone.unop ≅ p.flip.unop.cocone := +pushout_cocone.ext (iso.refl _) (by tidy) (by tidy) + +/-- The pullback cone obtained from the pushout cone associated to a +commutative square in the opposite category identifies to the cone associated +to the flipped square. -/ +def cocone_unop {W X Y Z : Cᵒᵖ} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + (p : comm_sq f g h i) : p.cocone.unop ≅ p.flip.unop.cone := +pullback_cone.ext (iso.refl _) (by tidy) (by tidy) + end comm_sq /-- The proposition that a square @@ -357,6 +364,17 @@ lemma of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} is_pullback h₁₁ v₁₁ v₁₂ h₂₁ := (of_bot s.flip p.symm t.flip).flip +lemma op (h : is_pullback fst snd f g) : is_pushout g.op f.op snd.op fst.op := +is_pushout.of_is_colimit (is_colimit.of_iso_colimit + (limits.pullback_cone.is_limit_equiv_is_colimit_op h.flip.cone h.flip.is_limit) + h.to_comm_sq.flip.cone_op) + +lemma unop {P X Y Z : Cᵒᵖ} {fst : P ⟶ X} {snd : P ⟶ Y} {f : X ⟶ Z} {g : Y ⟶ Z} + (h : is_pullback fst snd f g) : is_pushout g.unop f.unop snd.unop fst.unop := +is_pushout.of_is_colimit (is_colimit.of_iso_colimit + (limits.pullback_cone.is_limit_equiv_is_colimit_unop h.flip.cone h.flip.is_limit) + h.to_comm_sq.flip.cone_unop) + end is_pullback namespace is_pushout @@ -429,6 +447,17 @@ lemma of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} is_pushout h₁₂ v₁₂ v₁₃ h₂₂ := (of_bot s.flip p.symm t.flip).flip +lemma op (h : is_pushout f g inl inr) : is_pullback inr.op inl.op g.op f.op := +is_pullback.of_is_limit (is_limit.of_iso_limit + (limits.pushout_cocone.is_colimit_equiv_is_limit_op h.flip.cocone h.flip.is_colimit) + h.to_comm_sq.flip.cocone_op) + +lemma unop {Z X Y P : Cᵒᵖ} {f : Z ⟶ X} {g : Z ⟶ Y} {inl : X ⟶ P} {inr : Y ⟶ P} + (h : is_pushout f g inl inr) : is_pullback inr.unop inl.unop g.unop f.unop := +is_pullback.of_is_limit (is_limit.of_iso_limit + (limits.pushout_cocone.is_colimit_equiv_is_limit_unop h.flip.cocone h.flip.is_colimit) + h.to_comm_sq.flip.cocone_unop) + end is_pushout namespace functor @@ -436,9 +465,6 @@ namespace functor variables {D : Type u₂} [category.{v₂} D] variables (F : C ⥤ D) {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} -lemma map_comm_sq (s : comm_sq f g h i) : comm_sq (F.map f) (F.map g) (F.map h) (F.map i) := -⟨by simpa using congr_arg (λ k : W ⟶ Z, F.map k) s.w⟩ - lemma map_is_pullback [preserves_limit (cospan h i) F] (s : is_pullback f g h i) : is_pullback (F.map f) (F.map g) (F.map h) (F.map i) := -- This is made slightly awkward because `C` and `D` have different universes, @@ -465,7 +491,6 @@ end end functor -alias functor.map_comm_sq ← comm_sq.map alias functor.map_is_pullback ← is_pullback.map alias functor.map_is_pushout ← is_pushout.map From 39baa56ffc60e2a42e40738124d34ddfdddc96c5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 28 Jul 2022 01:01:34 +0000 Subject: [PATCH 119/173] feat(category_theory): construction of the localized category (#14422) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `W : morphism_property C` is a class of morphisms in a category `C`, this PR constructs the localized category `localization W` obtained by adding formal inverses to the morphisms in `W`. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- docs/references.bib | 9 + src/category_theory/arrow.lean | 3 + src/category_theory/category/Quiv.lean | 2 +- src/category_theory/eq_to_hom.lean | 7 + .../localization/construction.lean | 214 ++++++++++++++++++ src/category_theory/morphism_property.lean | 8 +- src/combinatorics/quiver/path.lean | 3 + 7 files changed, 244 insertions(+), 2 deletions(-) create mode 100644 src/category_theory/localization/construction.lean diff --git a/docs/references.bib b/docs/references.bib index 6dd8ea652d9ab..c374e3000a5be 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -762,6 +762,15 @@ @Article{ furedi-loeb1994 zbl = {0802.28002} } +@Book{ gabriel-zisman-1967, + author = {Gabriel, P. and Zisman, M.}, + title = {Calculus of fractions and homotopy theory}, + series = {Ergebnisse der Mathematik und ihrer Grenzgebiete, Band 35}, + publisher = {Springer-Verlag New York, Inc., New York}, + year = {1967}, + pages = {x+168}, +} + @InProceedings{ Gallier2011Notes, title = {Notes on Differential Geometry and Lie Groups}, author = {J. Gallier and J. Quaintance}, diff --git a/src/category_theory/arrow.lean b/src/category_theory/arrow.lean index 97c6eef12cdcc..d20e8c476c75c 100644 --- a/src/category_theory/arrow.lean +++ b/src/category_theory/arrow.lean @@ -51,6 +51,9 @@ def mk {X Y : T} (f : X ⟶ Y) : arrow T := right := Y, hom := f } +@[simp] lemma mk_eq (f : arrow T) : arrow.mk f.hom = f := +by { cases f, refl, } + theorem mk_injective (A B : T) : function.injective (arrow.mk : (A ⟶ B) → arrow T) := λ f g h, by { cases h, refl } diff --git a/src/category_theory/category/Quiv.lean b/src/category_theory/category/Quiv.lean index 423cb7d97b5a0..2bcf0b8f18308 100644 --- a/src/category_theory/category/Quiv.lean +++ b/src/category_theory/category/Quiv.lean @@ -71,7 +71,7 @@ namespace Quiv /-- Any prefunctor into a category lifts to a functor from the path category. -/ @[simps] -def lift {V : Type u} [quiver.{v+1} V] {C : Type u} [category.{v} C] +def lift {V : Type u} [quiver.{v+1} V] {C : Type*} [category C] (F : prefunctor V C) : paths V ⥤ C := { obj := λ X, F.obj X, map := λ X Y f, compose_path (F.map_path f), } diff --git a/src/category_theory/eq_to_hom.lean b/src/category_theory/eq_to_hom.lean index 65f70dda347ec..19e66aee7315d 100644 --- a/src/category_theory/eq_to_hom.lean +++ b/src/category_theory/eq_to_hom.lean @@ -140,6 +140,13 @@ lemma congr_hom {F G : C ⥤ D} (h : F = G) {X Y} (f : X ⟶ Y) : F.map f = eq_to_hom (congr_obj h X) ≫ G.map f ≫ eq_to_hom (congr_obj h Y).symm := by subst h; simp +lemma congr_inv_of_congr_hom (F G : C ⥤ D) {X Y : C} (e : X ≅ Y) + (hX : F.obj X = G.obj X) (hY : F.obj Y = G.obj Y) + (h₂ : F.map e.hom = eq_to_hom (by rw hX) ≫ G.map e.hom ≫ eq_to_hom (by rw hY)) : +F.map e.inv = eq_to_hom (by rw hY) ≫ G.map e.inv ≫ eq_to_hom (by rw hX) := +by simp only [← is_iso.iso.inv_hom e, functor.map_inv, h₂, is_iso.inv_comp, + inv_eq_to_hom, category.assoc] + section heq /- Composition of functors and maps w.r.t. heq -/ diff --git a/src/category_theory/localization/construction.lean b/src/category_theory/localization/construction.lean new file mode 100644 index 0000000000000..112084a4bf505 --- /dev/null +++ b/src/category_theory/localization/construction.lean @@ -0,0 +1,214 @@ +/- +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.morphism_property +import category_theory.category.Quiv + +/-! + +# Construction of the localized category + +This file constructs the localized category, obtained by formally inverting +a class of maps `W : morphism_property C` in a category `C`. + +We first construct a quiver `loc_quiver W` whose objects are the same as those +of `C` and whose maps are the maps in `C` and placeholders for the formal +inverses of the maps in `W`. + +The localized category `W.localization` is obtained by taking the quotient +of the path category of `loc_quiver W` by the congruence generated by four +types of relations. + +The obvious functor `Q W : C ⥤ W.localization` satisfies the universal property +of the localization. Indeed, if `G : C ⥤ D` sends morphisms in `W` to isomorphisms +in `D` (i.e. we have `hG : W.is_inverted_by G`), then there exists a unique functor +`G' : W.localization ⥤ D` such that `Q W ≫ G' = G`. This `G'` is `lift G hG`. +The expected property of `lift G hG` if expressed by the lemma `fac` and the +uniqueness is expressed by `uniq`. + +TODO: +1) show that for any category `E`, the composition of functors gives +an equivalence of categories between `W.localization ⥤ E` and the full +subcategory of `C ⥤ E` consisting of functors inverting `W`. (This only +requires an extension property for natural transformations of functors.) + +2) define a predicate `is_localization L W` for a functor `L : C ⥤ D` and +a class of morphisms `W` in `C` expressing that it is a localization with respect +to `W`, i.e. that it inverts `W` and that the obvious functor `W.localization ⥤ D` +induced by `L` is an equivalence of categories. (It is more straightforward +to define this predicate this way rather than by using a universal property which +may imply attempting at quantifying on all universes.) + +3) implement a constructor for `is_localization L W` which would take +as an input a *strict* universal property (`lift`/`fac`/`uniq`) similar to +what is obtained here for `W.localization`. (Practically speaking, this is +the easiest way to show that a functor is a localization.) + +4) when we have `is_localization L W`, then show that `D ⥤ E` identifies +to the full subcategory of `C ⥤ E` consisting of `W`-inverting functors. + +5) provide an API for the lifting of functors `C ⥤ E`, for which +`fac`/`uniq` assertions would be expressed as isomorphisms rather than +by equalities of functors. + +## References + +* [P. Gabriel, M. Zisman, *Calculus of fractions and homotopy theory*][gabriel-zisman-1967] + +-/ + +noncomputable theory + +open category_theory.category + +namespace category_theory + +variables {C D : Type*} [category C] [category D] (W : morphism_property C) + +namespace localization + +namespace construction + +/-- If `W : morphism_property C`, `loc_quiver W` is a quiver with the same objects +as `C`, and whose morphisms are those in `C` and placeholders for formal +inverses of the morphisms in `W`. -/ +@[nolint has_inhabited_instance] +structure loc_quiver (W : morphism_property C) := (obj : C) + +instance : quiver (loc_quiver W) := +{ hom := λ A B, (A.obj ⟶ B.obj) ⊕ { f : B.obj ⟶ A.obj // W f} } + +/-- The object in the path category of `loc_quiver W` attached to an object in +the category `C` -/ +def ι_paths (X : C) : paths (loc_quiver W) := ⟨X⟩ + +/-- The morphism in the path category associated to a morphism in the original category. -/ +@[simp] +def ψ₁ {X Y : C} (f : X ⟶ Y) : ι_paths W X ⟶ ι_paths W Y := paths.of.map (sum.inl f) + +/-- The morphism in the path category corresponding to a formal inverse. -/ +@[simp] +def ψ₂ {X Y : C} (w : X ⟶ Y) (hw : W w) : ι_paths W Y ⟶ ι_paths W X := +paths.of.map (sum.inr ⟨w, hw⟩) + +/-- The relations by which we take the quotient in order to get the localized category. -/ +inductive relations : hom_rel (paths (loc_quiver W)) +| id (X : C) : + relations (ψ₁ W (𝟙 X)) (𝟙 _) +| comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) : + relations (ψ₁ W (f ≫ g)) (ψ₁ W f ≫ ψ₁ W g) +| Winv₁ {X Y : C} (w : X ⟶ Y) (hw : W w) : + relations (ψ₁ W w ≫ ψ₂ W w hw) (𝟙 _) +| Winv₂ {X Y : C} (w : X ⟶ Y) (hw : W w) : + relations (ψ₂ W w hw ≫ ψ₁ W w) (𝟙 _) + +end construction + +end localization + +namespace morphism_property + +open localization.construction + +/-- The localized category obtained by formally inverting the morphisms +in `W : morphism_property C` -/ +@[derive category, nolint has_inhabited_instance] +def localization := category_theory.quotient (localization.construction.relations W) + +/-- The obvious functor `C ⥤ W.localization` -/ +def Q : C ⥤ W.localization := +{ obj := λ X, (quotient.functor _).obj (paths.of.obj ⟨X⟩), + map := λ X Y f, (quotient.functor _).map (ψ₁ W f), + map_id' := λ X, quotient.sound _ (relations.id X), + map_comp' := λ X Z Y f g, quotient.sound _ (relations.comp f g), } + +end morphism_property + +namespace localization + +namespace construction + +/-- The isomorphism in `W.localization` associated to a morphism `w` in W -/ +def Wiso {X Y : C} (w : X ⟶ Y) (hw : W w) : iso (W.Q.obj X) (W.Q.obj Y) := +{ hom := W.Q.map w, + inv := (quotient.functor _).map (paths.of.map (sum.inr ⟨w, hw⟩)), + hom_inv_id' := quotient.sound _ (relations.Winv₁ w hw), + inv_hom_id' := quotient.sound _ (relations.Winv₂ w hw), } + +end construction + +end localization + +namespace localization + +namespace construction + +variables {W} (G : C ⥤ D) (hG : W.is_inverted_by G) + +include G hG + +/-- The lifting of a functor to the path category of `loc_quiver W` -/ +@[simps] +def lift_to_path_category : paths (loc_quiver W) ⥤ D := +Quiv.lift +{ obj := λ X, G.obj X.obj, + map := λ X Y, begin + rintro (f|⟨g, hg⟩), + { exact G.map f, }, + { haveI := hG g hg, + exact inv (G.map g), }, + end, } + +/-- The lifting of a functor `C ⥤ D` inverting `W` as a functor `W.localization ⥤ D` -/ +@[simps] +def lift : W.localization ⥤ D := +quotient.lift (relations W) (lift_to_path_category G hG) +begin + rintro ⟨X⟩ ⟨Y⟩ f₁ f₂ r, + rcases r, + tidy, +end + +@[simp] +lemma fac : W.Q ⋙ lift G hG = G := +functor.ext (λ X, rfl) +begin + intros X Y f, + simp only [functor.comp_map, eq_to_hom_refl, comp_id, id_comp], + dsimp [lift, lift_to_path_category, morphism_property.Q], + rw compose_path_to_path, +end + +omit G hG + +lemma uniq (G₁ G₂ : W.localization ⥤ D) (h : W.Q ⋙ G₁ = W.Q ⋙ G₂) : + G₁ = G₂ := +begin + suffices h' : quotient.functor _ ⋙ G₁ = quotient.functor _ ⋙ G₂, + { refine functor.ext _ _, + { rintro ⟨⟨X⟩⟩, + apply functor.congr_obj h, }, + { rintros ⟨⟨X⟩⟩ ⟨⟨Y⟩⟩ ⟨f⟩, + apply functor.congr_hom h', }, }, + { refine paths.ext_functor _ _, + { ext X, + cases X, + apply functor.congr_obj h, }, + { rintro ⟨X⟩ ⟨Y⟩ (f|⟨w, hw⟩), + { simpa only using functor.congr_hom h f, }, + { have hw : W.Q.map w = (Wiso W w hw).hom := rfl, + have hw' := functor.congr_hom h w, + simp only [functor.comp_map, hw] at hw', + refine functor.congr_inv_of_congr_hom _ _ _ _ _ hw', + all_goals + { apply functor.congr_obj h, }, }, }, }, +end + +end construction + +end localization + +end category_theory diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index 8db835f72d3d3..102616acd5aeb 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -26,7 +26,7 @@ noncomputable theory namespace category_theory -variables (C : Type u) [category.{v} C] +variables (C : Type u) [category.{v} C] {D : Type*} [category D] /-- A `morphism_property C` is a class of morphisms between objects in `C`. -/ @[derive complete_lattice] @@ -121,6 +121,12 @@ begin hP.base_change_map hP' _ (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f') h₁], end +/-- If `P : morphism_property C` and `F : C ⥤ D`, then +`P.is_inverted_by F` means that all morphisms in `P` are mapped by `F` +to isomorphisms in `D`. -/ +def is_inverted_by (P : morphism_property C) (F : C ⥤ D) : Prop := +∀ ⦃X Y : C⦄ (f : X ⟶ Y) (hf : P f), is_iso (F.map f) + end morphism_property end category_theory diff --git a/src/combinatorics/quiver/path.lean b/src/combinatorics/quiver/path.lean index 62de1ffa1b554..430bbf2e3a708 100644 --- a/src/combinatorics/quiver/path.lean +++ b/src/combinatorics/quiver/path.lean @@ -84,4 +84,7 @@ def map_path {a : V} : | _ path.nil := rfl | _ (path.cons p e) := begin dsimp, rw [map_path_comp], end +@[simp] +lemma map_path_to_path {a b : V} (f : a ⟶ b) : F.map_path f.to_path = (F.map f).to_path := rfl + end prefunctor From 6b648812ca330e844edef50d96bbe209f9beac71 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 28 Jul 2022 02:26:42 +0000 Subject: [PATCH 120/173] chore(topology/algebra/module, analysis/normed_space/linear_isometry): dedup `submodule.subtypeL` and `continuous_linear_map.subtype_val` (#15700) To designate the continuous linear inclusion of a submodule into the ambient space, we currently have both `continuous_linear_map.subtype_val` (correct assumptions, name not consistent with `submodule.subtype`) and `submodule.subtypeL` (good name, but way too strong assumptions). This keeps the best of both worlds. --- src/analysis/calculus/implicit.lean | 8 +++--- src/analysis/normed_space/complemented.lean | 4 +-- .../normed_space/linear_isometry.lean | 14 ++-------- src/topology/algebra/module/basic.lean | 28 +++++++++++++------ 4 files changed, 26 insertions(+), 28 deletions(-) diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index dfeac8ee6776b..47aad7899c0a5 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -47,7 +47,7 @@ noncomputable theory open_locale topological_space open filter -open continuous_linear_map (fst snd subtype_val smul_right ker_prod) +open continuous_linear_map (fst snd smul_right ker_prod) open continuous_linear_equiv (of_bijective) /-! @@ -309,9 +309,9 @@ end lemma to_implicit_function_of_complemented (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : has_strict_fderiv_at (hf.implicit_function_of_complemented f f' hf' hker (f a)) - (subtype_val f'.ker) 0 := + f'.ker.subtypeL 0 := by convert (implicit_function_data_of_complemented f f' hf hf' - hker).implicit_function_has_strict_fderiv_at (subtype_val f'.ker) _ _; + hker).implicit_function_has_strict_fderiv_at f'.ker.subtypeL _ _; [skip, ext, ext]; simp [classical.some_spec hker] end complemented @@ -412,7 +412,7 @@ by apply eq_implicit_function_of_complemented lemma to_implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : has_strict_fderiv_at (hf.implicit_function f f' hf' (f a)) - (subtype_val f'.ker) 0 := + f'.ker.subtypeL 0 := by apply to_implicit_function_of_complemented end finite_dimensional diff --git a/src/analysis/normed_space/complemented.lean b/src/analysis/normed_space/complemented.lean index e666e595040c4..11c1018c88674 100644 --- a/src/analysis/normed_space/complemented.lean +++ b/src/analysis/normed_space/complemented.lean @@ -76,8 +76,6 @@ namespace subspace variables [complete_space E] (p q : subspace 𝕜 E) -open continuous_linear_map (subtype_val) - /-- If `q` is a closed complement of a closed subspace `p`, then `p × q` is continuously isomorphic to `E`. -/ def prod_equiv_of_closed_compl (h : is_compl p q) (hp : is_closed (p : set E)) @@ -85,7 +83,7 @@ def prod_equiv_of_closed_compl (h : is_compl p q) (hp : is_closed (p : set E)) begin haveI := hp.complete_space_coe, haveI := hq.complete_space_coe, refine (p.prod_equiv_of_is_compl q h).to_continuous_linear_equiv_of_continuous _, - exact ((subtype_val p).coprod (subtype_val q)).continuous + exact (p.subtypeL.coprod q.subtypeL).continuous end /-- Projection to a closed submodule along a closed complement. -/ diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 35497b695a19a..adeaf29e23ac7 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -251,18 +251,8 @@ def subtypeₗᵢ : p →ₗᵢ[R'] E := ⟨p.subtype, λ x, rfl⟩ @[simp] lemma subtypeₗᵢ_to_linear_map : p.subtypeₗᵢ.to_linear_map = p.subtype := rfl -/-- `submodule.subtype` as a `continuous_linear_map`. -/ -def subtypeL : p →L[R'] E := p.subtypeₗᵢ.to_continuous_linear_map - -@[simp] lemma coe_subtypeL : (p.subtypeL : p →ₗ[R'] E) = p.subtype := rfl - -@[simp] lemma coe_subtypeL' : ⇑p.subtypeL = p.subtype := rfl - -@[simp] lemma range_subtypeL : p.subtypeL.range = p := -range_subtype _ - -@[simp] lemma ker_subtypeL : p.subtypeL.ker = ⊥ := -ker_subtype _ +@[simp] lemma subtypeₗᵢ_to_continuous_linear_map : + p.subtypeₗᵢ.to_continuous_linear_map = p.subtypeL := rfl end submodule diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index e9f71739575ee..acf06b2b4f516 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -799,19 +799,29 @@ rfl ker (f.cod_restrict p h) = ker f := (f : M₁ →ₛₗ[σ₁₂] M₂).ker_cod_restrict p h -/-- Embedding of a submodule into the ambient space as a continuous linear map. -/ -def subtype_val (p : submodule R₁ M₁) : p →L[R₁] M₁ := +/-- `submodule.subtype` as a `continuous_linear_map`. -/ +def _root_.submodule.subtypeL (p : submodule R₁ M₁) : p →L[R₁] M₁ := { cont := continuous_subtype_val, to_linear_map := p.subtype } -@[simp, norm_cast] lemma coe_subtype_val (p : submodule R₁ M₁) : - (subtype_val p : p →ₗ[R₁] M₁) = p.subtype := +@[simp, norm_cast] lemma _root_.submodule.coe_subtypeL (p : submodule R₁ M₁) : + (p.subtypeL : p →ₗ[R₁] M₁) = p.subtype := rfl -@[simp, norm_cast] lemma subtype_val_apply (p : submodule R₁ M₁) (x : p) : - (subtype_val p : p → M₁) x = x := +@[simp] lemma _root_.submodule.coe_subtypeL' (p : submodule R₁ M₁) : + ⇑p.subtypeL = p.subtype := rfl +@[simp, norm_cast] lemma _root_.submodule.subtypeL_apply (p : submodule R₁ M₁) (x : p) : + p.subtypeL x = x := +rfl + +@[simp] lemma _root_.submodule.range_subtypeL (p : submodule R₁ M₁) : p.subtypeL.range = p := +submodule.range_subtype _ + +@[simp] lemma _root_.submodule.ker_subtypeL (p : submodule R₁ M₁) : p.subtypeL.ker = ⊥ := +submodule.ker_subtype _ + variables (R₁ M₁ M₂) /-- `prod.fst` as a `continuous_linear_map`. -/ @@ -1769,14 +1779,14 @@ end variables [module R M₂] [topological_add_group M] -open _root_.continuous_linear_map (id fst snd subtype_val mem_ker) +open _root_.continuous_linear_map (id fst snd mem_ker) /-- A pair of continuous linear maps such that `f₁ ∘ f₂ = id` generates a continuous linear equivalence `e` between `M` and `M₂ × f₁.ker` such that `(e x).2 = x` for `x ∈ f₁.ker`, `(e x).1 = f₁ x`, and `(e (f₂ y)).2 = 0`. The map is given by `e x = (f₁ x, x - f₂ (f₁ x))`. -/ def equiv_of_right_inverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : function.right_inverse f₂ f₁) : M ≃L[R] M₂ × f₁.ker := -equiv_of_inverse (f₁.prod (f₁.proj_ker_of_right_inverse f₂ h)) (f₂.coprod (subtype_val f₁.ker)) +equiv_of_inverse (f₁.prod (f₁.proj_ker_of_right_inverse f₂ h)) (f₂.coprod f₁.ker.subtypeL) (λ x, by simp) (λ ⟨x, y⟩, by simp [h x]) @@ -1934,7 +1944,7 @@ protected lemma closed_complemented.is_closed [topological_add_group M] [t1_spac is_closed (p : set M) := begin rcases h with ⟨f, hf⟩, - have : ker (id R M - (subtype_val p).comp f) = p := linear_map.ker_id_sub_eq_of_proj hf, + have : ker (id R M - p.subtypeL.comp f) = p := linear_map.ker_id_sub_eq_of_proj hf, exact this ▸ (is_closed_ker _) end From 79edb6d99424defb018db2d036d1668fc148c87b Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Thu, 28 Jul 2022 04:11:46 +0000 Subject: [PATCH 121/173] chore(category_theory/adjunction): fix a typo (#15720) --- src/category_theory/adjunction/comma.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/category_theory/adjunction/comma.lean b/src/category_theory/adjunction/comma.lean index 803d021f22548..19dc1a58801af 100644 --- a/src/category_theory/adjunction/comma.lean +++ b/src/category_theory/adjunction/comma.lean @@ -120,7 +120,7 @@ def adjunction_of_costructured_arrow_terminals : adjunction.adjunction_of_equiv_right _ _ /-- If each costructured arrow category on `G` has an terminal object, `G` is a left adjoint. -/ -def is_right_adjoint_of_costructured_arrow_terminals : is_left_adjoint G := +def is_left_adjoint_of_costructured_arrow_terminals : is_left_adjoint G := { right := right_adjoint_of_costructured_arrow_terminals G, adj := adjunction.adjunction_of_equiv_right _ _ } From 13166c6abcc7fbde2c22b3f9e85aa041f6fa1e3f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 28 Jul 2022 04:54:26 +0000 Subject: [PATCH 122/173] feat(order/bounded_order): `is_well_order` instances for `with_top` and `with_bot` (#15278) --- src/order/bounded_order.lean | 36 ++++++++++++++++++++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 6f31b9634e630..7a4e984d79b3f 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -1079,6 +1079,42 @@ lemma _root_.with_bot.well_founded_gt [preorder α] (h : @well_founded α (>)) : @well_founded (with_bot α) (>) := @with_top.well_founded_lt αᵒᵈ _ h +instance trichotomous.lt [preorder α] [is_trichotomous α (<)] : is_trichotomous (with_top α) (<) := +⟨begin + rintro (a | _) (b | _), + iterate 3 { simp }, + simpa [option.some_inj] using @trichotomous _ (<) _ a b +end⟩ + +instance is_well_order.lt [preorder α] [h : is_well_order α (<)] : is_well_order (with_top α) (<) := +{ wf := well_founded_lt h.wf } + +instance trichotomous.gt [preorder α] [is_trichotomous α (>)] : is_trichotomous (with_top α) (>) := +⟨begin + rintro (a | _) (b | _), + iterate 3 { simp }, + simpa [option.some_inj] using @trichotomous _ (>) _ a b +end⟩ + +instance is_well_order.gt [preorder α] [h : is_well_order α (>)] : is_well_order (with_top α) (>) := +{ wf := well_founded_gt h.wf } + +instance _root_.with_bot.trichotomous.lt [preorder α] [h : is_trichotomous α (<)] : + is_trichotomous (with_bot α) (<) := +@with_top.trichotomous.gt αᵒᵈ _ h + +instance _root_.with_bot.is_well_order.lt [preorder α] [h : is_well_order α (<)] : + is_well_order (with_bot α) (<) := +@with_top.is_well_order.gt αᵒᵈ _ h + +instance _root_.with_bot.trichotomous.gt [preorder α] [h : is_trichotomous α (>)] : + is_trichotomous (with_bot α) (>) := +@with_top.trichotomous.lt αᵒᵈ _ h + +instance _root_.with_bot.is_well_order.gt [preorder α] [h : is_well_order α (>)] : + is_well_order (with_bot α) (>) := +@with_top.is_well_order.lt αᵒᵈ _ h + instance [has_lt α] [densely_ordered α] [no_max_order α] : densely_ordered (with_top α) := order_dual.densely_ordered (with_bot αᵒᵈ) From 50b43a219677139eb1053e525252e38d4e83f7b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 28 Jul 2022 04:54:27 +0000 Subject: [PATCH 123/173] chore(measure_theory/tactic): remove the measurability attribute from two lemmas (#15295) The measurability tactic applied `set.finite.measurable_set` to every set (when the space had measurable singletons) and got stuck. --- src/measure_theory/tactic.lean | 2 -- test/measurability.lean | 6 ++++++ 2 files changed, 6 insertions(+), 2 deletions(-) diff --git a/src/measure_theory/tactic.lean b/src/measure_theory/tactic.lean index c737dcc5ae7a7..195407aa7b0ee 100644 --- a/src/measure_theory/tactic.lean +++ b/src/measure_theory/tactic.lean @@ -59,9 +59,7 @@ attribute [measurability] measurable_set.const measurable_set.insert measurable_set_eq - set.finite.measurable_set finset.measurable_set - set.countable.measurable_set measurable_space.measurable_set_top namespace tactic diff --git a/test/measurability.lean b/test/measurability.lean index 9b874480962ca..c7b66ea22e345 100644 --- a/test/measurability.lean +++ b/test/measurability.lean @@ -44,6 +44,12 @@ example (hf : measurable f) (hs₁ : measurable_set s₁) (ht₂ : measurable_se measurable_set ((f ⁻¹' t₂) ∩ s₁) := by measurability +/-- `ℝ` is a good test case because it verifies many assumptions, hence many lemmas apply and we +are more likely to detect a bad lemma. In a previous version of the tactic, `measurability` got +stuck trying to apply `set.finite.measurable_set` here. -/ +example {a b : ℝ} : measurable_set (set.Icc a b) := +by measurability + -- Tests on functions example [has_mul β] [has_measurable_mul₂ β] (hf : measurable f) (c : β) : From a42a71f0d70dc7b48cff696bfa889110067f7135 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 28 Jul 2022 04:54:28 +0000 Subject: [PATCH 124/173] feat(group_theory/finite_abelian): a finitely generated torsion abelian group is finite (#15402) --- src/group_theory/finite_abelian.lean | 47 ++++++++++++++++++++++++---- 1 file changed, 41 insertions(+), 6 deletions(-) diff --git a/src/group_theory/finite_abelian.lean b/src/group_theory/finite_abelian.lean index 04c269575ab3d..f2408e8ba4162 100644 --- a/src/group_theory/finite_abelian.lean +++ b/src/group_theory/finite_abelian.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Pierre-Alexandre Bazin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Pierre-Alexandre Bazin -/ -import data.zmod.quotient import algebra.module.pid +import data.zmod.quotient /-! # Structure of finite(ly generated) abelian groups @@ -17,15 +17,39 @@ import algebra.module.pid -/ - open_locale direct_sum +universe u + +namespace module + +variables (M : Type u) + +lemma finite_of_fg_torsion [add_comm_group M] [module ℤ M] [module.finite ℤ M] + (hM : module.is_torsion ℤ M) : _root_.finite M := +begin + rcases module.equiv_direct_sum_of_is_torsion hM with ⟨ι, _, p, h, e, ⟨l⟩⟩, + haveI : ∀ i : ι, fact $ 0 < (p i ^ e i).nat_abs := + λ i, fact.mk $ int.nat_abs_pos_of_ne_zero $ pow_ne_zero (e i) (h i).ne_zero, + haveI : ∀ i : ι, _root_.finite $ ℤ ⧸ submodule.span ℤ {p i ^ e i} := + λ i, finite.of_equiv _ (p i ^ e i).quotient_span_equiv_zmod.symm.to_equiv, + haveI : _root_.finite ⨁ i, ℤ ⧸ (submodule.span ℤ {p i ^ e i} : submodule ℤ ℤ) := + finite.of_equiv _ dfinsupp.equiv_fun_on_fintype.symm, + exact finite.of_equiv _ l.symm.to_equiv +end + +end module + +variables (G : Type u) + namespace add_comm_group +variable [add_comm_group G] + /-- **Structure theorem of finitely generated abelian groups** : 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`.-/ -theorem equiv_free_prod_direct_sum_zmod (G : Type*) [add_comm_group G] [hG : add_group.fg G] : +prime powers `p i ^ e i`. -/ +theorem equiv_free_prod_direct_sum_zmod [hG : add_group.fg G] : ∃ (n : ℕ) (ι : Type) [fintype ι] (p : ι → ℕ) [∀ i, nat.prime $ p i] (e : ι → ℕ), nonempty $ G ≃+ (fin n →₀ ℤ) × ⨁ (i : ι), zmod (p i ^ e i) := begin @@ -39,8 +63,8 @@ begin end /-- **Structure theorem of finite abelian groups** : Any finite abelian group is a direct sum of -some `zmod (p i ^ e i)` for some prime powers `p i ^ e i`.-/ -theorem equiv_direct_sum_zmod_of_fintype (G : Type*) [add_comm_group G] [fintype G] : +some `zmod (p i ^ e i)` for some prime powers `p i ^ e i`. -/ +theorem equiv_direct_sum_zmod_of_fintype [fintype G] : ∃ (ι : Type) [fintype ι] (p : ι → ℕ) [∀ i, nat.prime $ p i] (e : ι → ℕ), nonempty $ G ≃+ ⨁ (i : ι), zmod (p i ^ e i) := begin @@ -52,4 +76,15 @@ begin λ a, ⟨finsupp.single 0 a, finsupp.single_eq_same⟩).false.elim } end +lemma finite_of_fg_torsion [hG' : add_group.fg G] (hG : add_monoid.is_torsion G) : finite G := +@module.finite_of_fg_torsion _ _ _ (module.finite.iff_add_group_fg.mpr hG') $ + add_monoid.is_torsion_iff_is_torsion_int.mp hG + end add_comm_group + +namespace comm_group + +lemma finite_of_fg_torsion [comm_group G] [group.fg G] (hG : monoid.is_torsion G) : finite G := +@finite.of_equiv _ _ (add_comm_group.finite_of_fg_torsion (additive G) hG) multiplicative.of_add + +end comm_group From d6001a4cc3d9b09406bf37be91aa58bf11664340 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 28 Jul 2022 04:54:29 +0000 Subject: [PATCH 125/173] chore(field_theory/splitting_field): remove unneeded parameter (#15716) The definition of `splitting_field_aux` relied on a parameter `hfn : n = f.nat_degree` that was never actually used in the construction or its instances, only in the `splits` theorem. So we might as well delete it. This PR prepares for a redefinition of all instances on `splitting_field_aux` to fix instance diamonds, see the `splitting_field-diamond` branch for progress. Co-authored-by: Anne Baanen --- src/field_theory/splitting_field.lean | 113 ++++++++++++++------------ 1 file changed, 59 insertions(+), 54 deletions(-) diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index e6f2b4f51bc6d..a38f86aded4b2 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -479,24 +479,29 @@ theorem nat_degree_remove_factor' {f : K[X]} {n : ℕ} (hfn : f.nat_degree = n+1 f.remove_factor.nat_degree = n := by rw [nat_degree_remove_factor, hfn, n.add_sub_cancel] -/-- Auxiliary construction to a splitting field of a polynomial. Uses induction on the degree. -/ -def splitting_field_aux (n : ℕ) : Π {K : Type u} [field K], by exactI Π (f : K[X]), - f.nat_degree = n → Type u := -nat.rec_on n (λ K _ _ _, K) $ λ n ih K _ f hf, by exactI -ih f.remove_factor (nat_degree_remove_factor' hf) +/-- 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]) (hfn : f.nat_degree = n + 1) : - splitting_field_aux (n+1) f hfn = - splitting_field_aux n f.remove_factor (nat_degree_remove_factor' hfn) := rfl +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]} (hfn : f.nat_degree = n), field (splitting_field_aux n f hfn) := -nat.rec_on n (λ K _ _ _, ‹field K›) $ λ n ih K _ f hf, ih _ + Π {f : K[X]}, field (splitting_field_aux n f) := +nat.rec_on n (λ K _ _, ‹field K›) $ λ n ih K _ f, ih -instance inhabited {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n) : - inhabited (splitting_field_aux n f hfn) := ⟨37⟩ +instance inhabited {n : ℕ} {f : K[X]} : + inhabited (splitting_field_aux n f) := ⟨37⟩ /- Note that the recursive nature of this definition and `splitting_field_aux.field` creates @@ -519,66 +524,66 @@ example (x : ℕ) {α} (a₀ aₙ : α) : (cases_twice a₀ aₙ x).1 = (cases_t 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]} (hfn : f.nat_degree = n), - algebra R (splitting_field_aux n f hfn) := -nat.rec_on n (λ R K _ _ _ _ _, by exactI ‹algebra R K›) $ - λ n ih R K _ _ _ f hfn, by exactI ih R (nat_degree_remove_factor' hfn) + 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 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]} (hfn : f.nat_degree = n), - is_scalar_tower R₁ R₂ (splitting_field_aux n f hfn) := -nat.rec_on n (λ R₁ R₂ K _ _ _ _ _ _ _ _ _, by exactI ‹is_scalar_tower R₁ R₂ K›) $ - λ n ih R₁ R₂ K _ _ _ _ _ _ _ f hfn, by exactI ih R₁ R₂ (nat_degree_remove_factor' hfn) + 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₂ -instance algebra''' {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n + 1) : +instance algebra''' {n : ℕ} {f : K[X]} : algebra (adjoin_root f.factor) - (splitting_field_aux n f.remove_factor (nat_degree_remove_factor' hfn)) := -splitting_field_aux.algebra n _ _ + (splitting_field_aux n f.remove_factor) := +splitting_field_aux.algebra n _ -instance algebra' {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n + 1) : - algebra (adjoin_root f.factor) (splitting_field_aux n.succ f hfn) := -splitting_field_aux.algebra''' _ +instance algebra' {n : ℕ} {f : K[X]} : + algebra (adjoin_root f.factor) (splitting_field_aux n.succ f) := +splitting_field_aux.algebra''' -instance algebra'' {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n + 1) : - algebra K (splitting_field_aux n f.remove_factor (nat_degree_remove_factor' hfn)) := -splitting_field_aux.algebra n K _ +instance algebra'' {n : ℕ} {f : K[X]} : + algebra K (splitting_field_aux n f.remove_factor) := +splitting_field_aux.algebra n K -instance scalar_tower' {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n + 1) : +instance scalar_tower' {n : ℕ} {f : K[X]} : is_scalar_tower K (adjoin_root f.factor) - (splitting_field_aux n f.remove_factor (nat_degree_remove_factor' hfn)) := + (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) (nat_degree_remove_factor' hfn), + splitting_field_aux.is_scalar_tower n K (adjoin_root f.factor), end -instance scalar_tower {n : ℕ} {f : K[X]} (hfn : f.nat_degree = n + 1) : - is_scalar_tower K (adjoin_root f.factor) (splitting_field_aux _ f hfn) := -splitting_field_aux.scalar_tower' _ +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' -theorem algebra_map_succ (n : ℕ) (f : K[X]) (hfn : f.nat_degree = n + 1) : - by exact algebra_map K (splitting_field_aux _ _ hfn) = +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 (nat_degree_remove_factor' hfn))).comp + (splitting_field_aux n f.remove_factor)).comp (adjoin_root.of f.factor) := is_scalar_tower.algebra_map_eq _ _ _ protected theorem splits (n : ℕ) : ∀ {K : Type u} [field K], by exactI ∀ (f : K[X]) (hfn : f.nat_degree = n), - splits (algebra_map K $ splitting_field_aux n f hfn) f := + splits (algebra_map K $ splitting_field_aux n f) f := nat.rec_on n (λ K _ _ hf, by exactI splits_of_degree_le_one _ (le_trans degree_le_nat_degree $ hf.symm ▸ with_bot.coe_le_coe.2 zero_le_one)) $ λ n ih K _ f hf, by { resetI, rw [← splits_id_iff_splits, algebra_map_succ, ← map_map, splits_id_iff_splits, ← X_sub_C_mul_remove_factor f (λ h, by { rw h at hf, cases hf })], -exact splits_mul _ (splits_X_sub_C _) (ih _ _) } +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 hfn →+* L, + ∀ (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 }, @@ -596,37 +601,37 @@ let ⟨k, hk⟩ := ih f.remove_factor (nat_degree_remove_factor' hf) (adjoin_roo theorem adjoin_roots (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 hfn).roots.to_finset : - set (splitting_field_aux n f hfn)) = ⊤ := + algebra.adjoin K (↑(f.map $ algebra_map K $ splitting_field_aux n f).roots.to_finset : + 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 hfn)) f ≠ 0 := map_ne_zero hfn0, +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, 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 (nat_degree_remove_factor' hfn)), + (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 (nat_degree_remove_factor' hfn)), - ih, subalgebra.restrict_scalars_top] } + (splitting_field_aux n f.remove_factor), + ih _ (nat_degree_remove_factor' hfn), subalgebra.restrict_scalars_top] } end splitting_field_aux /-- A splitting field of a polynomial. -/ def splitting_field (f : K[X]) := -splitting_field_aux _ f rfl +splitting_field_aux f.nat_degree f namespace splitting_field variables (f : K[X]) instance : field (splitting_field f) := -splitting_field_aux.field _ _ +splitting_field_aux.field _ instance inhabited : inhabited (splitting_field f) := ⟨37⟩ @@ -656,25 +661,25 @@ since `K = ℕ` and `K = ℤ` are not possible due to the `field K` assumption. 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 _ _ _ +splitting_field_aux.algebra _ _ instance : algebra K (splitting_field f) := -splitting_field_aux.algebra _ _ _ +splitting_field_aux.algebra _ _ protected theorem splits : splits (algebra_map K (splitting_field f)) f := -splitting_field_aux.splits _ _ _ +splitting_field_aux.splits _ _ rfl 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 _ _ _ _ hb), +{ 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) } 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 _ _ _ +splitting_field_aux.adjoin_roots _ _ rfl theorem adjoin_root_set : algebra.adjoin K (f.root_set f.splitting_field) = ⊤ := adjoin_roots f From f90bbe65ff6adfb6f71e74eab2a4a6f4435fc57b Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 28 Jul 2022 13:50:36 +0000 Subject: [PATCH 126/173] feat(test/polyrith): better formatting for generated tests (#15728) The `polyrith` test suite is unusually structured to avoid making dozens of calls to Sage. A helper tactic calls `polyrith` and saves the output from Sage along with intermediate information, printing a full test as a "Try this:". This PR makes that helper tactic format its output better, spread over multiple lines as requested at https://github.com/leanprover-community/mathlib/pull/15292#discussion_r925838661 by @eric-wieser . Co-authored-by: Rob Lewis --- test/polyrith.lean | 386 +++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 354 insertions(+), 32 deletions(-) diff --git a/test/polyrith.lean b/test/polyrith.lean index 0662171f2d316..c8cf2c4f579ec 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Dhruv Bhatia. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Author(s): Dhruv Bhatia +Author(s): Dhruv Bhatia, Robert Y. Lewis -/ import tactic.polyrith @@ -12,7 +12,8 @@ import data.real.basic Each call to `polyrith` makes a call to the SageCell web API at . To avoid making many API calls from CI, we only test this communication in a few tests. -A full test suite is provided below the `#exit` command. + +A full test suite is provided at the bottom of the file. -/ @@ -37,8 +38,8 @@ meta def tactic.test_polyrith (only_on : bool) (hyps : list pexpr) guard (out = expected_out) <|> fail!"expected final output: {expected_out}\nbut produced: {out}" -meta def format_string_list (input : list string) : string := -to_string $ input.map (λ s, "\"" ++ s ++ "\"") +meta def format_string_list (input : list string) : format := +"[" ++ (format.join $ (input.map (λ s, ("\"" : format) ++ format.of_string s ++ "\"")).intersperse ("," ++ format.line)) ++ "]" setup_tactic_parser @@ -87,7 +88,8 @@ meta def tactic.interactive.create_polyrith_test (restr : parse (tk "only")?) let argstring := format_string_list args, let onl := if restr.is_some then "only " else "", let hyps := if hyps = [] then "" else to_string hyps, - trace!"test_polyrith {onl}{hyps} \"{sage_out}\" {argstring} \"{out}\"" + let trf := format.nest 2 $ format!"test_polyrith {onl}{hyps} \n\"{sage_out}\"\n{argstring}\n\"{out}\"", + trace!"Try this: {trf}" end tactic @@ -115,87 +117,213 @@ end example (x y : ℤ) (h1 : 3*x + 2*y = 10): 3*x + 2*y = 10 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "2", "[(((3 * var0) + (2 * var1)) - 10)]", "(((3 * var0) + (2 * var1)) - 10)"] "linear_combination h1" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "int", + "2", + "[(((3 * var0) + (2 * var1)) - 10)]", + "(((3 * var0) + (2 * var1)) - 10)"] + "linear_combination h1" example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : x*y = -2*y + 1 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" ["ff", "rat", "2", "[(((var0 * var1) + (2 * var0)) - 1), (var0 - var1)]", "((var0 * var1) - (((-1 * 2) * var1) + 1))"] "linear_combination h1 - 2 * h2" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + ["ff", + "rat", + "2", + "[(((var0 * var1) + (2 * var0)) - 1), (var0 - var1)]", + "((var0 * var1) - (((-1 * 2) * var1) + 1))"] + "linear_combination h1 - 2 * h2" example (x y : ℝ) (h1 : x + 2 = -3) (h2 : y = 10) : -y + 2*x + 4 = -16 := -by test_polyrith "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\"],\"success\":true}" ["ff", "real", "2", "[((var1 + 2) - (-1 * 3)), (var0 - 10)]", "((((-1 * var0) + (2 * var1)) + 4) - (-1 * 16))"] "linear_combination 2 * h1 - h2" +by test_polyrith + "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\"],\"success\":true}" + ["ff", + "real", + "2", + "[((var1 + 2) - (-1 * 3)), (var0 - 10)]", + "((((-1 * var0) + (2 * var1)) + 4) - (-1 * 16))"] + "linear_combination 2 * h1 - h2" example (x y z : ℝ) (ha : x + 2*y - z = 4) (hb : 2*x + y + z = -2) (hc : x + 2*y + z = 2) : -3*x - 3*y - 4*z = 2 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\",\"(poly.const -2/1)\"],\"success\":true}" ["ff", "real", "3", "[(((var0 + (2 * var1)) - var2) - 4), ((((2 * var0) + var1) + var2) - (-1 * 2)), (((var0 + (2 * var1)) + var2) - 2)]", "(((((-1 * 3) * var0) - (3 * var1)) - (4 * var2)) - 2)"] "linear_combination ha - hb - 2 * hc" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + ["ff", + "real", + "3", + "[(((var0 + (2 * var1)) - var2) - 4), ((((2 * var0) + var1) + var2) - (-1 * 2)), (((var0 + (2 * var1)) + var2) - 2)]", + "(((((-1 * 3) * var0) - (3 * var1)) - (4 * var2)) - 2)"] + "linear_combination ha - hb - 2 * hc" example (w x y z : ℝ) (h1 : x + 2.1*y + 2*z = 2) (h2 : x + 8*z + 5*w = -6.5) (h3 : x + y + 5*z + 5*w = 3) : x + 2.2*y + 2*z - 5*w = -8.5 := -by test_polyrith "{\"data\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" ["ff", "real", "4", "[(((var0 + (21/10 * var1)) + (2 * var2)) - 2), (((var0 + (8 * var2)) + (5 * var3)) - (-1 * 13/2)), ((((var0 + var1) + (5 * var2)) + (5 * var3)) - 3)]", "((((var0 + (11/5 * var1)) + (2 * var2)) - (5 * var3)) - (-1 * 17/2))"] "linear_combination 2 * h1 + h2 - 2 * h3" +by test_polyrith + "{\"data\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + ["ff", + "real", + "4", + "[(((var0 + (21/10 * var1)) + (2 * var2)) - 2), (((var0 + (8 * var2)) + (5 * var3)) - (-1 * 13/2)), ((((var0 + var1) + (5 * var2)) + (5 * var3)) - 3)]", + "((((var0 + (11/5 * var1)) + (2 * var2)) - (5 * var3)) - (-1 * 17/2))"] + "linear_combination 2 * h1 + h2 - 2 * h3" example (a b c d : ℚ) (h1 : a = 4) (h2 : 3 = b) (h3 : c*3 = d) (h4 : -d = a) : 2*a - 3 + 9*c + 3*d = 8 - b + 3*d - 3*a := -by test_polyrith "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\",\"(poly.const 3/1)\",\"(poly.const -3/1)\"],\"success\":true}" ["ff", "rat", "4", "[(var0 - 4), (3 - var3), ((var1 * 3) - var2), ((-1 * var2) - var0)]", "(((((2 * var0) - 3) + (9 * var1)) + (3 * var2)) - (((8 - var3) + (3 * var2)) - (3 * var0)))"] "linear_combination 2 * h1 - h2 + 3 * h3 - 3 * h4" +by test_polyrith + "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\",\"(poly.const 3/1)\",\"(poly.const -3/1)\"],\"success\":true}" + ["ff", + "rat", + "4", + "[(var0 - 4), (3 - var3), ((var1 * 3) - var2), ((-1 * var2) - var0)]", + "(((((2 * var0) - 3) + (9 * var1)) + (3 * var2)) - (((8 - var3) + (3 * var2)) - (3 * var0)))"] + "linear_combination 2 * h1 - h2 + 3 * h3 - 3 * h4" /-! ### Case with ambiguous identifiers-/ example («def evil» y : ℤ) (h1 : 3*«def evil» + 2*y = 10): 3*«def evil» + 2*y = 10 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "2", "[(((3 * var0) + (2 * var1)) - 10)]", "(((3 * var0) + (2 * var1)) - 10)"] "linear_combination h1" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "int", + "2", + "[(((3 * var0) + (2 * var1)) - 10)]", + "(((3 * var0) + (2 * var1)) - 10)"] + "linear_combination h1" example («¥» y : ℤ) (h1 : 3*«¥» + 2*y = 10): «¥» * (3*«¥» + 2*y) = 10 * «¥» := -by test_polyrith "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" ["ff", "int", "2", "[(((3 * var0) + (2 * var1)) - 10)]", "((var0 * ((3 * var0) + (2 * var1))) - (10 * var0))"] "linear_combination 1 * «¥» * h1" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" + ["ff", + "int", + "2", + "[(((3 * var0) + (2 * var1)) - 10)]", + "((var0 * ((3 * var0) + (2 * var1))) - (10 * var0))"] + "linear_combination 1 * «¥» * h1" /-! ### Cases with arbitrary coefficients -/ example (a b : ℤ) (h : a = b) : a * a = a * b := -by test_polyrith "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" ["ff", "int", "2", "[(var0 - var1)]", "((var0 * var0) - (var0 * var1))"] "linear_combination 1 * a * h" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" + ["ff", + "int", + "2", + "[(var0 - var1)]", + "((var0 * var0) - (var0 * var1))"] + "linear_combination 1 * a * h" example (a b c : ℤ) (h : a = b) : a * c = b * c := -by test_polyrith "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" ["ff", "int", "3", "[(var0 - var2)]", "((var0 * var1) - (var2 * var1))"] "linear_combination 1 * c * h" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + ["ff", + "int", + "3", + "[(var0 - var2)]", + "((var0 * var1) - (var2 * var1))"] + "linear_combination 1 * c * h" example (a b c : ℤ) (h1 : a = b) (h2 : b = 1) : c * a + b = c * b + 1 := -by test_polyrith "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\",\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "3", "[(var1 - var2), (var2 - 1)]", "(((var0 * var1) + var2) - ((var0 * var2) + 1))"] "linear_combination 1 * c * h1 + h2" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\",\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "int", + "3", + "[(var1 - var2), (var2 - 1)]", + "(((var0 * var1) + var2) - ((var0 * var2) + 1))"] + "linear_combination 1 * c * h1 + h2" example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : x*x*y + y*x*y + 6*x = 3*x*y + 14 := -by test_polyrith "{\"data\":[\"(poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1))\",\"(poly.const 2/1)\"],\"success\":true}" ["ff", "rat", "2", "[((var0 + var1) - 3), ((3 * var0) - 7)]", "(((((var0 * var0) * var1) + ((var1 * var0) * var1)) + (6 * var0)) - (((3 * var0) * var1) + 14))"] "linear_combination 1 * x * y * h1 + 2 * h2" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1))\",\"(poly.const 2/1)\"],\"success\":true}" + ["ff", + "rat", + "2", + "[((var0 + var1) - 3), ((3 * var0) - 7)]", + "(((((var0 * var0) * var1) + ((var1 * var0) * var1)) + (6 * var0)) - (((3 * var0) * var1) + 14))"] + "linear_combination 1 * x * y * h1 + 2 * h2" example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := -by test_polyrith "{\"data\":[\"(poly.add (poly.mul (poly.const 1/1) (poly.var 0)) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"success\":true}" ["ff", "rat", "4", "[(var1 - var3)]", "(((var0 * var1) + ((2 * var2) * var1)) - ((var0 * var3) + ((2 * var2) * var3)))"] "linear_combination (1 * x + 2 * y) * hzw" +by test_polyrith + "{\"data\":[\"(poly.add (poly.mul (poly.const 1/1) (poly.var 0)) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"success\":true}" + ["ff", + "rat", + "4", + "[(var1 - var3)]", + "(((var0 * var1) + ((2 * var2) * var1)) - ((var0 * var3) + ((2 * var2) * var3)))"] + "linear_combination (1 * x + 2 * y) * hzw" /-! ### Cases with non-hypothesis inputs/input restrictions -/ example (a b : ℝ) (ha : 2*a = 4) (hab : 2*b = a - b) (hignore : 3 = a + b) : b = 2 / 3 := -by test_polyrith only [ha, hab] "{\"data\":[\"(poly.const 1/6)\",\"(poly.const 1/3)\"],\"success\":true}" ["ff", "real", "2", "[((2 * var1) - 4), ((2 * var0) - (var1 - var0))]", "(var0 - 2/3)"] "linear_combination ha / 6 + hab / 3" +by test_polyrith only [ha, hab] + "{\"data\":[\"(poly.const 1/6)\",\"(poly.const 1/3)\"],\"success\":true}" + ["ff", + "real", + "2", + "[((2 * var1) - 4), ((2 * var0) - (var1 - var0))]", + "(var0 - 2/3)"] + "linear_combination ha / 6 + hab / 3" constant term : ∀ a b : ℚ, a + b = 0 example (a b c d : ℚ) (h : a + b = 0) (h2: b + c = 0): a + b + c + d = 0 := -by test_polyrith only [term c d, h] "{\"data\":[\"(poly.const 1/1)\",\"(poly.const 1/1)\"],\"success\":true}" ["ff", "rat", "4", "[((var2 + var3) - 0), ((var0 + var1) - 0)]", "((((var0 + var1) + var2) + var3) - 0)"] "linear_combination term c d + h" +by test_polyrith only [term c d, h] + "{\"data\":[\"(poly.const 1/1)\",\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "rat", + "4", + "[((var2 + var3) - 0), ((var0 + var1) - 0)]", + "((((var0 + var1) + var2) + var3) - 0)"] + "linear_combination term c d + h" constants (qc : ℚ) (hqc : qc = 2*qc) example (a b : ℚ) (h : ∀ p q : ℚ, p = q) : 3*a + qc = 3*b + 2*qc := -by test_polyrith [h a b, hqc] "{\"data\":[\"(poly.const 3/1)\",\"(poly.const 1/1)\"],\"success\":true}" ["ff", "rat", "3", "[(var0 - var2), (var1 - (2 * var1))]", "(((3 * var0) + var1) - ((3 * var2) + (2 * var1)))"] "linear_combination 3 * h a b + hqc" +by test_polyrith [h a b, hqc] + "{\"data\":[\"(poly.const 3/1)\",\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "rat", + "3", + "[(var0 - var2), (var1 - (2 * var1))]", + "(((3 * var0) + var1) - ((3 * var2) + (2 * var1)))"] + "linear_combination 3 * h a b + hqc" constant bad (q : ℚ) : q = 0 example (a b : ℚ) : a + b^3 = 0 := -by test_polyrith [bad a, bad (b^2)] "{\"data\":[\"(poly.const 1/1)\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" ["ff", "rat", "2", "[(var0 - 0), ((var1 ^ 2) - 0)]", "((var0 + (var1 ^ 3)) - 0)"] "linear_combination bad a + 1 * b * bad (b ^ 2)" +by test_polyrith [bad a, bad (b^2)] + "{\"data\":[\"(poly.const 1/1)\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + ["ff", + "rat", + "2", + "[(var0 - 0), ((var1 ^ 2) - 0)]", + "((var0 + (var1 ^ 3)) - 0)"] + "linear_combination bad a + 1 * b * bad (b ^ 2)" /-! ### Case over arbitrary field/ring -/ example {α} [h : comm_ring α] {a b c d e f : α} (h1 : a*d = b*c) (h2 : c*f = e*d) : c * (a*f - b*e) = 0 := -by test_polyrith "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 4))\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" ["ff", "α", "6", "[((var1 * var5) - (var3 * var0)), ((var0 * var2) - (var4 * var5))]", "((var0 * ((var1 * var2) - (var3 * var4))) - 0)"] "linear_combination 1 * e * h1 + 1 * a * h2" +by test_polyrith + "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 4))\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + ["ff", + "α", + "6", + "[((var1 * var5) - (var3 * var0)), ((var0 * var2) - (var4 * var5))]", + "((var0 * ((var1 * var2) - (var3 * var4))) - 0)"] + "linear_combination 1 * e * h1 + 1 * a * h2" example {K : Type*} [field K] [invertible 2] [invertible 3] {ω p q r s t x: K} (hp_nonzero : p ≠ 0) (hr : r ^ 2 = q ^ 2 + p ^ 3) (hs3 : s ^ 3 = q + r) @@ -205,11 +333,32 @@ example {K : Type*} [field K] [invertible 2] [invertible 3] begin have hs_nonzero : s ≠ 0, { contrapose! hp_nonzero with hs_nonzero, - test_polyrith "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.mul (poly.const 1/1) (poly.var 4))\"],\"success\":true}" ["ff", "K", "6", "[((var1 ^ 2) - ((var2 ^ 2) + (var0 ^ 3))), ((var3 ^ 3) - (var2 + var1)), ((var4 * var3) - var0), (((1 + var5) + (var5 ^ 2)) - 0), (var3 - 0)]", "(var0 - 0)"] "linear_combination -ht + 1 * t * hs_nonzero" }, + test_polyrith + "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.mul (poly.const 1/1) (poly.var 4))\"],\"success\":true}" + ["ff", + "K", + "6", + "[((var1 ^ 2) - ((var2 ^ 2) + (var0 ^ 3))), ((var3 ^ 3) - (var2 + var1)), ((var4 * var3) - var0), (((1 + var5) + (var5 ^ 2)) - 0), (var3 - 0)]", + "(var0 - 0)"] + "linear_combination -ht + 1 * t * hs_nonzero"}, have H' : 2 * q = s ^ 3 - t ^ 3, { rw ← mul_left_inj' (pow_ne_zero 3 hs_nonzero), - test_polyrith "{\"data\":[\"(poly.const -1/1)\",\"(poly.add (poly.add (poly.mul (poly.const -1/1) (poly.pow (poly.var 1) 3)) (poly.mul (poly.const 1/1) (poly.var 0))) (poly.mul (poly.const -1/1) (poly.var 3)))\",\"(poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 1) 2)) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 2)))\",\"(poly.const 0/1)\"],\"success\":true}" ["ff", "K", "6", "[((var3 ^ 2) - ((var0 ^ 2) + (var4 ^ 3))), ((var1 ^ 3) - (var0 + var3)), ((var2 * var1) - var4), (((1 + var5) + (var5 ^ 2)) - 0)]", "(((2 * var0) * (var1 ^ 3)) - (((var1 ^ 3) - (var2 ^ 3)) * (var1 ^ 3)))"] "linear_combination -hr + ((-1) * s ^ 3 + 1 * q + (-1) * r) * hs3 + (1 * s ^ 2 * t ^ 2 + 1 * s * t * p + 1 * p ^ 2) * ht" }, -test_polyrith "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.pow (poly.var 5) 4)) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 3) 2)) (poly.var 5)) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 3) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 4) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 1)) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 0) 2)) (poly.var 3))) (poly.mul (poly.const -1/1) (poly.pow (poly.var 3) 3))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 0) 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 3))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1)))\",\"(poly.const -1/1)\"],\"success\":true}" ["ff", "K", "7", "[((var6 ^ 2) - ((var2 ^ 2) + (var1 ^ 3))), ((var3 ^ 3) - (var2 + var6)), ((var4 * var3) - var1), (((1 + var5) + (var5 ^ 2)) - 0), ((2 * var2) - ((var3 ^ 3) - (var4 ^ 3)))]", "((((var0 ^ 3) + ((3 * var1) * var0)) - (2 * var2)) - (((var0 - (var3 - var4)) * (var0 - ((var3 * var5) - (var4 * (var5 ^ 2))))) * (var0 - ((var3 * (var5 ^ 2)) - (var4 * var5)))))"] "linear_combination (1 * x * ω ^ 4 + (-1) * s * ω ^ 4 + 1 * t * ω ^ 4 + (-1) * s * ω ^ 3 + 1 * t * ω ^ 3 + + test_polyrith + "{\"data\":[\"(poly.const -1/1)\",\"(poly.add (poly.add (poly.mul (poly.const -1/1) (poly.pow (poly.var 1) 3)) (poly.mul (poly.const 1/1) (poly.var 0))) (poly.mul (poly.const -1/1) (poly.var 3)))\",\"(poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 1) 2)) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 2)))\",\"(poly.const 0/1)\"],\"success\":true}" + ["ff", + "K", + "6", + "[((var3 ^ 2) - ((var0 ^ 2) + (var4 ^ 3))), ((var1 ^ 3) - (var0 + var3)), ((var2 * var1) - var4), (((1 + var5) + (var5 ^ 2)) - 0)]", + "(((2 * var0) * (var1 ^ 3)) - (((var1 ^ 3) - (var2 ^ 3)) * (var1 ^ 3)))"] + "linear_combination -hr + ((-1) * s ^ 3 + 1 * q + (-1) * r) * hs3 + (1 * s ^ 2 * t ^ 2 + 1 * s * t * p + 1 * p ^ 2) * ht"}, + test_polyrith + "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.pow (poly.var 5) 4)) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 3) 2)) (poly.var 5)) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 3) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 4) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 1)) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 0) 2)) (poly.var 3))) (poly.mul (poly.const -1/1) (poly.pow (poly.var 3) 3))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 0) 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 3))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1)))\",\"(poly.const -1/1)\"],\"success\":true}" + ["ff", + "K", + "7", + "[((var6 ^ 2) - ((var2 ^ 2) + (var1 ^ 3))), ((var3 ^ 3) - (var2 + var6)), ((var4 * var3) - var1), (((1 + var5) + (var5 ^ 2)) - 0), ((2 * var2) - ((var3 ^ 3) - (var4 ^ 3)))]", + "((((var0 ^ 3) + ((3 * var1) * var0)) - (2 * var2)) - (((var0 - (var3 - var4)) * (var0 - ((var3 * var5) - (var4 * (var5 ^ 2))))) * (var0 - ((var3 * (var5 ^ 2)) - (var4 * var5)))))"] + "linear_combination (1 * x * ω ^ 4 + (-1) * s * ω ^ 4 + 1 * t * ω ^ 4 + (-1) * s * ω ^ 3 + 1 * t * ω ^ 3 + 3 * x * ω ^ 2 + (-1) * s * ω ^ 2 + 1 * t * ω ^ 2 + @@ -229,22 +378,64 @@ end /-! ## Degenerate cases -/ example {K : Type*} [field K] [char_zero K] {s : K} (hs : 3 * s + 1 = 4) : s = 1 := -by test_polyrith "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" ["ff", "K", "1", "[(((3 * var0) + 1) - 4)]", "(var0 - 1)"] "linear_combination hs / 3" +by test_polyrith + "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" + ["ff", + "K", + "1", + "[(((3 * var0) + 1) - 4)]", + "(var0 - 1)"] + "linear_combination hs / 3" example {x : ℤ} (h1 : x + 4 = 2) : x = -2 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "1", "[((var0 + 4) - 2)]", "(var0 - (-1 * 2))"] "linear_combination h1" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "int", + "1", + "[((var0 + 4) - 2)]", + "(var0 - (-1 * 2))"] + "linear_combination h1" example {w : ℚ} (h1 : 3 * w + 1 = 4) : w = 1 := -by test_polyrith "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" ["ff", "rat", "1", "[(((3 * var0) + 1) - 4)]", "(var0 - 1)"] "linear_combination h1 / 3" +by test_polyrith + "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" + ["ff", + "rat", + "1", + "[(((3 * var0) + 1) - 4)]", + "(var0 - 1)"] + "linear_combination h1 / 3" example {x : ℤ} (h1 : 2 * x + 3 = x) : x = -3 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "1", "[(((2 * var0) + 3) - var0)]", "(var0 - (-1 * 3))"] "linear_combination h1" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "int", + "1", + "[(((2 * var0) + 3) - var0)]", + "(var0 - (-1 * 3))"] + "linear_combination h1" example {c : ℚ} (h1 : 4 * c + 1 = 3 * c - 2) : c = -3 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" ["ff", "rat", "1", "[(((4 * var0) + 1) - ((3 * var0) - 2))]", "(var0 - (-1 * 3))"] "linear_combination h1" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + ["ff", + "rat", + "1", + "[(((4 * var0) + 1) - ((3 * var0) - 2))]", + "(var0 - (-1 * 3))"] + "linear_combination h1" example (z : ℤ) (h1 : z + 1 = 2) (h2 : z + 2 = 2) : (1 : ℤ) = 2 := -by test_polyrith "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\"],\"success\":true}" ["ff", "int", "1", "[((var0 + 1) - 2), ((var0 + 2) - 2)]", "(1 - 2)"] "linear_combination h1 - h2" +by test_polyrith + "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\"],\"success\":true}" + ["ff", + "int", + "1", + "[((var0 + 1) - 2), ((var0 + 2) - 2)]", + "(1 - 2)"] + "linear_combination h1 - h2" -- We comment the following tests so that we don't overwhelm the SageCell API. @@ -393,3 +584,134 @@ begin linear_combination 3 * h a b + hqc, end -/ + + +-- the following can be uncommented to regenerate the tests above. + +/- + + +/-! ### Standard Cases over ℤ, ℚ, and ℝ -/ + +example (x y : ℤ) (h1 : 3*x + 2*y = 10): + 3*x + 2*y = 10 := +by create_polyrith_test + +example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : + x*y = -2*y + 1 := +by create_polyrith_test + +example (x y : ℝ) (h1 : x + 2 = -3) (h2 : y = 10) : + -y + 2*x + 4 = -16 := +by create_polyrith_test + +example (x y z : ℝ) (ha : x + 2*y - z = 4) (hb : 2*x + y + z = -2) + (hc : x + 2*y + z = 2) : + -3*x - 3*y - 4*z = 2 := +by create_polyrith_test + +example (w x y z : ℝ) (h1 : x + 2.1*y + 2*z = 2) (h2 : x + 8*z + 5*w = -6.5) + (h3 : x + y + 5*z + 5*w = 3) : + x + 2.2*y + 2*z - 5*w = -8.5 := +by create_polyrith_test + +example (a b c d : ℚ) (h1 : a = 4) (h2 : 3 = b) (h3 : c*3 = d) (h4 : -d = a) : + 2*a - 3 + 9*c + 3*d = 8 - b + 3*d - 3*a := +by create_polyrith_test + +/-! ### Case with ambiguous identifiers-/ + +example («def evil» y : ℤ) (h1 : 3*«def evil» + 2*y = 10): + 3*«def evil» + 2*y = 10 := +by create_polyrith_test + +example («¥» y : ℤ) (h1 : 3*«¥» + 2*y = 10): + «¥» * (3*«¥» + 2*y) = 10 * «¥» := +by create_polyrith_test + +/-! ### Cases with arbitrary coefficients -/ + +example (a b : ℤ) (h : a = b) : + a * a = a * b := +by create_polyrith_test + +example (a b c : ℤ) (h : a = b) : + a * c = b * c := +by create_polyrith_test + +example (a b c : ℤ) (h1 : a = b) (h2 : b = 1) : + c * a + b = c * b + 1 := +by create_polyrith_test + +example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : + x*x*y + y*x*y + 6*x = 3*x*y + 14 := +by create_polyrith_test + +example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := +by create_polyrith_test + +/-! ### Cases with non-hypothesis inputs/input restrictions -/ + +example (a b : ℝ) (ha : 2*a = 4) (hab : 2*b = a - b) (hignore : 3 = a + b) : + b = 2 / 3 := +by create_polyrith_test only [ha, hab] + +constant term : ∀ a b : ℚ, a + b = 0 + +example (a b c d : ℚ) (h : a + b = 0) (h2: b + c = 0): a + b + c + d = 0 := +by create_polyrith_test only [term c d, h] + +constants (qc : ℚ) (hqc : qc = 2*qc) + +example (a b : ℚ) (h : ∀ p q : ℚ, p = q) : 3*a + qc = 3*b + 2*qc := +by create_polyrith_test [h a b, hqc] + +constant bad (q : ℚ) : q = 0 + +example (a b : ℚ) : a + b^3 = 0 := +by create_polyrith_test [bad a, bad (b^2)] + +/-! ### Case over arbitrary field/ring -/ + +example {α} [h : comm_ring α] {a b c d e f : α} (h1 : a*d = b*c) (h2 : c*f = e*d) : + c * (a*f - b*e) = 0 := +by create_polyrith_test + +example {K : Type*} [field K] [invertible 2] [invertible 3] + {ω p q r s t x: K} (hp_nonzero : p ≠ 0) (hr : r ^ 2 = q ^ 2 + p ^ 3) (hs3 : s ^ 3 = q + r) + (ht : t * s = p) (x : K) (H : 1 + ω + ω ^ 2 = 0) : + x ^ 3 + 3 * p * x - 2 * q = + (x - (s - t)) * (x - (s * ω - t * ω ^ 2)) * (x - (s * ω ^ 2 - t * ω)) := +begin + have hs_nonzero : s ≠ 0, + { contrapose! hp_nonzero with hs_nonzero, + create_polyrith_test }, + have H' : 2 * q = s ^ 3 - t ^ 3, + { rw ← mul_left_inj' (pow_ne_zero 3 hs_nonzero), + create_polyrith_test }, + create_polyrith_test +end + + +/-! ## Degenerate cases -/ + +example {K : Type*} [field K] [char_zero K] {s : K} (hs : 3 * s + 1 = 4) : s = 1 := +by create_polyrith_test + +example {x : ℤ} (h1 : x + 4 = 2) : x = -2 := +by create_polyrith_test + +example {w : ℚ} (h1 : 3 * w + 1 = 4) : w = 1 := +by create_polyrith_test + +example {x : ℤ} (h1 : 2 * x + 3 = x) : x = -3 := +by create_polyrith_test + +example {c : ℚ} (h1 : 4 * c + 1 = 3 * c - 2) : c = -3 := +by create_polyrith_test + +example (z : ℤ) (h1 : z + 1 = 2) (h2 : z + 2 = 2) : (1 : ℤ) = 2 := +by create_polyrith_test + + +-/ From 6efbd672f55fcf6d1df071fc1a0d69f1d09a1575 Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Thu, 28 Jul 2022 15:05:15 +0000 Subject: [PATCH 127/173] feat(data/pi/algebra): sum_elim lemmas (#15588) --- src/data/int/cast.lean | 5 +++++ src/data/nat/cast.lean | 5 +++++ src/data/pi/algebra.lean | 36 ++++++++++++++++++++++++++++++++++++ src/data/sum/basic.lean | 36 ++++++++++++++++++++++++++++++++++++ 4 files changed, 82 insertions(+) diff --git a/src/data/int/cast.lean b/src/data/int/cast.lean index b0b247760d5a1..dd553a423cb9f 100644 --- a/src/data/int/cast.lean +++ b/src/data/int/cast.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import data.int.basic import data.nat.cast import tactic.pi_instances +import data.sum.basic /-! # Cast of integers (additional theorems) @@ -270,6 +271,10 @@ lemma int_apply (n : ℤ) (a : α) : (n : ∀ a, β a) a = n := rfl end pi +lemma sum.elim_int_cast_int_cast {α β γ : Type*} [has_int_cast γ] (n : ℤ) : + sum.elim (n : α → γ) (n : β → γ) = n := +@sum.elim_lam_const_lam_const α β γ n + namespace pi variables {α : Type*} {β : α → Type*} [∀ a, add_group_with_one (β a)] diff --git a/src/data/nat/cast.lean b/src/data/nat/cast.lean index 972a4b3822453..17c0199b286ec 100644 --- a/src/data/nat/cast.lean +++ b/src/data/nat/cast.lean @@ -7,6 +7,7 @@ import data.nat.basic import data.nat.cast.defs import algebra.group.pi import tactic.pi_instances +import data.sum.basic /-! # Cast of natural numbers (additional theorems) @@ -274,6 +275,10 @@ lemma nat_apply (n : ℕ) (a : α) : (n : ∀ a, β a) a = n := rfl end pi +lemma sum.elim_nat_cast_nat_cast {α β γ : Type*} [has_nat_cast γ] (n : ℕ) : + sum.elim (n : α → γ) (n : β → γ) = n := +@sum.elim_lam_const_lam_const α β γ n + namespace pi variables {α : Type*} {β : α → Type*} [∀ a, add_monoid_with_one (β a)] diff --git a/src/data/pi/algebra.lean b/src/data/pi/algebra.lean index 964db17b1cbb2..be18220129284 100644 --- a/src/data/pi/algebra.lean +++ b/src/data/pi/algebra.lean @@ -9,6 +9,7 @@ import logic.unique import tactic.congr import tactic.simpa import tactic.split_ifs +import data.sum.basic /-! # Instances and theorems on pi types @@ -225,3 +226,38 @@ lemma subsingleton.pi_mul_single_eq {α : Type*} [decidable_eq I] [subsingleton (i : I) (x : α) : pi.mul_single i x = λ _, x := funext $ λ j, by rw [subsingleton.elim j i, pi.mul_single_eq_same] + +namespace sum +variables (a a' : α → γ) (b b' : β → γ) + +@[simp, to_additive] +lemma elim_one_one [has_one γ] : + sum.elim (1 : α → γ) (1 : β → γ) = 1 := +sum.elim_const_const 1 + +@[simp, to_additive] +lemma elim_mul_single_one [decidable_eq α] [decidable_eq β] [has_one γ] (i : α) (c : γ) : + sum.elim (pi.mul_single i c) (1 : β → γ) = pi.mul_single (sum.inl i) c := +by simp only [pi.mul_single, sum.elim_update_left, elim_one_one] + +@[simp, to_additive] +lemma elim_one_mul_single [decidable_eq α] [decidable_eq β] [has_one γ] (i : β) (c : γ) : + sum.elim (1 : α → γ) (pi.mul_single i c) = pi.mul_single (sum.inr i) c := +by simp only [pi.mul_single, sum.elim_update_right, elim_one_one] + +@[to_additive] +lemma elim_inv_inv [has_inv γ] : + sum.elim a⁻¹ b⁻¹ = (sum.elim a b)⁻¹ := +(sum.comp_elim has_inv.inv a b).symm + +@[to_additive] +lemma elim_mul_mul [has_mul γ] : + sum.elim (a * a') (b * b') = sum.elim a b * sum.elim a' b' := +by { ext x, cases x; refl } + +@[to_additive] +lemma elim_div_div [has_div γ] : + sum.elim (a / a') (b / b') = sum.elim a b / sum.elim a' b' := +by { ext x, cases x; refl } + +end sum diff --git a/src/data/sum/basic.lean b/src/data/sum/basic.lean index 590a4957cca64..0a87ab9020196 100644 --- a/src/data/sum/basic.lean +++ b/src/data/sum/basic.lean @@ -346,6 +346,42 @@ lemma surjective.sum_map {f : α → β} {g : α' → β'} (hf : surjective f) ( end function +namespace sum +open function + +lemma elim_const_const (c : γ) : + sum.elim (const _ c : α → γ) (const _ c : β → γ) = const _ c := +by { ext x, cases x; refl } + +@[simp] +lemma elim_lam_const_lam_const (c : γ) : + sum.elim (λ (_ : α), c) (λ (_ : β), c) = λ _, c := +sum.elim_const_const c + +lemma elim_update_left [decidable_eq α] [decidable_eq β] + (f : α → γ) (g : β → γ) (i : α) (c : γ) : + sum.elim (function.update f i c) g = function.update (sum.elim f g) (inl i) c := +begin + ext x, cases x, + { by_cases h : x = i, + { subst h, simp }, + { simp [h] } }, + { simp } +end + +lemma elim_update_right [decidable_eq α] [decidable_eq β] + (f : α → γ) (g : β → γ) (i : β) (c : γ) : + sum.elim f (function.update g i c) = function.update (sum.elim f g) (inr i) c := +begin + ext x, cases x, + { simp }, + { by_cases h : x = i, + { subst h, simp }, + { simp [h] } } +end + +end sum + /-! ### Ternary sum From 05244cc7e0fd544e4e105d44d2970f6616ace235 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Thu, 28 Jul 2022 15:05:16 +0000 Subject: [PATCH 128/173] feat(algebra/module): sub_mem_sup for modules over rings (#15733) --- src/algebra/module/submodule/lattice.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/algebra/module/submodule/lattice.lean b/src/algebra/module/submodule/lattice.lean index 9edf95ad147fb..fc514aadcca18 100644 --- a/src/algebra/module/submodule/lattice.lean +++ b/src/algebra/module/submodule/lattice.lean @@ -213,6 +213,14 @@ show T ≤ S ⊔ T, from le_sup_right lemma add_mem_sup {S T : submodule R M} {s t : M} (hs : s ∈ S) (ht : t ∈ T) : s + t ∈ S ⊔ T := add_mem (mem_sup_left hs) (mem_sup_right ht) +lemma sub_mem_sup {R' M' : Type*} [ring R'] [add_comm_group M'] [module R' M'] + {S T : submodule R' M'} {s t : M'} (hs : s ∈ S) (ht : t ∈ T) : + s - t ∈ S ⊔ T := +begin + rw sub_eq_add_neg, + exact add_mem_sup hs (neg_mem ht), +end + lemma mem_supr_of_mem {ι : Sort*} {b : M} {p : ι → submodule R M} (i : ι) (h : b ∈ p i) : b ∈ (⨆i, p i) := have p i ≤ (⨆i, p i) := le_supr p i, From c04b5a9223b245a68a4d8af80d53d8769c283708 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 28 Jul 2022 17:47:20 +0000 Subject: [PATCH 129/173] feat(data/{pi, prod}): add missing `has_pow` instances for `pi` type (#15478) This generalizes the existing powers by `nat` and `int` defined in later files to be defined for arbitrary powers. As well as eliminating the need for separate lemmas for these different power operations, this also means that tuples of real numbers now inherit a power operation by the real numbers. Co-authored-by: Eric Wieser --- src/algebra/group/pi.lean | 4 --- src/data/pi/algebra.lean | 35 ++++++++++++++++++- src/group_theory/group_action/pi.lean | 10 ------ src/group_theory/group_action/prod.lean | 20 ++++++++++- .../continuous_function/polynomial.lean | 2 +- 5 files changed, 54 insertions(+), 17 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 5a8c7e8fa2e9e..d1685c1e91563 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -43,10 +43,6 @@ instance monoid [∀ i, monoid $ f i] : monoid (Π i : I, f i) := by refine_struct { one := (1 : Π i, f i), mul := (*), npow := λ n x i, (x i) ^ n }; tactic.pi_instance_derive_field --- the attributes are intentionally out of order. `smul_apply` proves `nsmul_apply`. -@[to_additive, simp] -lemma pow_apply [∀ i, monoid $ f i] (n : ℕ) : (x^n) i = (x i)^n := rfl - @[to_additive] instance comm_monoid [∀ i, comm_monoid $ f i] : comm_monoid (Π i : I, f i) := by refine_struct { one := (1 : Π i, f i), mul := (*), npow := monoid.npow }; diff --git a/src/data/pi/algebra.lean b/src/data/pi/algebra.lean index be18220129284..4276fe8c609fc 100644 --- a/src/data/pi/algebra.lean +++ b/src/data/pi/algebra.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Eric Wieser -/ import algebra.group.to_additive +import algebra.group.defs import data.prod.basic import logic.unique import tactic.congr @@ -30,7 +31,7 @@ variables (x y : Π i, f i) (i : I) namespace pi -/-! `1`, `0`, `+`, `*`, `-`, `⁻¹`, and `/` are defined pointwise. -/ +/-! `1`, `0`, `+`, `*`, `+ᵥ`, `•`, `^`, `-`, `⁻¹`, and `/` are defined pointwise. -/ @[to_additive] instance has_one [∀ i, has_one $ f i] : has_one (Π i : I, f i) := @@ -59,6 +60,38 @@ instance has_mul [∀ i, has_mul $ f i] : @[to_additive] lemma mul_comp [has_mul γ] (x y : β → γ) (z : α → β) : (x * y) ∘ z = x ∘ z * y ∘ z := rfl +@[to_additive pi.has_vadd] instance has_smul [Π i, has_smul α $ f i] : has_smul α (Π i : I, f i) := +⟨λ s x, λ i, s • (x i)⟩ + +@[simp, to_additive] lemma smul_apply [Π i, has_smul α $ f i] (s : α) (x : Π i, f i) (i : I) : + (s • x) i = s • x i := rfl + +@[to_additive] lemma smul_def [Π i, has_smul α $ f i] (s : α) (x : Π i, f i) : + s • x = λ i, s • x i := rfl + +@[simp, to_additive] +lemma smul_const [has_smul α β] (a : α) (b : β) : a • const I b = const I (a • b) := rfl + +@[to_additive] +lemma smul_comp [has_smul α γ] (a : α) (x : β → γ) (y : I → β) : (a • x) ∘ y = a • (x ∘ y) := rfl + +@[to_additive pi.has_smul] +instance has_pow [Π i, has_pow (f i) β] : has_pow (Π i, f i) β := +⟨λ x b i, (x i) ^ b⟩ + +@[simp, to_additive pi.smul_apply, to_additive_reorder 5] +lemma pow_apply [Π i, has_pow (f i) β] (x : Π i, f i) (b : β) (i : I) : (x ^ b) i = (x i) ^ b := rfl + +@[to_additive pi.smul_def, to_additive_reorder 5] +lemma pow_def [Π i, has_pow (f i) β] (x : Π i, f i) (b : β) : x ^ b = λ i, (x i) ^ b := rfl + +-- `to_additive` generates bad output if we take `has_pow α β`. +@[simp, to_additive smul_const, to_additive_reorder 5] +lemma const_pow [has_pow β α] (b : β) (a : α) : const I b ^ a = const I (b ^ a) := rfl + +@[to_additive smul_comp, to_additive_reorder 6] +lemma pow_comp [has_pow γ α] (x : β → γ) (a : α) (y : I → β) : (x ^ a) ∘ y = (x ∘ y) ^ a := rfl + @[simp] lemma bit0_apply [Π i, has_add $ f i] : (bit0 x) i = bit0 (x i) := rfl @[simp] lemma bit1_apply [Π i, has_add $ f i] [Π i, has_one $ f i] : (bit1 x) i = bit1 (x i) := rfl diff --git a/src/group_theory/group_action/pi.lean b/src/group_theory/group_action/pi.lean index 3c9e58aa85643..0686e87d3466e 100644 --- a/src/group_theory/group_action/pi.lean +++ b/src/group_theory/group_action/pi.lean @@ -26,16 +26,6 @@ variables (x y : Π i, f i) (i : I) namespace pi -@[to_additive pi.has_vadd] -instance has_smul {α : Type*} [Π i, has_smul α $ f i] : - has_smul α (Π i : I, f i) := -⟨λ s x, λ i, s • (x i)⟩ - -@[to_additive] -lemma smul_def {α : Type*} [Π i, has_smul α $ f i] (s : α) : s • x = λ i, s • x i := rfl -@[simp, to_additive] -lemma smul_apply {α : Type*} [Π i, has_smul α $ f i] (s : α) : (s • x) i = s • x i := rfl - @[to_additive pi.has_vadd'] instance has_smul' {g : I → Type*} [Π i, has_smul (f i) (g i)] : has_smul (Π i, f i) (Π i : I, g i) := diff --git a/src/group_theory/group_action/prod.lean b/src/group_theory/group_action/prod.lean index 007e861c241a7..9b2a8d66b702a 100644 --- a/src/group_theory/group_action/prod.lean +++ b/src/group_theory/group_action/prod.lean @@ -25,7 +25,7 @@ scalar multiplication as a homomorphism from `α × β` to `β`. * `group_theory.group_action.sum` -/ -variables {M N P α β : Type*} +variables {M N P E α β : Type*} namespace prod @@ -41,6 +41,24 @@ variables [has_smul M α] [has_smul M β] [has_smul N α] [has_smul N β] (a : M @[to_additive] theorem smul_def (a : M) (x : α × β) : a • x = (a • x.1, a • x.2) := rfl @[simp, to_additive] theorem smul_swap : (a • x).swap = a • x.swap := rfl + +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] +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 +/- Note that the `c` arguments to this lemmas cannot be in the more natural right-most positions due +to limitations in `to_additive` and `to_additive_reorder`, which will silently fail to reorder more +than two adjacent arguments -/ +@[simp, to_additive smul_mk, to_additive_reorder 6] +lemma pow_mk (c : E) (a : α) (b : β) : (prod.mk a b) ^ c = prod.mk (a ^ c) (b ^ c) := rfl +@[to_additive smul_def, to_additive_reorder 6] +lemma pow_def (p : α × β) (c : E) : p ^ c = (p.1 ^ c, p.2 ^ c) := rfl +@[simp, to_additive smul_swap, to_additive_reorder 6] +lemma pow_swap (p : α × β) (c : E) : (p ^ c).swap = p.swap ^ c := rfl + instance [has_smul M N] [is_scalar_tower M N α] [is_scalar_tower M N β] : is_scalar_tower M N (α × β) := ⟨λ x y z, mk.inj_iff.mpr ⟨smul_assoc _ _ _, smul_assoc _ _ _⟩⟩ diff --git a/src/topology/continuous_function/polynomial.lean b/src/topology/continuous_function/polynomial.lean index b88303370be39..04ae3993512da 100644 --- a/src/topology/continuous_function/polynomial.lean +++ b/src/topology/continuous_function/polynomial.lean @@ -61,7 +61,7 @@ variables {α : Type*} [topological_space α] begin apply polynomial.induction_on' g, { intros p q hp hq, simp [hp, hq], }, - { intros n a, simp [pi.pow_apply f x n], }, + { intros n a, simp [pi.pow_apply], }, end end From e4c7baa4693a4d2cd304cc097a459f034aa187f8 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 28 Jul 2022 20:15:18 +0000 Subject: [PATCH 130/173] fix(tactic/polyrith): remove unnecessary 1* from polyrith output (#15292) Some `polyrith` certificates were printing with a superfluous `1*` in front. Co-authored-by: Rob Lewis --- scripts/polyrith_sage_helper.py | 49 +++++++++++++++--- src/tactic/polyrith.lean | 21 ++++++-- test/polyrith.lean | 88 ++++++++++++++++----------------- 3 files changed, 102 insertions(+), 56 deletions(-) diff --git a/scripts/polyrith_sage_helper.py b/scripts/polyrith_sage_helper.py index 08a93a62854b7..8da7c5d0c881b 100644 --- a/scripts/polyrith_sage_helper.py +++ b/scripts/polyrith_sage_helper.py @@ -1,5 +1,7 @@ # this file will be run by the remote sage server, so should not import local files. from functools import reduce +from typing import Optional, Iterator +from attr import dataclass from sage.rings.polynomial.polydict import ETuple def mk_app(*args: str) -> str: @@ -15,20 +17,51 @@ def power_to_string(var: int, pow: int) -> str: return var_s return mk_app("poly.pow", var_s, str(pow)) -def sum_to_string(terms: list[str]) -> str: - terms = iter(terms) +@dataclass +class MonomForm: + """ + `MonomForm` stores the string representation of a monomial. + + To cleanly format sums of monomials, we need to be careful with negation: + if the first monomial in a sum is negative, we print the negation symbol; + if a subsequent monomial is negative, we subtract the non-negated version. + `MonomForm` always stores the *positive* representation of a monomial in the `pos_form` field. + If the monomial is in fact negative, it also stores the full (negative) representation + in the `neg_form` field. For instance, putting `-2*x` into a `MonomForm` would store + the representation of `2*x` in `pos_form` and `-2*x` in `neg_form`. + """ + pos_form: str + neg_form: Optional[str] = None + +def sum_to_string(terms: Iterator[MonomForm]) -> str: try: first = next(terms) except StopIteration: return const_to_string(QQ(0)) else: - return reduce(lambda s, t: mk_app("poly.add", s, t), terms, first) + first_form = first.neg_form if first.neg_form is not None else first.pos_form + return reduce( + lambda old, nxt: mk_app("poly.sub" if nxt.neg_form is not None else "poly.add", old, nxt.pos_form), + terms, first_form) + +def monomial_to_string(etuple: ETuple, coeff: QQ) -> MonomForm: + etuple = list(etuple.sparse_iter()) + if abs(coeff) == 1 and len(etuple) > 0: + powforms = [power_to_string(t[0], t[1]) for t in etuple] + pos_form = reduce( + lambda s, t: mk_app("poly.mul", s, t), powforms) + return MonomForm(pos_form, mk_app("poly.neg", pos_form) if coeff < 0 else None) + else: + pos_form = reduce( + lambda s, t: mk_app("poly.mul", s, power_to_string(t[0], t[1])), + etuple, + const_to_string(abs(coeff))) + neg_form = reduce( + lambda s, t: mk_app("poly.mul", s, power_to_string(t[0], t[1])), + etuple, + const_to_string(coeff)) if coeff < 0 else None + return MonomForm(pos_form, neg_form) -def monomial_to_string(etuple: ETuple, coeff: QQ) -> str: - return reduce( - lambda s, t: mk_app("poly.mul", s, power_to_string(t[0], t[1])), - etuple.sparse_iter(), - const_to_string(coeff)) def polynomial_to_string(p) -> str: return sum_to_string(monomial_to_string(pows, coeff) for pows, coeff in p.dict().items()) diff --git a/src/tactic/polyrith.lean b/src/tactic/polyrith.lean index 3c21f0797e209..17c6ad9872e9c 100644 --- a/src/tactic/polyrith.lean +++ b/src/tactic/polyrith.lean @@ -74,6 +74,7 @@ inductive poly | sub : poly → poly → poly | mul : poly → poly → poly | pow : poly → ℕ → poly +| neg : poly → poly /-- This converts a poly object into a string representing it. The string @@ -89,11 +90,13 @@ meta def poly.mk_string : poly → string | (poly.sub p q) := "(" ++ poly.mk_string p ++ " - " ++ poly.mk_string q ++ ")" | (poly.mul p q) := "(" ++ poly.mk_string p ++ " * " ++ poly.mk_string q ++ ")" | (poly.pow p n) := to_string $ format!"({poly.mk_string p} ^ {n})" +| (poly.neg p) := ("-" ++ poly.mk_string p) meta instance : has_add poly := ⟨poly.add⟩ meta instance : has_sub poly := ⟨poly.sub⟩ meta instance : has_mul poly := ⟨poly.mul⟩ meta instance : has_pow poly ℕ := ⟨poly.pow⟩ +meta instance : has_neg poly := ⟨poly.neg⟩ meta instance : has_repr poly := ⟨poly.mk_string⟩ meta instance : has_to_format poly := ⟨to_fmt ∘ poly.mk_string⟩ meta instance : inhabited poly := ⟨poly.const 0⟩ @@ -149,7 +152,7 @@ meta def poly_form_of_expr (red : transparency) : list expr → expr → tactic return (m', comp1 - comp2) | m `(-%%e) := do (m', comp) ← poly_form_of_expr m e, - return (m', (poly.const (-1)) * comp) + return (m', - comp) | m p@`(@has_pow.pow _ ℕ _ %%e %%n) := match n.to_nat with | some k := @@ -207,7 +210,10 @@ meta def poly.to_pexpr : list expr → poly → tactic pexpr do p_pexpr ← poly.to_pexpr m p, return ``(%%p_pexpr ^ %%n.to_pexpr) - +| m (poly.neg p) := + do + p_pexpr ← poly.to_pexpr m p, + return ``(- %%p_pexpr) /-! # Parsing SageMath output into a poly @@ -268,12 +274,21 @@ and `n` is a natural number. meta def pow_parser (cont : parser poly) : parser poly := str "poly.pow " >> poly.pow <$> cont <*> (ch ' ' >> nat) +/-- +A parser object that parses `string`s of the form `"poly.neg p"` +to the appropriate `poly` object representing the negation of a `poly`. +Here, `p` is the string form of a `poly`. +-/ +meta def neg_parser (cont : parser poly) : parser poly := +str "poly.neg " >> poly.neg <$> cont + /-- A parser for `poly` that uses an s-essresion style formats such as `(poly.add (poly.var 0) (poly.const 1)`. -/ meta def poly_parser : parser poly := ch '(' *> (var_parser <|> const_fraction_parser <|> add_parser poly_parser - <|> sub_parser poly_parser <|> mul_parser poly_parser <|> pow_parser poly_parser) + <|> sub_parser poly_parser <|> mul_parser poly_parser <|> pow_parser poly_parser + <|> neg_parser poly_parser) <* ch ')' meta instance : non_null_json_serializable poly := diff --git a/test/polyrith.lean b/test/polyrith.lean index c8cf2c4f579ec..c3b71514e1e42 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -113,6 +113,8 @@ begin linear_combination 2 * h1 + h2 - 2 * h3 end + + /-! ### Standard Cases over ℤ, ℚ, and ℝ -/ example (x y : ℤ) (h1 : 3*x + 2*y = 10): @@ -134,7 +136,7 @@ by test_polyrith "rat", "2", "[(((var0 * var1) + (2 * var0)) - 1), (var0 - var1)]", - "((var0 * var1) - (((-1 * 2) * var1) + 1))"] + "((var0 * var1) - ((-2 * var1) + 1))"] "linear_combination h1 - 2 * h2" example (x y : ℝ) (h1 : x + 2 = -3) (h2 : y = 10) : @@ -144,8 +146,8 @@ by test_polyrith ["ff", "real", "2", - "[((var1 + 2) - (-1 * 3)), (var0 - 10)]", - "((((-1 * var0) + (2 * var1)) + 4) - (-1 * 16))"] + "[((var1 + 2) - -3), (var0 - 10)]", + "(((-var0 + (2 * var1)) + 4) - -16)"] "linear_combination 2 * h1 - h2" example (x y z : ℝ) (ha : x + 2*y - z = 4) (hb : 2*x + y + z = -2) @@ -156,8 +158,8 @@ by test_polyrith ["ff", "real", "3", - "[(((var0 + (2 * var1)) - var2) - 4), ((((2 * var0) + var1) + var2) - (-1 * 2)), (((var0 + (2 * var1)) + var2) - 2)]", - "(((((-1 * 3) * var0) - (3 * var1)) - (4 * var2)) - 2)"] + "[(((var0 + (2 * var1)) - var2) - 4), ((((2 * var0) + var1) + var2) - -2), (((var0 + (2 * var1)) + var2) - 2)]", + "((((-3 * var0) - (3 * var1)) - (4 * var2)) - 2)"] "linear_combination ha - hb - 2 * hc" example (w x y z : ℝ) (h1 : x + 2.1*y + 2*z = 2) (h2 : x + 8*z + 5*w = -6.5) @@ -168,8 +170,8 @@ by test_polyrith ["ff", "real", "4", - "[(((var0 + (21/10 * var1)) + (2 * var2)) - 2), (((var0 + (8 * var2)) + (5 * var3)) - (-1 * 13/2)), ((((var0 + var1) + (5 * var2)) + (5 * var3)) - 3)]", - "((((var0 + (11/5 * var1)) + (2 * var2)) - (5 * var3)) - (-1 * 17/2))"] + "[(((var0 + (21/10 * var1)) + (2 * var2)) - 2), (((var0 + (8 * var2)) + (5 * var3)) - -13/2), ((((var0 + var1) + (5 * var2)) + (5 * var3)) - 3)]", + "((((var0 + (11/5 * var1)) + (2 * var2)) - (5 * var3)) - -17/2)"] "linear_combination 2 * h1 + h2 - 2 * h3" example (a b c d : ℚ) (h1 : a = 4) (h2 : 3 = b) (h3 : c*3 = d) (h4 : -d = a) : @@ -179,7 +181,7 @@ by test_polyrith ["ff", "rat", "4", - "[(var0 - 4), (3 - var3), ((var1 * 3) - var2), ((-1 * var2) - var0)]", + "[(var0 - 4), (3 - var3), ((var1 * 3) - var2), (-var2 - var0)]", "(((((2 * var0) - 3) + (9 * var1)) + (3 * var2)) - (((8 - var3) + (3 * var2)) - (3 * var0)))"] "linear_combination 2 * h1 - h2 + 3 * h3 - 3 * h4" @@ -199,69 +201,69 @@ by test_polyrith example («¥» y : ℤ) (h1 : 3*«¥» + 2*y = 10): «¥» * (3*«¥» + 2*y) = 10 * «¥» := by test_polyrith - "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" + "{\"data\":[\"(poly.var 0)\"],\"success\":true}" ["ff", "int", "2", "[(((3 * var0) + (2 * var1)) - 10)]", "((var0 * ((3 * var0) + (2 * var1))) - (10 * var0))"] - "linear_combination 1 * «¥» * h1" + "linear_combination «¥» * h1" /-! ### Cases with arbitrary coefficients -/ example (a b : ℤ) (h : a = b) : a * a = a * b := by test_polyrith - "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\"],\"success\":true}" + "{\"data\":[\"(poly.var 0)\"],\"success\":true}" ["ff", "int", "2", "[(var0 - var1)]", "((var0 * var0) - (var0 * var1))"] - "linear_combination 1 * a * h" + "linear_combination a * h" example (a b c : ℤ) (h : a = b) : a * c = b * c := by test_polyrith - "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + "{\"data\":[\"(poly.var 1)\"],\"success\":true}" ["ff", "int", "3", "[(var0 - var2)]", "((var0 * var1) - (var2 * var1))"] - "linear_combination 1 * c * h" + "linear_combination c * h" example (a b c : ℤ) (h1 : a = b) (h2 : b = 1) : c * a + b = c * b + 1 := by test_polyrith - "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 0))\",\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":[\"(poly.var 0)\",\"(poly.const 1/1)\"],\"success\":true}" ["ff", "int", "3", "[(var1 - var2), (var2 - 1)]", "(((var0 * var1) + var2) - ((var0 * var2) + 1))"] - "linear_combination 1 * c * h1 + h2" + "linear_combination c * h1 + h2" example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : x*x*y + y*x*y + 6*x = 3*x*y + 14 := by test_polyrith - "{\"data\":[\"(poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1))\",\"(poly.const 2/1)\"],\"success\":true}" + "{\"data\":[\"(poly.mul (poly.var 0) (poly.var 1))\",\"(poly.const 2/1)\"],\"success\":true}" ["ff", "rat", "2", "[((var0 + var1) - 3), ((3 * var0) - 7)]", "(((((var0 * var0) * var1) + ((var1 * var0) * var1)) + (6 * var0)) - (((3 * var0) * var1) + 14))"] - "linear_combination 1 * x * y * h1 + 2 * h2" + "linear_combination x * y * h1 + 2 * h2" example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := by test_polyrith - "{\"data\":[\"(poly.add (poly.mul (poly.const 1/1) (poly.var 0)) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"success\":true}" + "{\"data\":[\"(poly.add (poly.var 0) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"success\":true}" ["ff", "rat", "4", "[(var1 - var3)]", "(((var0 * var1) + ((2 * var2) * var1)) - ((var0 * var3) + ((2 * var2) * var3)))"] - "linear_combination (1 * x + 2 * y) * hzw" + "linear_combination (x + 2 * y) * hzw" /-! ### Cases with non-hypothesis inputs/input restrictions -/ @@ -304,26 +306,26 @@ constant bad (q : ℚ) : q = 0 example (a b : ℚ) : a + b^3 = 0 := by test_polyrith [bad a, bad (b^2)] - "{\"data\":[\"(poly.const 1/1)\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + "{\"data\":[\"(poly.const 1/1)\",\"(poly.var 1)\"],\"success\":true}" ["ff", "rat", "2", "[(var0 - 0), ((var1 ^ 2) - 0)]", "((var0 + (var1 ^ 3)) - 0)"] - "linear_combination bad a + 1 * b * bad (b ^ 2)" + "linear_combination bad a + b * bad (b ^ 2)" /-! ### Case over arbitrary field/ring -/ example {α} [h : comm_ring α] {a b c d e f : α} (h1 : a*d = b*c) (h2 : c*f = e*d) : c * (a*f - b*e) = 0 := by test_polyrith - "{\"data\":[\"(poly.mul (poly.const 1/1) (poly.var 4))\",\"(poly.mul (poly.const 1/1) (poly.var 1))\"],\"success\":true}" + "{\"data\":[\"(poly.var 4)\",\"(poly.var 1)\"],\"success\":true}" ["ff", "α", "6", "[((var1 * var5) - (var3 * var0)), ((var0 * var2) - (var4 * var5))]", "((var0 * ((var1 * var2) - (var3 * var4))) - 0)"] - "linear_combination 1 * e * h1 + 1 * a * h2" + "linear_combination e * h1 + a * h2" example {K : Type*} [field K] [invertible 2] [invertible 3] {ω p q r s t x: K} (hp_nonzero : p ≠ 0) (hr : r ^ 2 = q ^ 2 + p ^ 3) (hs3 : s ^ 3 = q + r) @@ -334,43 +336,39 @@ begin have hs_nonzero : s ≠ 0, { contrapose! hp_nonzero with hs_nonzero, test_polyrith - "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.mul (poly.const 1/1) (poly.var 4))\"],\"success\":true}" + "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.var 4)\"],\"success\":true}" ["ff", "K", "6", "[((var1 ^ 2) - ((var2 ^ 2) + (var0 ^ 3))), ((var3 ^ 3) - (var2 + var1)), ((var4 * var3) - var0), (((1 + var5) + (var5 ^ 2)) - 0), (var3 - 0)]", "(var0 - 0)"] - "linear_combination -ht + 1 * t * hs_nonzero"}, + "linear_combination -ht + t * hs_nonzero"}, have H' : 2 * q = s ^ 3 - t ^ 3, { rw ← mul_left_inj' (pow_ne_zero 3 hs_nonzero), test_polyrith - "{\"data\":[\"(poly.const -1/1)\",\"(poly.add (poly.add (poly.mul (poly.const -1/1) (poly.pow (poly.var 1) 3)) (poly.mul (poly.const 1/1) (poly.var 0))) (poly.mul (poly.const -1/1) (poly.var 3)))\",\"(poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 1) 2)) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 2)))\",\"(poly.const 0/1)\"],\"success\":true}" + "{\"data\":[\"(poly.const -1/1)\",\"(poly.sub (poly.add (poly.neg (poly.pow (poly.var 1) 3)) (poly.var 0)) (poly.var 3))\",\"(poly.add (poly.add (poly.mul (poly.pow (poly.var 1) 2) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.var 1) (poly.var 2)) (poly.var 4))) (poly.pow (poly.var 4) 2))\",\"(poly.const 0/1)\"],\"success\":true}" ["ff", "K", "6", "[((var3 ^ 2) - ((var0 ^ 2) + (var4 ^ 3))), ((var1 ^ 3) - (var0 + var3)), ((var2 * var1) - var4), (((1 + var5) + (var5 ^ 2)) - 0)]", "(((2 * var0) * (var1 ^ 3)) - (((var1 ^ 3) - (var2 ^ 3)) * (var1 ^ 3)))"] - "linear_combination -hr + ((-1) * s ^ 3 + 1 * q + (-1) * r) * hs3 + (1 * s ^ 2 * t ^ 2 + 1 * s * t * p + 1 * p ^ 2) * ht"}, + "linear_combination -hr + (-s ^ 3 + q - r) * hs3 + (s ^ 2 * t ^ 2 + s * t * p + p ^ 2) * ht"}, test_polyrith - "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.pow (poly.var 5) 4)) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 4))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const -1/1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.add (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 3) 2)) (poly.var 5)) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 3) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 4) 3)) (poly.var 5))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 0)) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 1)) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.mul (poly.const 1/1) (poly.var 1)) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 1/1) (poly.pow (poly.var 0) 2)) (poly.var 3))) (poly.mul (poly.const -1/1) (poly.pow (poly.var 3) 3))) (poly.mul (poly.mul (poly.const -1/1) (poly.pow (poly.var 0) 2)) (poly.var 4))) (poly.mul (poly.const 1/1) (poly.pow (poly.var 4) 3))) (poly.mul (poly.mul (poly.mul (poly.const -1/1) (poly.var 0)) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1)))\",\"(poly.const -1/1)\"],\"success\":true}" + "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.mul (poly.var 0) (poly.pow (poly.var 5) 4)) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 3))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.add (poly.sub (poly.add (poly.sub (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.sub (poly.add (poly.neg (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 3) 2)) (poly.var 5))) (poly.mul (poly.pow (poly.var 3) 3) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.pow (poly.var 4) 3) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 3))) (poly.pow (poly.var 3) 3)) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 4))) (poly.pow (poly.var 4) 3)) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1)))\",\"(poly.const -1/1)\"],\"success\":true}" ["ff", "K", "7", "[((var6 ^ 2) - ((var2 ^ 2) + (var1 ^ 3))), ((var3 ^ 3) - (var2 + var6)), ((var4 * var3) - var1), (((1 + var5) + (var5 ^ 2)) - 0), ((2 * var2) - ((var3 ^ 3) - (var4 ^ 3)))]", "((((var0 ^ 3) + ((3 * var1) * var0)) - (2 * var2)) - (((var0 - (var3 - var4)) * (var0 - ((var3 * var5) - (var4 * (var5 ^ 2))))) * (var0 - ((var3 * (var5 ^ 2)) - (var4 * var5)))))"] - "linear_combination (1 * x * ω ^ 4 + (-1) * s * ω ^ 4 + 1 * t * ω ^ 4 + (-1) * s * ω ^ 3 + 1 * t * ω ^ 3 + - 3 * x * ω ^ 2 + - (-1) * s * ω ^ 2 + - 1 * t * ω ^ 2 + - 2 * x * ω) * ht + ((-1) * x * s ^ 2 * ω + 1 * s ^ 3 * ω + (-1) * x * t ^ 2 * ω + (-1) * t ^ 3 * ω + - 1 * x * p * ω ^ 2 + - (-1) * p * s * ω ^ 2 + - 1 * p * t * ω ^ 2 + - 1 * x ^ 2 * s + - (-1) * s ^ 3 + - (-1) * x ^ 2 * t + - 1 * t ^ 3 + - (-1) * x * p * ω + + "linear_combination (x * ω ^ 4 - s * ω ^ 4 + t * ω ^ 4 - s * ω ^ 3 + t * ω ^ 3 + 3 * x * ω ^ 2 - s * ω ^ 2 + + t * ω ^ 2 + + 2 * x * ω) * ht + (-(x * s ^ 2 * ω) + s ^ 3 * ω - x * t ^ 2 * ω - t ^ 3 * ω + x * p * ω ^ 2 - p * s * ω ^ 2 + + p * t * ω ^ 2 + + x ^ 2 * s - + s ^ 3 - + x ^ 2 * t + + t ^ 3 - + x * p * ω + 3 * x * p) * H - H'" end @@ -394,7 +392,7 @@ by test_polyrith "int", "1", "[((var0 + 4) - 2)]", - "(var0 - (-1 * 2))"] + "(var0 - -2)"] "linear_combination h1" example {w : ℚ} (h1 : 3 * w + 1 = 4) : w = 1 := @@ -414,7 +412,7 @@ by test_polyrith "int", "1", "[(((2 * var0) + 3) - var0)]", - "(var0 - (-1 * 3))"] + "(var0 - -3)"] "linear_combination h1" example {c : ℚ} (h1 : 4 * c + 1 = 3 * c - 2) : c = -3 := @@ -424,7 +422,7 @@ by test_polyrith "rat", "1", "[(((4 * var0) + 1) - ((3 * var0) - 2))]", - "(var0 - (-1 * 3))"] + "(var0 - -3)"] "linear_combination h1" example (z : ℤ) (h1 : z + 1 = 2) (h2 : z + 2 = 2) : (1 : ℤ) = 2 := From 74ff61b5fcabf81a110f84b6bdee637d7cc85b59 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Thu, 28 Jul 2022 20:15:19 +0000 Subject: [PATCH 131/173] feat(probability/strong_law): Lp version of the strong law of large numbers (#15392) This PR proves the Lp version of the strong law of large numbers which states that $\frac{1}{n}\sum_{i < n} X_i$ converges to $\mathbb{E}[X_0]$ in the Lp-norm where $(X_n)$ is iid. and Lp. --- .../function/uniform_integrable.lean | 56 ++++++++++++++++- src/probability/ident_distrib.lean | 61 +++++++++++++++++++ src/probability/strong_law.lean | 36 +++++++++++ 3 files changed, 152 insertions(+), 1 deletion(-) diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index 865face4a0deb..7288f0efe7dff 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -42,7 +42,7 @@ uniform integrable, uniformly absolutely continuous integral, Vitali convergence -/ noncomputable theory -open_locale classical measure_theory nnreal ennreal topological_space +open_locale classical measure_theory nnreal ennreal topological_space big_operators namespace measure_theory @@ -943,6 +943,60 @@ lemma uniform_integrable_iff [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ⟨λ h, ⟨h.1, λ ε, h.spec (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm hp'⟩, λ h, uniform_integrable_of hp hp' h.1 h.2⟩ + +/-- The averaging of a uniformly integrable sequence is also uniformly integrable. -/ +lemma uniform_integrable_average (hp : 1 ≤ p) + {f : ℕ → α → ℝ} (hf : uniform_integrable f p μ) : + uniform_integrable (λ n, (∑ i in finset.range n, f i) / n) p μ := +begin + obtain ⟨hf₁, hf₂, hf₃⟩ := hf, + refine ⟨λ n, _, λ ε hε, _, _⟩, + { simp_rw div_eq_mul_inv, + exact (finset.ae_strongly_measurable_sum' _ (λ i _, hf₁ i)).mul + (ae_strongly_measurable_const : ae_strongly_measurable (λ x, (↑n : ℝ)⁻¹) μ) }, + { obtain ⟨δ, hδ₁, hδ₂⟩ := hf₂ hε, + refine ⟨δ, hδ₁, λ n s hs hle, _⟩, + simp_rw [div_eq_mul_inv, finset.sum_mul, set.indicator_finset_sum], + refine le_trans (snorm_sum_le (λ i hi, ((hf₁ i).mul_const (↑n)⁻¹).indicator hs) hp) _, + have : ∀ i, s.indicator (f i * (↑n)⁻¹) = (↑n : ℝ)⁻¹ • s.indicator (f i), + { intro i, + rw [mul_comm, (_ : (↑n)⁻¹ * f i = λ ω, (↑n : ℝ)⁻¹ • f i ω)], + { rw set.indicator_const_smul s (↑n)⁻¹ (f i), + refl }, + { refl } }, + simp_rw [this, snorm_const_smul, ← finset.mul_sum, nnnorm_inv, real.nnnorm_coe_nat], + by_cases hn : (↑(↑n : ℝ≥0)⁻¹ : ℝ≥0∞) = 0, + { simp only [hn, zero_mul, zero_le] }, + refine le_trans _ (_ : ↑(↑n : ℝ≥0)⁻¹ * (n • ennreal.of_real ε) ≤ ennreal.of_real ε), + { refine (ennreal.mul_le_mul_left hn ennreal.coe_ne_top).2 _, + conv_rhs { rw ← finset.card_range n }, + exact finset.sum_le_card_nsmul _ _ _ (λ i hi, hδ₂ _ _ hs hle) }, + { simp only [ennreal.coe_eq_zero, inv_eq_zero, nat.cast_eq_zero] at hn, + rw [nsmul_eq_mul, ← mul_assoc, ennreal.coe_inv, ennreal.coe_nat, + ennreal.inv_mul_cancel _ ennreal.coe_nat_ne_top, one_mul], + { exact le_rfl }, + all_goals { simpa only [ne.def, nat.cast_eq_zero] } } }, + { obtain ⟨C, hC⟩ := hf₃, + simp_rw [div_eq_mul_inv, finset.sum_mul], + refine ⟨C, λ n, (snorm_sum_le (λ i hi, (hf₁ i).mul_const (↑n)⁻¹) hp).trans _⟩, + have : ∀ i, (λ ω, f i ω * (↑n)⁻¹) = (↑n : ℝ)⁻¹ • λ ω, f i ω, + { intro i, + ext ω, + simp only [mul_comm, pi.smul_apply, algebra.id.smul_eq_mul] }, + simp_rw [this, snorm_const_smul, ← finset.mul_sum, nnnorm_inv, real.nnnorm_coe_nat], + by_cases hn : (↑(↑n : ℝ≥0)⁻¹ : ℝ≥0∞) = 0, + { simp only [hn, zero_mul, zero_le] }, + refine le_trans _ (_ : ↑(↑n : ℝ≥0)⁻¹ * (n • C : ℝ≥0∞) ≤ C), + { refine (ennreal.mul_le_mul_left hn ennreal.coe_ne_top).2 _, + conv_rhs { rw ← finset.card_range n }, + exact finset.sum_le_card_nsmul _ _ _ (λ i hi, hC i) }, + { simp only [ennreal.coe_eq_zero, inv_eq_zero, nat.cast_eq_zero] at hn, + rw [nsmul_eq_mul, ← mul_assoc, ennreal.coe_inv, ennreal.coe_nat, + ennreal.inv_mul_cancel _ ennreal.coe_nat_ne_top, one_mul], + { exact le_rfl }, + all_goals { simpa only [ne.def, nat.cast_eq_zero] } } } +end + end uniform_integrable end measure_theory diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index b86122e18c847..e397568c6c6f4 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import probability.variance +import measure_theory.function.uniform_integrable /-! # Identically distributed random variables @@ -107,6 +108,12 @@ protected lemma comp {u : γ → δ} (h : ident_distrib f g μ ν) (hu : measura ident_distrib (u ∘ f) (u ∘ g) μ ν := h.comp_of_ae_measurable hu.ae_measurable +protected lemma of_ae_eq {g : α → γ} (hf : ae_measurable f μ) (heq : f =ᵐ[μ] g) : + ident_distrib f g μ μ := +{ ae_measurable_fst := hf, + ae_measurable_snd := hf.congr heq, + map_eq := measure.map_congr heq } + lemma measure_mem_eq (h : ident_distrib f g μ ν) {s : set γ} (hs : measurable_set s) : μ (f ⁻¹' s) = ν (g ⁻¹' s) := by rw [← measure.map_apply_of_ae_measurable h.ae_measurable_fst hs, @@ -283,4 +290,58 @@ end end ident_distrib +section uniform_integrable + +open topological_space + +variables {E : Type*} [measurable_space E] [normed_add_comm_group E] [borel_space E] + [second_countable_topology E] {μ : measure α} [is_finite_measure μ] + +/-- This lemma is superceded by `mem_ℒp.uniform_integrable_of_ident_distrib` which only require +`ae_strongly_measurable`. -/ +lemma mem_ℒp.uniform_integrable_of_ident_distrib_aux {ι : Type*} {f : ι → α → E} + {j : ι} {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) + (hℒp : mem_ℒp (f j) p μ) (hfmeas : ∀ i, strongly_measurable (f i)) + (hf : ∀ i, ident_distrib (f i) (f j) μ μ) : + uniform_integrable f p μ := +begin + refine uniform_integrable_of' hp hp' hfmeas (λ ε hε, _), + by_cases hι : nonempty ι, + swap, { exact ⟨0, λ i, false.elim (hι $ nonempty.intro i)⟩ }, + obtain ⟨C, hC₁, hC₂⟩ := hℒp.snorm_indicator_norm_ge_pos_le μ (hfmeas _) hε, + have hmeas : ∀ i, measurable_set {x | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥f i x∥₊} := + λ i, measurable_set_le measurable_const (hfmeas _).measurable.nnnorm, + refine ⟨⟨C, hC₁.le⟩, λ i, le_trans (le_of_eq _) hC₂⟩, + have : {x : α | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥f i x∥₊}.indicator (f i) = + (λ x : E, if (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥x∥₊ then x else 0) ∘ (f i), + { ext x, + simp only [set.indicator, set.mem_set_of_eq] }, + simp_rw [coe_nnnorm, this], + rw [← snorm_map_measure _ (hf i).ae_measurable_fst, (hf i).map_eq, + snorm_map_measure _ (hf j).ae_measurable_fst], + { refl }, + all_goals { exact ae_strongly_measurable_id.indicator + (measurable_set_le measurable_const measurable_nnnorm) }, +end + +/-- A sequence of identically distributed Lᵖ functions is p-uniformly integrable. -/ +lemma mem_ℒp.uniform_integrable_of_ident_distrib {ι : Type*} {f : ι → α → E} + {j : ι} {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) + (hℒp : mem_ℒp (f j) p μ) (hf : ∀ i, ident_distrib (f i) (f j) μ μ) : + uniform_integrable f p μ := +begin + have hfmeas : ∀ i, ae_strongly_measurable (f i) μ := + λ i, (hf i).ae_strongly_measurable_iff.2 hℒp.1, + set g : ι → α → E := λ i, (hfmeas i).some, + have hgmeas : ∀ i, strongly_measurable (g i) := λ i, (Exists.some_spec $ hfmeas i).1, + have hgeq : ∀ i, g i =ᵐ[μ] f i := λ i, (Exists.some_spec $ hfmeas i).2.symm, + have hgℒp : mem_ℒp (g j) p μ := hℒp.ae_eq (hgeq j).symm, + exact uniform_integrable.ae_eq (mem_ℒp.uniform_integrable_of_ident_distrib_aux hp hp' + hgℒp hgmeas $ + λ i, (ident_distrib.of_ae_eq (hgmeas i).ae_measurable (hgeq i)).trans ((hf i).trans + $ ident_distrib.of_ae_eq (hfmeas j).ae_measurable (hgeq j).symm)) hgeq, +end + +end uniform_integrable + end probability_theory diff --git a/src/probability/strong_law.lean b/src/probability/strong_law.lean index 441688403a213..ecdad27ce878c 100644 --- a/src/probability/strong_law.lean +++ b/src/probability/strong_law.lean @@ -19,6 +19,10 @@ If `X n` is a sequence of independent identically distributed integrable real-va variables, then `∑ i in range n, X i / n` converges almost surely to `𝔼[X 0]`. We give here the strong version, due to Etemadi, that only requires pairwise independence. +This file also contains the Lᵖ version of the strong law of large numbers provided by +`probability_theory.strong_law_Lp` which shows `∑ i in range n, X i / n` converges in Lᵖ to +`𝔼[X 0]` provided `X n` is independent identically distributed and is Lᵖ. + ## Implementation We follow the proof by Etemadi @@ -728,4 +732,36 @@ end end strong_law_ae +section strong_law_Lp + +variables {Ω : Type*} [measure_space Ω] [is_probability_measure (ℙ : measure Ω)] + +/-- *Strong law of large numbers*, Lᵖ version: if `X n` is a sequence of independent +identically distributed real-valued random variables in Lᵖ, then `∑ i in range n, X i / n` +converges in Lᵖ to `𝔼[X 0]`. -/ +theorem strong_law_Lp + {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) + (X : ℕ → Ω → ℝ) (hℒp : mem_ℒp (X 0) p) + (hindep : pairwise (λ i j, indep_fun (X i) (X j))) + (hident : ∀ i, ident_distrib (X i) (X 0)) : + tendsto (λ n, snorm (λ ω, (∑ i in range n, X i ω) / n - 𝔼[X 0]) p ℙ) at_top (𝓝 0) := +begin + have hmeas : ∀ i, ae_strongly_measurable (X i) ℙ := + λ i, (hident i).ae_strongly_measurable_iff.2 hℒp.1, + have hint : integrable (X 0) ℙ := hℒp.integrable hp, + have havg : ∀ n, ae_strongly_measurable (λ ω, (∑ i in range n, X i ω) / n) ℙ, + { intro n, + simp_rw div_eq_mul_inv, + exact ae_strongly_measurable.mul_const (ae_strongly_measurable_sum _ (λ i _, hmeas i)) _ }, + refine tendsto_Lp_of_tendsto_in_measure _ hp hp' havg (mem_ℒp_const _) _ + (tendsto_in_measure_of_tendsto_ae havg (strong_law_ae _ hint hindep hident)), + rw (_ : (λ n ω, (∑ i in range n, X i ω) / ↑n) = λ n, (∑ i in range n, X i) / ↑n), + { exact (uniform_integrable_average hp $ + mem_ℒp.uniform_integrable_of_ident_distrib hp hp' hℒp hident).2.1 }, + { ext n ω, + simp only [pi.coe_nat, pi.div_apply, sum_apply] } +end + +end strong_law_Lp + end probability_theory From d6b7aed6807c5498e43394651ff1d4917ed79974 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 28 Jul 2022 21:09:51 +0000 Subject: [PATCH 132/173] feat(analysis/normed_space/compact_operator): definition and basic facts about compact operators (#15467) --- .../normed_space/compact_operator.lean | 430 ++++++++++++++++++ 1 file changed, 430 insertions(+) create mode 100644 src/analysis/normed_space/compact_operator.lean diff --git a/src/analysis/normed_space/compact_operator.lean b/src/analysis/normed_space/compact_operator.lean new file mode 100644 index 0000000000000..6b9d7a865dd12 --- /dev/null +++ b/src/analysis/normed_space/compact_operator.lean @@ -0,0 +1,430 @@ +/- +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 + +/-! +# Compact operators + +In this file we define compact linear operators between two topological vector spaces (TVS). + +## Main definitions + +* `is_compact_operator` : predicate for compact operators + +## Main statements + +* `is_compact_operator_iff_is_compact_closure_image_closed_ball` : the usual characterization of + compact operators from a normed space to a T2 TVS. +* `is_compact_operator.comp_clm` : precomposing a compact operator by a continuous linear map gives + a compact operator +* `is_compact_operator.clm_comp` : postcomposing a compact operator by a continuous linear map + gives a compact operator +* `is_compact_operator.continuous` : compact operators are automatically continuous +* `is_closed_set_of_is_compact_operator` : the set of compact operators is closed for the operator + norm + +## Implementation details + +We define `is_compact_operator` as a predicate, because the space of compact operators inherits all +of its structure from the space of continuous linear maps (e.g we want to have the usual operator +norm on compact operators). + +The two natural options then would be to make it a predicate over linear maps or continuous linear +maps. Instead we define it as a predicate over bare functions, although it really only makes sense +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) + +## Tags + +Compact operator +-/ + +open function set filter bornology metric + +open_locale pointwise big_operators topological_space + +/-- A compact operator between two topological vector spaces. This definition is usually +given as "there exists a neighborhood of zero whose image is contained in a compact set", +but we choose a definition which involves fewer existential quantifiers and replaces images +with preimages. + +We prove the equivalence in `is_compact_operator_iff_exists_mem_nhds_image_subset_compact`. -/ +def is_compact_operator {M₁ M₂ : Type*} [has_zero M₁] [topological_space M₁] + [topological_space M₂] (f : M₁ → M₂) : Prop := +∃ K, is_compact K ∧ f ⁻¹' K ∈ (𝓝 0 : filter M₁) + +lemma is_compact_operator_zero {M₁ M₂ : Type*} [has_zero M₁] [topological_space M₁] + [topological_space M₂] [has_zero M₂] : is_compact_operator (0 : M₁ → M₂) := +⟨{0}, is_compact_singleton, mem_of_superset univ_mem (λ x _, rfl)⟩ + +section characterizations + +section + +variables {R₁ R₂ : Type*} [semiring R₁] [semiring R₂] {σ₁₂ : R₁ →+* R₂} {M₁ M₂ : Type*} + [topological_space M₁] [add_comm_monoid M₁] [topological_space M₂] + +lemma is_compact_operator_iff_exists_mem_nhds_image_subset_compact (f : M₁ → M₂) : + is_compact_operator f ↔ ∃ V ∈ (𝓝 0 : filter M₁), ∃ (K : set M₂), is_compact K ∧ f '' V ⊆ K := +⟨λ ⟨K, hK, hKf⟩, ⟨f ⁻¹' K, hKf, K, hK, image_preimage_subset _ _⟩, + λ ⟨V, hV, K, hK, hVK⟩, ⟨K, hK, mem_of_superset hV (image_subset_iff.mp hVK)⟩⟩ + +lemma is_compact_operator_iff_exists_mem_nhds_is_compact_closure_image [t2_space M₂] + (f : M₁ → M₂) : + is_compact_operator f ↔ ∃ V ∈ (𝓝 0 : filter M₁), is_compact (closure $ f '' V) := +begin + rw is_compact_operator_iff_exists_mem_nhds_image_subset_compact, + exact ⟨λ ⟨V, hV, K, hK, hKV⟩, ⟨V, hV, compact_closure_of_subset_compact hK hKV⟩, + λ ⟨V, hV, hVc⟩, ⟨V, hV, closure (f '' V), hVc, subset_closure⟩⟩, +end + +end + +section bounded + +variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [semi_normed_ring 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} + {M₁ M₂ : Type*} [topological_space M₁] [add_comm_monoid M₁] [topological_space M₂] + [add_comm_monoid M₂] [module 𝕜₁ M₁] [module 𝕜₂ M₂] [has_continuous_const_smul 𝕜₂ M₂] + +lemma is_compact_operator.image_subset_compact_of_vonN_bounded {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) {S : set M₁} (hS : is_vonN_bounded 𝕜₁ S) : + ∃ (K : set M₂), is_compact K ∧ f '' S ⊆ K := +let ⟨K, hK, hKf⟩ := hf, + ⟨r, hr, hrS⟩ := hS hKf, + ⟨c, hc⟩ := normed_field.exists_lt_norm 𝕜₁ r, + this := ne_zero_of_norm_ne_zero (hr.trans hc).ne.symm in +⟨σ₁₂ c • K, hK.image $ continuous_id.const_smul (σ₁₂ c), + by rw [image_subset_iff, preimage_smul_setₛₗ _ _ _ f this.is_unit]; exact hrS c hc.le⟩ + +lemma is_compact_operator.is_compact_closure_image_of_vonN_bounded [t2_space M₂] + {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) {S : set M₁} + (hS : is_vonN_bounded 𝕜₁ S) : is_compact (closure $ f '' S) := +let ⟨K, hK, hKf⟩ := hf.image_subset_compact_of_vonN_bounded hS in +compact_closure_of_subset_compact hK hKf + +end bounded + +section normed_space + +variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [semi_normed_ring 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} + {M₁ M₂ M₃ : Type*} [seminormed_add_comm_group M₁] [topological_space M₂] + [add_comm_monoid M₂] [normed_space 𝕜₁ M₁] [module 𝕜₂ M₂] + +lemma is_compact_operator.image_subset_compact_of_bounded [has_continuous_const_smul 𝕜₂ M₂] + {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) {S : set M₁} (hS : metric.bounded S) : + ∃ (K : set M₂), is_compact K ∧ f '' S ⊆ K := +hf.image_subset_compact_of_vonN_bounded +(by rwa [normed_space.is_vonN_bounded_iff, ← metric.bounded_iff_is_bounded]) + +lemma is_compact_operator.is_compact_closure_image_of_bounded [has_continuous_const_smul 𝕜₂ M₂] + [t2_space M₂] {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) {S : set M₁} + (hS : metric.bounded S) : is_compact (closure $ f '' S) := +hf.is_compact_closure_image_of_vonN_bounded +(by rwa [normed_space.is_vonN_bounded_iff, ← metric.bounded_iff_is_bounded]) + +lemma is_compact_operator.image_ball_subset_compact [has_continuous_const_smul 𝕜₂ M₂] + {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) (r : ℝ) : + ∃ (K : set M₂), is_compact K ∧ f '' metric.ball 0 r ⊆ K := +hf.image_subset_compact_of_vonN_bounded (normed_space.is_vonN_bounded_ball 𝕜₁ M₁ r) + +lemma is_compact_operator.image_closed_ball_subset_compact [has_continuous_const_smul 𝕜₂ M₂] + {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) (r : ℝ) : + ∃ (K : set M₂), is_compact K ∧ f '' metric.closed_ball 0 r ⊆ K := +hf.image_subset_compact_of_vonN_bounded (normed_space.is_vonN_bounded_closed_ball 𝕜₁ M₁ r) + +lemma is_compact_operator.is_compact_closure_image_ball [has_continuous_const_smul 𝕜₂ M₂] + [t2_space M₂] {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) (r : ℝ) : + is_compact (closure $ f '' metric.ball 0 r) := +hf.is_compact_closure_image_of_vonN_bounded (normed_space.is_vonN_bounded_ball 𝕜₁ M₁ r) + +lemma is_compact_operator.is_compact_closure_image_closed_ball [has_continuous_const_smul 𝕜₂ M₂] + [t2_space M₂] {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : is_compact_operator f) (r : ℝ) : + is_compact (closure $ f '' metric.closed_ball 0 r) := +hf.is_compact_closure_image_of_vonN_bounded (normed_space.is_vonN_bounded_closed_ball 𝕜₁ M₁ r) + +lemma is_compact_operator_iff_image_ball_subset_compact [has_continuous_const_smul 𝕜₂ M₂] + (f : M₁ →ₛₗ[σ₁₂] M₂) {r : ℝ} (hr : 0 < r) : + is_compact_operator f ↔ ∃ (K : set M₂), is_compact K ∧ f '' metric.ball 0 r ⊆ K := +⟨λ hf, hf.image_ball_subset_compact r, + λ ⟨K, hK, hKr⟩, (is_compact_operator_iff_exists_mem_nhds_image_subset_compact f).mpr + ⟨metric.ball 0 r, ball_mem_nhds _ hr, K, hK, hKr⟩⟩ + +lemma is_compact_operator_iff_image_closed_ball_subset_compact [has_continuous_const_smul 𝕜₂ M₂] + (f : M₁ →ₛₗ[σ₁₂] M₂) {r : ℝ} (hr : 0 < r) : + is_compact_operator f ↔ ∃ (K : set M₂), is_compact K ∧ f '' metric.closed_ball 0 r ⊆ K := +⟨λ hf, hf.image_closed_ball_subset_compact r, + λ ⟨K, hK, hKr⟩, (is_compact_operator_iff_exists_mem_nhds_image_subset_compact f).mpr + ⟨metric.closed_ball 0 r, closed_ball_mem_nhds _ hr, K, hK, hKr⟩⟩ + +lemma is_compact_operator_iff_is_compact_closure_image_ball [has_continuous_const_smul 𝕜₂ M₂] + [t2_space M₂] (f : M₁ →ₛₗ[σ₁₂] M₂) {r : ℝ} (hr : 0 < r) : + is_compact_operator f ↔ is_compact (closure $ f '' metric.ball 0 r) := +⟨λ hf, hf.is_compact_closure_image_ball r, + λ hf, (is_compact_operator_iff_exists_mem_nhds_is_compact_closure_image f).mpr + ⟨metric.ball 0 r, ball_mem_nhds _ hr, hf⟩⟩ + +lemma is_compact_operator_iff_is_compact_closure_image_closed_ball + [has_continuous_const_smul 𝕜₂ M₂] [t2_space M₂] (f : M₁ →ₛₗ[σ₁₂] M₂) {r : ℝ} (hr : 0 < r) : + is_compact_operator f ↔ is_compact (closure $ f '' metric.closed_ball 0 r) := +⟨λ hf, hf.is_compact_closure_image_closed_ball r, + λ hf, (is_compact_operator_iff_exists_mem_nhds_is_compact_closure_image f).mpr + ⟨metric.closed_ball 0 r, closed_ball_mem_nhds _ hr, hf⟩⟩ + +end normed_space + +end characterizations + +section operations + +variables {R₁ R₂ R₃ R₄ : Type*} [semiring R₁] [semiring R₂] [comm_semiring R₃] [comm_semiring R₄] + {σ₁₂ : R₁ →+* R₂} {σ₁₄ : R₁ →+* R₄} {σ₃₄ : R₃ →+* R₄} {M₁ M₂ M₃ M₄ : Type*} + [topological_space M₁] [add_comm_monoid M₁] [topological_space M₂] [add_comm_monoid M₂] + [topological_space M₃] [add_comm_group M₃] [topological_space M₄] [add_comm_group M₄] + +lemma is_compact_operator.smul {S : Type*} [monoid S] [distrib_mul_action S M₂] + [has_continuous_const_smul S M₂] {f : M₁ → M₂} + (hf : is_compact_operator f) (c : S) : + is_compact_operator (c • f) := +let ⟨K, hK, hKf⟩ := hf in ⟨c • K, hK.image $ continuous_id.const_smul c, + mem_of_superset hKf (λ x hx, smul_mem_smul_set hx)⟩ + +lemma is_compact_operator.add [has_continuous_add M₂] {f g : M₁ → M₂} + (hf : is_compact_operator f) (hg : is_compact_operator g) : + is_compact_operator (f + g) := +let ⟨A, hA, hAf⟩ := hf, ⟨B, hB, hBg⟩ := hg in +⟨A + B, hA.add hB, mem_of_superset (inter_mem hAf hBg) (λ x ⟨hxA, hxB⟩, set.add_mem_add hxA hxB)⟩ + +lemma is_compact_operator.neg [has_continuous_neg M₄] {f : M₁ → M₄} + (hf : is_compact_operator f) : is_compact_operator (-f) := +let ⟨K, hK, hKf⟩ := hf in +⟨-K, hK.neg, mem_of_superset hKf $ λ x (hx : f x ∈ K), set.neg_mem_neg.mpr hx⟩ + +lemma is_compact_operator.sub [topological_add_group M₄] {f g : M₁ → M₄} + (hf : is_compact_operator f) (hg : is_compact_operator g) : is_compact_operator (f - g) := +by rw sub_eq_add_neg; exact hf.add hg.neg + +variables (σ₁₄ M₁ M₄) + +/-- The submodule of compact continuous linear maps. -/ +def compact_operator [module R₁ M₁] [module R₄ M₄] [has_continuous_const_smul R₄ M₄] + [topological_add_group M₄] : + submodule R₄ (M₁ →SL[σ₁₄] M₄) := +{ carrier := {f | is_compact_operator f}, + add_mem' := λ f g hf hg, hf.add hg, + zero_mem' := is_compact_operator_zero, + smul_mem' := λ c f hf, hf.smul c } + +end operations + +section comp + +variables {R₁ R₂ R₃ : Type*} [semiring R₁] [semiring R₂] [semiring R₃] {σ₁₂ : R₁ →+* R₂} + {σ₂₃ : R₂ →+* R₃} {M₁ M₂ M₃ : Type*} [topological_space M₁] [topological_space M₂] + [topological_space M₃] [add_comm_monoid M₁] [module R₁ M₁] + +lemma is_compact_operator.comp_clm [add_comm_monoid M₂] [module R₂ M₂] {f : M₂ → M₃} + (hf : is_compact_operator f) (g : M₁ →SL[σ₁₂] M₂) : + is_compact_operator (f ∘ g) := +begin + have := g.continuous.tendsto 0, + rw map_zero at this, + rcases hf with ⟨K, hK, hKf⟩, + exact ⟨K, hK, this hKf⟩ +end + +lemma is_compact_operator.continuous_comp {f : M₁ → M₂} (hf : is_compact_operator f) {g : M₂ → M₃} + (hg : continuous g) : + is_compact_operator (g ∘ f) := +begin + rcases hf with ⟨K, hK, hKf⟩, + refine ⟨g '' K, hK.image hg, mem_of_superset hKf _⟩, + nth_rewrite 1 preimage_comp, + exact preimage_mono (subset_preimage_image _ _) +end + +lemma is_compact_operator.clm_comp [add_comm_monoid M₂] [module R₂ M₂] [add_comm_monoid M₃] + [module R₃ M₃] {f : M₁ → M₂} (hf : is_compact_operator f) (g : M₂ →SL[σ₂₃] M₃) : + is_compact_operator (g ∘ f) := +hf.continuous_comp g.continuous + +end comp + +section cod_restrict + +variables {R₁ R₂ : Type*} [semiring R₁] [semiring R₂] {σ₁₂ : R₁ →+* R₂} + {M₁ M₂ : Type*} [topological_space M₁] [topological_space M₂] + [add_comm_monoid M₁] [add_comm_monoid M₂] [module R₁ M₁] [module R₂ M₂] + +lemma is_compact_operator.cod_restrict {f : M₁ → M₂} (hf : is_compact_operator f) + {V : submodule R₂ M₂} (hV : ∀ x, f x ∈ V) (h_closed : is_closed (V : set M₂)): + is_compact_operator (set.cod_restrict f V hV) := +let ⟨K, hK, hKf⟩ := hf in +⟨coe ⁻¹' K, (closed_embedding_subtype_coe h_closed).is_compact_preimage hK, hKf⟩ + +end cod_restrict + +section restrict + +variables {R₁ R₂ R₃ : Type*} [semiring R₁] [semiring R₂] [semiring R₃] {σ₁₂ : R₁ →+* R₂} + {σ₂₃ : R₂ →+* R₃} {M₁ M₂ M₃ : Type*} [topological_space M₁] [uniform_space M₂] + [topological_space M₃] [add_comm_monoid M₁] [add_comm_monoid M₂] [add_comm_monoid M₃] + [module R₁ M₁] [module R₂ M₂] [module R₃ M₃] + +/-- If a compact operator preserves a closed submodule, its restriction to that submodule is +compact. + +Note that, following mathlib's convention in linear algebra, `restrict` designates the restriction +of an endomorphism `f : E →ₗ E` to an endomorphism `f' : ↥V →ₗ ↥V`. To prove that the restriction +`f' : ↥U →ₛₗ ↥V` of a compact operator `f : E →ₛₗ F` is compact, apply +`is_compact_operator.cod_restrict` to `f ∘ U.subtypeL`, which is compact by +`is_compact_operator.comp_clm`. -/ +lemma is_compact_operator.restrict {f : M₁ →ₗ[R₁] M₁} (hf : is_compact_operator f) + {V : submodule R₁ M₁} (hV : ∀ v ∈ V, f v ∈ V) (h_closed : is_closed (V : set M₁)): + is_compact_operator (f.restrict hV) := +(hf.comp_clm V.subtypeL).cod_restrict (set_like.forall.2 hV) h_closed + +/-- If a compact operator preserves a complete submodule, its restriction to that submodule is +compact. + +Note that, following mathlib's convention in linear algebra, `restrict` designates the restriction +of an endomorphism `f : E →ₗ E` to an endomorphism `f' : ↥V →ₗ ↥V`. To prove that the restriction +`f' : ↥U →ₛₗ ↥V` of a compact operator `f : E →ₛₗ F` is compact, apply +`is_compact_operator.cod_restrict` to `f ∘ U.subtypeL`, which is compact by +`is_compact_operator.comp_clm`. -/ +lemma is_compact_operator.restrict' [separated_space M₂] {f : M₂ →ₗ[R₂] M₂} + (hf : is_compact_operator f) {V : submodule R₂ M₂} (hV : ∀ v ∈ V, f v ∈ V) + [hcomplete : complete_space V] : is_compact_operator (f.restrict hV) := +hf.restrict hV (complete_space_coe_iff_is_complete.mp hcomplete).is_closed + +end restrict + +section continuous + +variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [nontrivially_normed_field 𝕜₂] + {σ₁₂ : 𝕜₁ →+* 𝕜₂} [ring_hom_isometric σ₁₂] {M₁ M₂ : Type*} [topological_space M₁] + [add_comm_group M₁] [topological_space M₂] [add_comm_group M₂] [module 𝕜₁ M₁] [module 𝕜₂ M₂] + [topological_add_group M₁] [has_continuous_const_smul 𝕜₁ M₁] + [topological_add_group M₂] [has_continuous_smul 𝕜₂ M₂] + +@[continuity] lemma is_compact_operator.continuous {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) : continuous f := +begin + letI : uniform_space M₂ := topological_add_group.to_uniform_space _, + haveI : uniform_add_group M₂ := topological_add_group_is_uniform, + -- Since `f` is linear, we only need to show that it is continuous at zero. + -- Let `U` be a neighborhood of `0` in `M₂`. + refine continuous_of_continuous_at_zero f (λ U hU, _), + rw map_zero at hU, + -- The compactness of `f` gives us a compact set `K : set M₂` such that `f ⁻¹' K` is a + -- neighborhood of `0` in `M₁`. + rcases hf with ⟨K, hK, hKf⟩, + -- But any compact set is totally bounded, hence Von-Neumann bounded. Thus, `K` absorbs `U`. + -- This gives `r > 0` such that `∀ a : 𝕜₂, r ≤ ∥a∥ → K ⊆ a • U`. + rcases hK.totally_bounded.is_vonN_bounded 𝕜₂ hU with ⟨r, hr, hrU⟩, + -- Choose `c : 𝕜₂` with `r < ∥c∥`. + rcases normed_field.exists_lt_norm 𝕜₁ r with ⟨c, hc⟩, + have hcnz : c ≠ 0 := ne_zero_of_norm_ne_zero (hr.trans hc).ne.symm, + -- We have `f ⁻¹' ((σ₁₂ c⁻¹) • K) = c⁻¹ • f ⁻¹' K ∈ 𝓝 0`. Thus, showing that + -- `(σ₁₂ c⁻¹) • K ⊆ U` is enough to deduce that `f ⁻¹' U ∈ 𝓝 0`. + suffices : (σ₁₂ $ c⁻¹) • K ⊆ U, + { refine mem_of_superset _ this, + have : is_unit c⁻¹ := hcnz.is_unit.inv, + rwa [mem_map, preimage_smul_setₛₗ _ _ _ f this, set_smul_mem_nhds_zero_iff (inv_ne_zero hcnz)], + apply_instance }, + -- Since `σ₁₂ c⁻¹` = `(σ₁₂ c)⁻¹`, we have to prove that `K ⊆ σ₁₂ c • U`. + rw [σ₁₂.map_inv, ← subset_set_smul_iff₀ (σ₁₂.map_ne_zero.mpr hcnz)], + -- But `σ₁₂` is isometric, so `∥σ₁₂ c∥ = ∥c∥ > r`, which concludes the argument since + -- `∀ a : 𝕜₂, r ≤ ∥a∥ → K ⊆ a • U`. + refine hrU (σ₁₂ c) _, + rw ring_hom_isometric.is_iso, + exact hc.le +end + +/-- Upgrade a compact `linear_map` to a `continuous_linear_map`. -/ +def continuous_linear_map.mk_of_is_compact_operator {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) : M₁ →SL[σ₁₂] M₂ := +⟨f, hf.continuous⟩ + +@[simp] lemma continuous_linear_map.mk_of_is_compact_operator_to_linear_map {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) : + (continuous_linear_map.mk_of_is_compact_operator hf : M₁ →ₛₗ[σ₁₂] M₂) = f := +rfl + +@[simp] lemma continuous_linear_map.coe_mk_of_is_compact_operator {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) : + (continuous_linear_map.mk_of_is_compact_operator hf : M₁ → M₂) = f := +rfl + +lemma continuous_linear_map.mk_of_is_compact_operator_mem_compact_operator {f : M₁ →ₛₗ[σ₁₂] M₂} + (hf : is_compact_operator f) : + continuous_linear_map.mk_of_is_compact_operator hf ∈ compact_operator σ₁₂ M₁ M₂ := +hf + +end continuous + +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₂] : + 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, + 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 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⟩, + have hTv : v '' closed_ball 0 1 ⊆ _ := subset_closure.trans hTv, + refine ⟨T, hT, _⟩, + rw image_subset_iff at ⊢ hTv, + intros x hx, + specialize hTv hx, + rw [mem_preimage, 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 +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₂] : + (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)) : + is_compact_operator f := +is_closed_set_of_is_compact_operator.mem_of_tendsto hf hF From 049062a6187583624a33ea5bf2d32ce0b14f4802 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 29 Jul 2022 02:07:00 +0000 Subject: [PATCH 133/173] =?UTF-8?q?refactor(linear=5Falgebra,=20analysis/i?= =?UTF-8?q?nner=5Fproduct=5Fspace):=20use=20`=E2=8A=A4=20=E2=89=A4`=20inst?= =?UTF-8?q?ead=20of=20`=3D=20=E2=8A=A4`=20in=20bases=20constructors=20(#15?= =?UTF-8?q?697)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit All the existing proof just need a `.ge` to be fixed, and this allows to remove a lot of `rw [eq_top_iff]`, and it sticks with the convention of having weakest forms in assumptions and stronger forms in conclusion. Zulip thread: https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60.E2.8A.A4.20.E2.89.A4.20span.60.20in.20basis.20constructors --- src/analysis/fourier.lean | 2 +- .../gram_schmidt_ortho.lean | 4 +- .../inner_product_space/l2_space.lean | 23 +++++----- src/analysis/inner_product_space/pi_L2.lean | 7 +-- .../inner_product_space/projection.lean | 2 +- src/linear_algebra/affine_space/basis.lean | 2 +- src/linear_algebra/basis.lean | 21 ++++----- src/linear_algebra/determinant.lean | 4 +- src/linear_algebra/finite_dimensional.lean | 45 +++++++++---------- src/number_theory/ramification_inertia.lean | 2 +- src/ring_theory/adjoin/power_basis.lean | 3 +- src/ring_theory/adjoin_root.lean | 3 +- src/ring_theory/discriminant.lean | 2 +- src/ring_theory/localization/module.lean | 2 +- 14 files changed, 58 insertions(+), 64 deletions(-) diff --git a/src/analysis/fourier.lean b/src/analysis/fourier.lean index afff0ce40880e..8c568d14f027d 100644 --- a/src/analysis/fourier.lean +++ b/src/analysis/fourier.lean @@ -215,7 +215,7 @@ section fourier /-- We define `fourier_series` to be a `ℤ`-indexed Hilbert basis for `Lp ℂ 2 haar_circle`, which by definition is an isometric isomorphism from `Lp ℂ 2 haar_circle` to `ℓ²(ℤ, ℂ)`. -/ def fourier_series : hilbert_basis ℤ ℂ (Lp ℂ 2 haar_circle) := -hilbert_basis.mk orthonormal_fourier (span_fourier_Lp_closure_eq_top (by norm_num)) +hilbert_basis.mk orthonormal_fourier (span_fourier_Lp_closure_eq_top (by norm_num)).ge /-- The elements of the Hilbert basis `fourier_series` for `Lp ℂ 2 haar_circle` are the functions `fourier_Lp 2`, the monomials `λ z, z ^ n` on the circle considered as elements of `L2`. -/ diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index 554ce5e0359fb..91957da899ecd 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -202,7 +202,7 @@ linear_independent_of_ne_zero_of_inner_eq_zero noncomputable def gram_schmidt_basis (b : basis ι 𝕜 E) : basis ι 𝕜 E := basis.mk (gram_schmidt_linear_independent 𝕜 b b.linear_independent) - ((span_gram_schmidt 𝕜 b).trans b.span_eq) + ((span_gram_schmidt 𝕜 b).trans b.span_eq).ge lemma coe_gram_schmidt_basis (b : basis ι 𝕜 E) : (gram_schmidt_basis 𝕜 b : ι → E) = gram_schmidt 𝕜 b := basis.coe_mk _ _ @@ -259,4 +259,4 @@ noncomputable def gram_schmidt_orthonormal_basis [fintype ι] (b : basis ι 𝕜 orthonormal_basis ι 𝕜 E := orthonormal_basis.mk (gram_schmidt_orthonormal 𝕜 b b.linear_independent) - (((span_gram_schmidt_normed_range 𝕜 b).trans (span_gram_schmidt 𝕜 b)).trans b.span_eq) + (((span_gram_schmidt_normed_range 𝕜 b).trans (span_gram_schmidt 𝕜 b)).trans b.span_eq).ge diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 11eb17b1f7b03..0b26b94f4c823 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -247,22 +247,20 @@ a isometric isomorphism from E to `lp 2` of the subspaces. Note that this goes in the opposite direction from `orthogonal_family.linear_isometry`. -/ noncomputable def linear_isometry_equiv [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) : E ≃ₗᵢ[𝕜] lp G 2 := linear_isometry_equiv.symm $ linear_isometry_equiv.of_surjective hV.linear_isometry begin rw [←linear_isometry.coe_to_linear_map], - refine linear_map.range_eq_top.mp _, - rw ← hV', - rw hV.range_linear_isometry, + exact linear_map.range_eq_top.mp (eq_top_iff.mpr $ hV'.trans_eq hV.range_linear_isometry.symm) end /-- In the canonical isometric isomorphism `E ≃ₗᵢ[𝕜] lp G 2` induced by an orthogonal family `G`, a vector `w : lp G 2` is the image of the infinite sum of the associated elements in `E`. -/ protected lemma linear_isometry_equiv_symm_apply [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) (w : lp G 2) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) (w : lp G 2) : (hV.linear_isometry_equiv hV').symm w = ∑' i, V i (w i) := by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.linear_isometry_apply] @@ -270,7 +268,7 @@ by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.linear_isome 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 has_sum_linear_isometry_equiv_symm [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) (w : lp G 2) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) (w : lp G 2) : has_sum (λ i, V i (w i)) ((hV.linear_isometry_equiv hV').symm w) := by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.has_sum_linear_isometry] @@ -278,7 +276,7 @@ by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.has_sum_line family `G`, an "elementary basis vector" in `lp G 2` supported at `i : ι` is the image of the associated element in `E`. -/ @[simp] protected lemma linear_isometry_equiv_symm_apply_single [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) {i : ι} (x : G i) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) {i : ι} (x : G i) : (hV.linear_isometry_equiv hV').symm (lp.single 2 i x) = V i x := by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.linear_isometry_apply_single] @@ -287,7 +285,7 @@ family `G`, a finitely-supported vector in `lp G 2` is the image of the associat elements of `E`. -/ @[simp] protected lemma linear_isometry_equiv_symm_apply_dfinsupp_sum_single [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) (W₀ : Π₀ (i : ι), G i) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) (W₀ : Π₀ (i : ι), G i) : (hV.linear_isometry_equiv hV').symm (W₀.sum (lp.single 2)) = (W₀.sum (λ i, V i)) := by simp [orthogonal_family.linear_isometry_equiv, orthogonal_family.linear_isometry_apply_dfinsupp_sum_single] @@ -297,7 +295,7 @@ family `G`, a finitely-supported vector in `lp G 2` is the image of the associat elements of `E`. -/ @[simp] protected lemma linear_isometry_equiv_apply_dfinsupp_sum_single [Π i, complete_space (G i)] - (hV' : (⨆ i, (V i).to_linear_map.range).topological_closure = ⊤) (W₀ : Π₀ (i : ι), G i) : + (hV' : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) (W₀ : Π₀ (i : ι), G i) : (hV.linear_isometry_equiv hV' (W₀.sum (λ i, V i)) : Π i, G i) = W₀ := begin rw ← hV.linear_isometry_equiv_symm_apply_dfinsupp_sum_single hV', @@ -411,6 +409,7 @@ protected def to_orthonormal_basis [fintype ι] (b : hilbert_basis ι 𝕜 E) : orthonormal_basis ι 𝕜 E := orthonormal_basis.mk b.orthonormal begin + refine eq.ge _, have := (span 𝕜 (finset.univ.image b : set E)).closed_of_finite_dimensional, simpa only [finset.coe_image, finset.coe_univ, set.image_univ, hilbert_basis.dense_span] using this.submodule_topological_closure_eq.symm @@ -424,7 +423,7 @@ variables {v : ι → E} (hv : orthonormal 𝕜 v) include hv cplt /-- An orthonormal family of vectors whose span is dense in the whole module is a Hilbert basis. -/ -protected def mk (hsp : (span 𝕜 (set.range v)).topological_closure = ⊤) : +protected def mk (hsp : ⊤ ≤ (span 𝕜 (set.range v)).topological_closure) : hilbert_basis ι 𝕜 E := hilbert_basis.of_repr $ hv.orthogonal_family.linear_isometry_equiv @@ -438,7 +437,7 @@ lemma _root_.orthonormal.linear_isometry_equiv_symm_apply_single_one (h i) : by rw [orthogonal_family.linear_isometry_equiv_symm_apply_single, linear_isometry.to_span_singleton_apply, one_smul] -@[simp] protected lemma coe_mk (hsp : (span 𝕜 (set.range v)).topological_closure = ⊤) : +@[simp] protected lemma coe_mk (hsp : ⊤ ≤ (span 𝕜 (set.range v)).topological_closure) : ⇑(hilbert_basis.mk hv hsp) = v := funext $ orthonormal.linear_isometry_equiv_symm_apply_single_one hv _ @@ -446,7 +445,7 @@ funext $ orthonormal.linear_isometry_equiv_symm_apply_single_one hv _ basis. -/ protected def mk_of_orthogonal_eq_bot (hsp : (span 𝕜 (set.range v))ᗮ = ⊥) : hilbert_basis ι 𝕜 E := hilbert_basis.mk hv -(by rw [← orthogonal_orthogonal_eq_closure, orthogonal_eq_top_iff, hsp]) +(by rw [← orthogonal_orthogonal_eq_closure, ← eq_top_iff, orthogonal_eq_top_iff, hsp]) @[simp] protected lemma coe_of_orthogonal_eq_bot_mk (hsp : (span 𝕜 (set.range v))ᗮ = ⊥) : ⇑(hilbert_basis.mk_of_orthogonal_eq_bot hv hsp) = v := diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 70f7499ce8496..559570c3b4a6c 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -386,12 +386,12 @@ calc (v.to_orthonormal_basis hv : ι → E) = ((v.to_orthonormal_basis hv).to_ba variable {v : ι → E} /-- A finite orthonormal set that spans is an orthonormal basis -/ -protected def mk (hon : orthonormal 𝕜 v) (hsp: submodule.span 𝕜 (set.range v) = ⊤): +protected def mk (hon : orthonormal 𝕜 v) (hsp: ⊤ ≤ submodule.span 𝕜 (set.range v)): orthonormal_basis ι 𝕜 E := (basis.mk (orthonormal.linear_independent hon) hsp).to_orthonormal_basis (by rwa basis.coe_mk) @[simp] -protected lemma coe_mk (hon : orthonormal 𝕜 v) (hsp: submodule.span 𝕜 (set.range v) = ⊤) : +protected lemma coe_mk (hon : orthonormal 𝕜 v) (hsp: ⊤ ≤ submodule.span 𝕜 (set.range v)) : ⇑(orthonormal_basis.mk hon hsp) = v := by classical; rw [orthonormal_basis.mk, _root_.basis.coe_to_orthonormal_basis, basis.coe_mk] @@ -405,7 +405,7 @@ let convert orthonormal_span (h.comp (coe : s → ι') subtype.coe_injective), ext, simp [e₀', basis.span_apply], - end e₀'.span_eq, + end e₀'.span_eq.ge, φ : span 𝕜 (s.image v' : set E) ≃ₗᵢ[𝕜] span 𝕜 (range (v' ∘ (coe : s → ι'))) := linear_isometry_equiv.of_eq _ _ begin @@ -429,6 +429,7 @@ protected def mk_of_orthogonal_eq_bot (hon : orthonormal 𝕜 v) (hsp : (span orthonormal_basis ι 𝕜 E := orthonormal_basis.mk hon begin + refine eq.ge _, haveI : finite_dimensional 𝕜 (span 𝕜 (range v)) := finite_dimensional.span_of_finite 𝕜 (finite_range v), haveI : complete_space (span 𝕜 (range v)) := finite_dimensional.complete 𝕜 _, diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 834e4c9914b92..c74cf11b4e0d1 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -1143,7 +1143,7 @@ begin have hv_coe : range (coe : v → E) = v := by simp, split, { refine λ h, ⟨basis.mk hv.linear_independent _, basis.coe_mk _ _⟩, - convert h }, + convert h.ge }, { rintros ⟨h, coe_h⟩, rw [← h.span_eq, coe_h, hv_coe] } end diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index f599bb429dfeb..f0aa891ea4546 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -76,7 +76,7 @@ basis.mk ((affine_independent_iff_linear_independent_vsub k b.points 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), - { rw [this, ← direction_affine_span, b.tot, affine_subspace.direction_top], }, + { 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), congr, diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 7305b4da13c7c..fda0c874741b2 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -898,7 +898,7 @@ end section mk -variables (hli : linear_independent R v) (hsp : span R (range v) = ⊤) +variables (hli : linear_independent R v) (hsp : ⊤ ≤ span R (range v)) /-- A linear independent family of vectors spanning the whole module is a basis. -/ protected noncomputable def mk : basis ι R M := @@ -906,10 +906,10 @@ basis.of_repr { inv_fun := finsupp.total _ _ _ v, left_inv := λ x, hli.total_repr ⟨x, _⟩, right_inv := λ x, hli.repr_eq rfl, - .. hli.repr.comp (linear_map.id.cod_restrict _ (λ h, hsp.symm ▸ submodule.mem_top)) } + .. hli.repr.comp (linear_map.id.cod_restrict _ (λ h, hsp submodule.mem_top)) } @[simp] lemma mk_repr : - (basis.mk hli hsp).repr x = hli.repr ⟨x, hsp.symm ▸ submodule.mem_top⟩ := + (basis.mk hli hsp).repr x = hli.repr ⟨x, hsp submodule.mem_top⟩ := rfl lemma mk_apply (i : ι) : basis.mk hli hsp i = v i := @@ -954,7 +954,6 @@ variables (hli : linear_independent R v) protected noncomputable def span : basis ι R (span R (range v)) := basis.mk (linear_independent_span hli) $ begin - rw eq_top_iff, intros x _, have h₁ : (coe : span R (range v) → M) '' set.range (λ i, subtype.mk (v i) _) = range v, { rw ← set.range_comp, @@ -996,13 +995,13 @@ def group_smul {G : Type*} [group G] [distrib_mul_action G R] [distrib_mul_actio [is_scalar_tower G R M] [smul_comm_class G R M] (v : basis ι R M) (w : ι → G) : basis ι R M := @basis.mk ι R M (w • v) _ _ _ - (v.linear_independent.group_smul w) (group_smul_span_eq_top v.span_eq) + (v.linear_independent.group_smul w) (group_smul_span_eq_top v.span_eq).ge lemma group_smul_apply {G : Type*} [group G] [distrib_mul_action G R] [distrib_mul_action G M] [is_scalar_tower G R M] [smul_comm_class G R M] {v : basis ι R M} {w : ι → G} (i : ι) : v.group_smul w i = (w • v : ι → M) i := mk_apply - (v.linear_independent.group_smul w) (group_smul_span_eq_top v.span_eq) i + (v.linear_independent.group_smul w) (group_smul_span_eq_top v.span_eq).ge i lemma units_smul_span_eq_top {v : ι → M} (hv : submodule.span R (set.range v) = ⊤) {w : ι → Rˣ} : submodule.span R (set.range (w • v)) = ⊤ := @@ -1013,12 +1012,12 @@ provides the basis corresponding to `w • v`. -/ def units_smul (v : basis ι R M) (w : ι → Rˣ) : basis ι R M := @basis.mk ι R M (w • v) _ _ _ - (v.linear_independent.units_smul w) (units_smul_span_eq_top v.span_eq) + (v.linear_independent.units_smul w) (units_smul_span_eq_top v.span_eq).ge lemma units_smul_apply {v : basis ι R M} {w : ι → Rˣ} (i : ι) : v.units_smul w i = w i • v i := mk_apply - (v.linear_independent.units_smul w) (units_smul_span_eq_top v.span_eq) i + (v.linear_independent.units_smul w) (units_smul_span_eq_top v.span_eq).ge i /-- A version of `smul_of_units` that uses `is_unit`. -/ def is_unit_smul (v : basis ι R M) {w : ι → R} (hw : ∀ i, is_unit (w i)): @@ -1043,8 +1042,7 @@ have span_b : submodule.span R (set.range (N.subtype ∘ b)) = N, @basis.mk _ _ _ (fin.cons y (N.subtype ∘ b) : fin (n + 1) → M) _ _ _ ((b.linear_independent.map' N.subtype (submodule.ker_subtype _)) .fin_cons' _ _ $ by { rintros c ⟨x, hx⟩ hc, rw span_b at hx, exact hli c x hx hc }) - (eq_top_iff.mpr (λ x _, - by { rw [fin.range_cons, submodule.mem_span_insert', span_b], exact hsp x })) + (λ x _, by { rw [fin.range_cons, submodule.mem_span_insert', span_b], exact hsp x }) @[simp] lemma coe_mk_fin_cons {n : ℕ} {N : submodule R M} (y : M) (b : basis (fin n) R N) (hli : ∀ (c : R) (x ∈ N), c • y + x = 0 → c = 0) @@ -1151,8 +1149,7 @@ noncomputable def extend (hs : linear_independent K (coe : s → V)) : basis _ K V := basis.mk (@linear_independent.restrict_of_comp_subtype _ _ _ id _ _ _ _ (hs.linear_independent_extend _)) - (eq_top_iff.mpr $ set_like.coe_subset_coe.mp $ - by simpa using hs.subset_span_extend (subset_univ s)) + (set_like.coe_subset_coe.mp $ by simpa using hs.subset_span_extend (subset_univ s)) lemma extend_apply_self (hs : linear_independent K (coe : s → V)) (x : hs.extend _) : diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index f09292b5cfca6..cf98839b81982 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -464,7 +464,7 @@ lemma is_basis_iff_det {v : ι → M} : begin split, { rintro ⟨hli, hspan⟩, - set v' := basis.mk hli hspan with v'_eq, + set v' := basis.mk hli hspan.ge with v'_eq, rw e.det_apply, convert linear_equiv.is_unit_det (linear_equiv.refl _ _) v' e using 2, ext i j, @@ -535,7 +535,7 @@ end /-- If we fix a background basis `e`, then for any other basis `v`, we can characterise the coordinates provided by `v` in terms of determinants relative to `e`. -/ lemma basis.det_smul_mk_coord_eq_det_update {v : ι → M} - (hli : linear_independent R v) (hsp : span R (range v) = ⊤) (i : ι) : + (hli : linear_independent R v) (hsp : ⊤ ≤ span R (range v)) (i : ι) : (e.det v) • (basis.mk hli hsp).coord i = e.det.to_multilinear_map.to_linear_map v i := begin apply (basis.mk hli hsp).ext, diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 79fe00f572d85..78a5476bd9dfb 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -1379,15 +1379,15 @@ section basis section division_ring variables [division_ring K] [add_comm_group V] [module K V] -lemma linear_independent_of_span_eq_top_of_card_eq_finrank {ι : Type*} [fintype ι] {b : ι → V} - (span_eq : span K (set.range b) = ⊤) (card_eq : fintype.card ι = finrank K V) : +lemma linear_independent_of_top_le_span_of_card_eq_finrank {ι : Type*} [fintype ι] {b : ι → V} + (spans : ⊤ ≤ span K (set.range b)) (card_eq : fintype.card ι = finrank K V) : linear_independent K b := linear_independent_iff'.mpr $ λ s g dependent i i_mem_s, begin by_contra gx_ne_zero, -- We'll derive a contradiction by showing `b '' (univ \ {i})` of cardinality `n - 1` -- spans a vector space of dimension `n`. - refine ne_of_lt (span_lt_top_of_card_lt_finrank + refine not_le_of_gt (span_lt_top_of_card_lt_finrank (show (b '' (set.univ \ {i})).to_finset.card < finrank K V, from _)) _, { calc (b '' (set.univ \ {i})).to_finset.card = ((set.univ \ {i}).to_finset.image b).card : by rw [set.to_finset_card, fintype.card_of_finset] @@ -1398,7 +1398,7 @@ begin -- We already have that `b '' univ` spans the whole space, -- so we only need to show that the span of `b '' (univ \ {i})` contains each `b j`. - refine trans (le_antisymm (span_mono (set.image_subset_range _ _)) (span_le.mpr _)) span_eq, + refine spans.trans (span_le.mpr _), rintros _ ⟨j, rfl, rfl⟩, -- The case that `j ≠ i` is easy because `b j ∈ b '' (univ \ {i})`. by_cases j_eq : j = i, @@ -1438,9 +1438,8 @@ begin let f := (submodule.subtype (span K (set.range b))), let b' : ι → span K (set.range b) := λ i, ⟨b i, mem_span.2 (λ p hp, hp (set.mem_range_self _))⟩, - have hs : span K (set.range b') = ⊤, - { rw eq_top_iff', - intro x, + have hs : ⊤ ≤ span K (set.range b'), + { intro x, have h : span K (f '' (set.range b')) = map f (span K (set.range b')) := span_image f, have hf : f '' (set.range b') = set.range b, { ext x, simp [set.mem_image, set.mem_range] }, rw hf at h, @@ -1448,36 +1447,36 @@ begin conv at hx { congr, skip, rw h }, simpa [mem_map] using hx }, have hi : f.ker = ⊥ := ker_subtype _, - convert (linear_independent_of_span_eq_top_of_card_eq_finrank hs hc).map' _ hi } + convert (linear_independent_of_top_le_span_of_card_eq_finrank hs hc).map' _ hi } end /-- A family of `finrank K V` vectors forms a basis if they span the whole space. -/ -noncomputable def basis_of_span_eq_top_of_card_eq_finrank {ι : Type*} [fintype ι] (b : ι → V) - (span_eq : span K (set.range b) = ⊤) (card_eq : fintype.card ι = finrank K V) : +noncomputable def basis_of_top_le_span_of_card_eq_finrank {ι : Type*} [fintype ι] (b : ι → V) + (le_span : ⊤ ≤ span K (set.range b)) (card_eq : fintype.card ι = finrank K V) : basis ι K V := -basis.mk (linear_independent_of_span_eq_top_of_card_eq_finrank span_eq card_eq) span_eq +basis.mk (linear_independent_of_top_le_span_of_card_eq_finrank le_span card_eq) le_span -@[simp] lemma coe_basis_of_span_eq_top_of_card_eq_finrank {ι : Type*} [fintype ι] (b : ι → V) - (span_eq : span K (set.range b) = ⊤) (card_eq : fintype.card ι = finrank K V) : - ⇑(basis_of_span_eq_top_of_card_eq_finrank b span_eq card_eq) = b := +@[simp] lemma coe_basis_of_top_le_span_of_card_eq_finrank {ι : Type*} [fintype ι] (b : ι → V) + (le_span : ⊤ ≤ span K (set.range b)) (card_eq : fintype.card ι = finrank K V) : + ⇑(basis_of_top_le_span_of_card_eq_finrank b le_span card_eq) = b := basis.coe_mk _ _ /-- A finset of `finrank K V` vectors forms a basis if they span the whole space. -/ @[simps] -noncomputable def finset_basis_of_span_eq_top_of_card_eq_finrank {s : finset V} - (span_eq : span K (s : set V) = ⊤) (card_eq : s.card = finrank K V) : +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 := -basis_of_span_eq_top_of_card_eq_finrank (coe : (s : set V) → V) - ((@subtype.range_coe_subtype _ (λ x, x ∈ s)).symm ▸ span_eq) +basis_of_top_le_span_of_card_eq_finrank (coe : (s : set V) → V) + ((@subtype.range_coe_subtype _ (λ x, x ∈ s)).symm ▸ le_span) (trans (fintype.card_coe _) card_eq) /-- A set of `finrank K V` vectors forms a basis if they span the whole space. -/ @[simps] -noncomputable def set_basis_of_span_eq_top_of_card_eq_finrank {s : set V} [fintype s] - (span_eq : span K s = ⊤) (card_eq : s.to_finset.card = finrank K V) : +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 := -basis_of_span_eq_top_of_card_eq_finrank (coe : s → V) - ((@subtype.range_coe_subtype _ s).symm ▸ span_eq) +basis_of_top_le_span_of_card_eq_finrank (coe : s → V) + ((@subtype.range_coe_subtype _ s).symm ▸ le_span) (trans s.to_finset_card.symm card_eq) end division_ring @@ -1510,7 +1509,7 @@ noncomputable def basis_of_linear_independent_of_card_eq_finrank (lin_ind : linear_independent K b) (card_eq : fintype.card ι = finrank K V) : basis ι K V := basis.mk lin_ind $ -span_eq_top_of_linear_independent_of_card_eq_finrank lin_ind card_eq +(span_eq_top_of_linear_independent_of_card_eq_finrank lin_ind card_eq).ge @[simp] lemma coe_basis_of_linear_independent_of_card_eq_finrank {ι : Type*} [nonempty ι] [fintype ι] {b : ι → V} diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 6aa3751809ef1..8b4c99f64392d 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -396,7 +396,7 @@ begin have b''_li : linear_independent _ b'' := _, have b''_sp : submodule.span _ (set.range b'') = ⊤ := _, -- Since the two bases have the same index set, the spaces have the same dimension. - let c : basis ι K L := basis.mk b''_li b''_sp, + let c : basis ι K L := basis.mk b''_li b''_sp.ge, rw [finrank_eq_card_basis b, finrank_eq_card_basis c], -- It remains to show that the basis is indeed linear independent and spans the whole space. { rw set.range_comp, diff --git a/src/ring_theory/adjoin/power_basis.lean b/src/ring_theory/adjoin/power_basis.lean index bd9316b97588f..a2b132111e4d0 100644 --- a/src/ring_theory/adjoin/power_basis.lean +++ b/src/ring_theory/adjoin/power_basis.lean @@ -41,8 +41,7 @@ begin (adjoin K {x}) (λ i, ⟨x, subset_adjoin (set.mem_singleton x)⟩ ^ (i : ℕ)), { have := hx'.linear_independent_pow, rwa minpoly_eq at this }, - { rw _root_.eq_top_iff, - rintros ⟨y, hy⟩ _, + { rintros ⟨y, hy⟩ _, have := hx'.mem_span_pow, rw minpoly_eq at this, apply this, diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index a13d23f0d0864..ef8f030aa82c5 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -340,8 +340,7 @@ begin 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 }, - { rw _root_.eq_top_iff, - rintros y -, + { rintros y -, rw [← deg_f', ← minpoly_eq], apply (is_integral_root hf).mem_span_pow, obtain ⟨g⟩ := y, diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 71461eea136b9..8e5e05041555a 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -128,7 +128,7 @@ begin { have := span_eq_top_of_linear_independent_of_card_eq_finrank b.linear_independent (finrank_eq_card_basis b).symm, rw [discr_def, trace_matrix_def], - simp_rw [← basis.mk_apply b.linear_independent this], + 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], exact trace_form_nondegenerate _ _ }, end diff --git a/src/ring_theory/localization/module.lean b/src/ring_theory/localization/module.lean index e47d7f15b914a..bca0672c79990 100644 --- a/src/ring_theory/localization/module.lean +++ b/src/ring_theory/localization/module.lean @@ -60,7 +60,7 @@ variables {M : Type*} [add_comm_group M] [module R M] [module Rₛ M] [is_scalar /-- 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 [← @submodule.restrict_scalars_eq_top_iff Rₛ R, eq_top_iff, ← b.span_eq], +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 } end add_comm_group From e846e899ec78d4d303302ff9ea4c575ddc8a3090 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 29 Jul 2022 05:53:23 +0000 Subject: [PATCH 134/173] feat(linear_algebra/clifford_algebra/even): Universal property and isomorphisms for the even subalgebra (#14790) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main three results here are: * `even.lift : even_hom Q A ≃ (clifford_algebra.even Q →ₐ[R] A)` * `equiv_even : clifford_algebra Q ≃ₐ[R] clifford_algebra.even (Q' Q)` * `even_equiv_even_neg : clifford_algebra.even Q ≃ₐ[R] clifford_algebra.even (-Q)` --- src/linear_algebra/clifford_algebra/even.lean | 240 +++++++++++++++++ .../clifford_algebra/even_equiv.lean | 247 ++++++++++++++++++ 2 files changed, 487 insertions(+) create mode 100644 src/linear_algebra/clifford_algebra/even.lean create mode 100644 src/linear_algebra/clifford_algebra/even_equiv.lean diff --git a/src/linear_algebra/clifford_algebra/even.lean b/src/linear_algebra/clifford_algebra/even.lean new file mode 100644 index 0000000000000..665fe2a12f547 --- /dev/null +++ b/src/linear_algebra/clifford_algebra/even.lean @@ -0,0 +1,240 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import linear_algebra.clifford_algebra.fold +import linear_algebra.clifford_algebra.grading + +/-! +# The universal property of the even subalgebra + +## Main definitions + +* `clifford_algebra.even Q`: The even subalgebra of `clifford_algebra Q`. +* `clifford_algebra.even_hom`: The type of bilinear maps that satisfy the universal property of the + even subalgebra +* `clifford_algebra.even.lift`: The universal property of the even subalgebra, which states + that every bilinear map `f` with `f v v = Q v` and `f u v * f v w = Q v • f u w` is in unique + correspondence with an algebra morphism from `clifford_algebra.even Q`. + +## Implementation notes + +The approach here is outlined in "Computing with the universal properties of the Clifford algebra +and the even subalgebra" (to appear). + +The broad summary is that we have two tricks available to us for implementing complex recursors on +top of `clifford_algebra.lift`: the first is to use morphisms as the output type, such as +`A = module.End R N` which is how we obtained `clifford_algebra.foldr`; and the second is to use +`N = (N', S)` where `N'` is the value we wish to compute, and `S` is some auxiliary state passed +between one recursor invocation and the next. +For the universal property of the even subalgebra, we apply a variant of the first trick again by +choosing `S` to itself be a submodule of morphisms. +-/ + +namespace clifford_algebra + +variables {R M : Type*} [comm_ring R] [add_comm_group M] [module R M] +variables {Q : quadratic_form R M} +-- put this after `Q` since we want to talk about morphisms from `clifford_algebra Q` to `A` and +-- that order is more natural +variables {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] + +open_locale direct_sum + +variables (Q) + +/-- The even submodule `clifford_algebra.even_odd Q 0` is also a subalgebra. -/ +def even : subalgebra R (clifford_algebra Q) := +(even_odd Q 0).to_subalgebra + set_like.graded_monoid.one_mem + (λ x y hx hy, add_zero (0 : zmod 2) ▸ set_like.graded_monoid.mul_mem hx hy) + +@[simp] lemma even_to_submodule : (even Q).to_submodule = even_odd Q 0 := +rfl + +variables (A) + +/-- The type of bilinear maps which are accepted by `clifford_algebra.even.lift`. -/ +@[ext] +structure even_hom : Type (max u_2 u_3) := +(bilin : M →ₗ[R] M →ₗ[R] A) +(contract (m : M) : bilin m m = algebra_map R A (Q m)) +(contract_mid (m₁ m₂ m₃ : M) : bilin m₁ m₂ * bilin m₂ m₃ = Q m₂ • bilin m₁ m₃) + +variables {A Q} + +/-- Compose an `even_hom` with an `alg_hom` on the output. -/ +@[simps] +def even_hom.compr₂ (g : even_hom Q A) (f : A →ₐ[R] B) : even_hom Q B := +{ bilin := g.bilin.compr₂ f.to_linear_map, + contract := λ m, (f.congr_arg $ g.contract _).trans $ f.commutes _, + contract_mid := λ m₁ m₂ m₃, (f.map_mul _ _).symm.trans $ + (f.congr_arg $ g.contract_mid _ _ _).trans $ f.map_smul _ _ } + +variables (Q) + +/-- The embedding of pairs of vectors into the even subalgebra, as a bilinear map. -/ +@[simps bilin_apply_apply_coe] +def even.ι : even_hom Q (even Q) := +{ bilin := linear_map.mk₂ R (λ m₁ m₂, ⟨ι Q m₁ * ι Q m₂, ι_mul_ι_mem_even_odd_zero _ _ _⟩) + (λ _ _ _, by { simp only [linear_map.map_add, add_mul], refl }) + (λ _ _ _, by { simp only [linear_map.map_smul, smul_mul_assoc], refl }) + (λ _ _ _, by { simp only [linear_map.map_add, mul_add], refl }) + (λ _ _ _, by { simp only [linear_map.map_smul, mul_smul_comm], refl }), + contract := λ m, subtype.ext $ ι_sq_scalar Q m, + contract_mid := λ m₁ m₂ m₃, subtype.ext $ + calc ι Q m₁ * ι Q m₂ * (ι Q m₂ * ι Q m₃) + = ι Q m₁ * ((ι Q m₂ * ι Q m₂) * ι Q m₃) : by simp only [mul_assoc] + ... = Q m₂ • (ι Q m₁ * ι Q m₃) : by rw [algebra.smul_def, ι_sq_scalar, algebra.left_comm] } + +instance : inhabited (even_hom Q (even Q)) := ⟨even.ι Q⟩ + +variables (f : even_hom Q A) + +/-- Two algebra morphisms from the even subalgebra are equal if they agree on pairs of generators. + +See note [partially-applied ext lemmas]. -/ +@[ext] +lemma even.alg_hom_ext ⦃f g : even Q →ₐ[R] A⦄ + (h : (even.ι Q).compr₂ f = (even.ι Q).compr₂ g) : + f = g := +begin + rw even_hom.ext_iff at h, + ext ⟨x, hx⟩, + refine even_induction _ _ _ _ _ hx, + { intro r, + exact (f.commutes r).trans (g.commutes r).symm }, + { intros x y hx hy ihx ihy, + have := congr_arg2 (+) ihx ihy, + exact (f.map_add _ _).trans (this.trans $ (g.map_add _ _).symm) }, + { intros m₁ m₂ x hx ih, + have := congr_arg2 (*) (linear_map.congr_fun (linear_map.congr_fun h m₁) m₂) ih, + exact (f.map_mul _ _).trans (this.trans $ (g.map_mul _ _).symm) }, +end + +variables {Q} + +namespace even.lift + +/-- An auxiliary submodule used to store the half-applied values of `f`. +This is the span of elements `f'` such that `∃ x m₂, ∀ m₁, f' m₁ = f m₁ m₂ * x`. -/ +private def S : submodule R (M →ₗ[R] A) := +submodule.span R {f' | ∃ x m₂, f' = linear_map.lcomp R _ (f.bilin.flip m₂) (algebra.lmul_right R x)} + +/-- An auxiliary bilinear map that is later passed into `clifford_algebra.fold`. Our desired result +is stored in the `A` part of the accumulator, while auxiliary recursion state is stored in the `S f` +part. -/ +private def f_fold : M →ₗ[R] (A × S f) →ₗ[R] (A × S f) := +linear_map.mk₂ R (λ m acc, + /- We could write this `snd` term in a point-free style as follows, but it wouldn't help as we + don't have any prod or subtype combinators to deal with n-linear maps of this degree. + ```lean + (linear_map.lcomp R _ (algebra.lmul R A).to_linear_map.flip).comp $ + (linear_map.llcomp R M A A).flip.comp f.flip : M →ₗ[R] A →ₗ[R] M →ₗ[R] A) + ``` + -/ + (acc.2 m, ⟨(algebra.lmul_right R acc.1).comp (f.bilin.flip m), + submodule.subset_span $ ⟨_, _, rfl⟩⟩)) + (λ m₁ m₂ a, prod.ext + (linear_map.map_add _ m₁ m₂) + (subtype.ext $ linear_map.ext $ λ m₃, + show f.bilin m₃ (m₁ + m₂) * a.1 = f.bilin m₃ m₁ * a.1 + f.bilin m₃ m₂ * a.1, + by rw [map_add, add_mul])) + (λ c m a, prod.ext + (linear_map.map_smul _ c m) + (subtype.ext $ linear_map.ext $ λ m₃, + show f.bilin m₃ (c • m) * a.1 = c • (f.bilin m₃ m * a.1), + by rw [linear_map.map_smul, smul_mul_assoc])) + (λ m a₁ a₂, prod.ext rfl (subtype.ext $ linear_map.ext $ λ m₃, mul_add _ _ _)) + (λ c m a, prod.ext rfl (subtype.ext $ linear_map.ext $ λ m₃, mul_smul_comm _ _ _)) + +@[simp] private lemma fst_f_fold_f_fold (m₁ m₂ : M) (x : A × S f) : + (f_fold f m₁ (f_fold f m₂ x)).fst = f.bilin m₁ m₂ * x.fst := rfl + +@[simp] private lemma snd_f_fold_f_fold (m₁ m₂ m₃ : M) (x : A × S f) : + ((f_fold f m₁ (f_fold f m₂ x)).snd : M →ₗ[R] A) m₃ = f.bilin m₃ m₁ * (x.snd : M →ₗ[R] A) m₂ := rfl + +private lemma f_fold_f_fold (m : M) (x : A × S f) : + f_fold f m (f_fold f m x) = Q m • x := +begin + obtain ⟨a, ⟨g, hg⟩⟩ := x, + ext : 2, + { change f.bilin m m * a = Q m • a, + rw [algebra.smul_def, f.contract] }, + { ext m₁, + change f.bilin _ _ * g m = Q m • g m₁, + apply submodule.span_induction' _ _ _ _ hg, + { rintros _ ⟨b, m₃, rfl⟩, + change f.bilin _ _ * (f.bilin _ _ * b) = Q m • (f.bilin _ _ * b), + rw [←smul_mul_assoc, ←mul_assoc, f.contract_mid] }, + { change f.bilin m₁ m * 0 = Q m • 0, + rw [mul_zero, smul_zero] }, + { rintros x hx y hy ihx ihy, + rw [linear_map.add_apply, linear_map.add_apply, mul_add, smul_add, ihx, ihy] }, + { rintros x hx c ihx, + rw [linear_map.smul_apply, linear_map.smul_apply, mul_smul_comm, ihx, smul_comm] } }, +end + +/-- The final auxiliary construction for `clifford_algebra.even.lift`. This map is the forwards +direction of that equivalence, but not in the fully-bundled form. -/ +@[simps apply {attrs := []}] def aux (f : even_hom Q A) : clifford_algebra.even Q →ₗ[R] A := +begin + refine _ ∘ₗ (even Q).val.to_linear_map, + exact linear_map.fst _ _ _ ∘ₗ foldr Q (f_fold f) (f_fold_f_fold f) (1, 0), +end + +@[simp] lemma aux_one : aux f 1 = 1 := +(congr_arg prod.fst (foldr_one _ _ _ _)) + +@[simp] lemma aux_ι (m₁ m₂ : M) : aux f ((even.ι Q).bilin m₁ m₂) = f.bilin m₁ m₂ := +(congr_arg prod.fst (foldr_mul _ _ _ _ _ _)).trans begin + rw [foldr_ι, foldr_ι], + exact mul_one _, +end + +@[simp] lemma aux_algebra_map (r) (hr) : + aux f ⟨algebra_map R _ r, hr⟩ = algebra_map R _ r := +(congr_arg prod.fst (foldr_algebra_map _ _ _ _ _)).trans (algebra.algebra_map_eq_smul_one r).symm + +@[simp] lemma aux_mul (x y : even Q) : + aux f (x * y) = aux f x * aux f y := +begin + cases x, + cases y, + refine (congr_arg prod.fst (foldr_mul _ _ _ _ _ _)).trans _, + dsimp only, + refine even_induction Q _ _ _ _ x_property, + { intros r, + rw [foldr_algebra_map, aux_algebra_map], + exact (algebra.smul_def r _), }, + { intros x y hx hy ihx ihy, + rw [linear_map.map_add, prod.fst_add, ihx, ihy, ←add_mul, ←linear_map.map_add], + refl, }, + { rintros m₁ m₂ x (hx : x ∈ even Q) ih, + rw [aux_apply, foldr_mul, foldr_mul, foldr_ι, foldr_ι, fst_f_fold_f_fold, ih, + ←mul_assoc, subtype.coe_mk, foldr_mul, foldr_mul, foldr_ι, foldr_ι, fst_f_fold_f_fold], + refl } +end + +end even.lift + +open even.lift + +variables (Q) {A} + +/-- Every algebra morphism from the even subalgebra is in one-to-one correspondence with a +bilinear map that sends duplicate arguments to the quadratic form, and contracts across +multiplication. -/ +@[simps symm_apply_bilin] +def even.lift : even_hom Q A ≃ (clifford_algebra.even Q →ₐ[R] A) := +{ to_fun := λ f, alg_hom.of_linear_map (aux f) (aux_one f) (aux_mul f), + inv_fun := λ F, (even.ι Q).compr₂ F, + left_inv := λ f, even_hom.ext _ _ $ linear_map.ext₂ $ even.lift.aux_ι f, + right_inv := λ F, even.alg_hom_ext Q $ even_hom.ext _ _ $ linear_map.ext₂ $ even.lift.aux_ι _ } + +@[simp] lemma even.lift_ι (f : even_hom Q A) (m₁ m₂ : M) : + even.lift Q f ((even.ι Q).bilin m₁ m₂) = f.bilin m₁ m₂ := +even.lift.aux_ι _ _ _ + +end clifford_algebra diff --git a/src/linear_algebra/clifford_algebra/even_equiv.lean b/src/linear_algebra/clifford_algebra/even_equiv.lean new file mode 100644 index 0000000000000..b9ec5de35dad8 --- /dev/null +++ b/src/linear_algebra/clifford_algebra/even_equiv.lean @@ -0,0 +1,247 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import linear_algebra.clifford_algebra.conjugation +import linear_algebra.clifford_algebra.even +import linear_algebra.quadratic_form.prod +/-! +# Isomorphisms with the even subalgebra of a Clifford algebra + +This file provides some notable isomorphisms regarding the even subalgebra, `clifford_algebra.even`. + +## Main definitions + +* `clifford_algebra.equiv_even`: Every Clifford algebra is isomorphic as an algebra to the even + subalgebra of a Clifford algebra with one more dimension. + * `clifford_algebra.even_equiv.Q'`: The quadratic form used by this "one-up" algebra. + * `clifford_algebra.to_even`: The simp-normal form of the forward direction of this isomorphism. + * `clifford_algebra.of_even`: The simp-normal form of the reverse direction of this isomorphism. + +* `clifford_algebra.even_equiv_even_neg`: Every even subalgebra is isomorphic to the even subalgebra + of the Clifford algebra with negated quadratic form. + * `clifford_algebra.even_to_neg`: The simp-normal form of each direction of this isomorphism. + +## Main results + +* `clifford_algebra.coe_to_even_reverse_involute`: the behavior of `clifford_algebra.to_even` on the + "Clifford conjugate", that is `clifford_algebra.reverse` composed with + `clifford_algebra.involute`. +-/ + +namespace clifford_algebra + +variables {R M : Type*} [comm_ring R] [add_comm_group M] [module R M] +variables (Q : quadratic_form R M) + +/-! ### Constructions needed for `clifford_algebra.equiv_even` -/ + +namespace equiv_even + +/-- The quadratic form on the augmented vector space `M × R` sending `v + r•e0` to `Q v - r^2`. -/ +@[reducible] +def Q' : quadratic_form R (M × R) := (Q.prod $ -@quadratic_form.sq R _) + +lemma Q'_apply (m : M × R) : Q' Q m = Q m.1 - m.2 * m.2 := (sub_eq_add_neg _ _).symm + +/-- The unit vector in the new dimension -/ +def e0 : clifford_algebra (Q' Q) := ι (Q' Q) (0, 1) + +/-- The embedding from the existing vector space -/ +def v : M →ₗ[R] clifford_algebra (Q' Q) := (ι (Q' Q)) ∘ₗ linear_map.inl _ _ _ + +lemma ι_eq_v_add_smul_e0 (m : M) (r : R) : ι (Q' Q) (m, r) = v Q m + r • e0 Q := +by rw [e0, v, linear_map.comp_apply, linear_map.inl_apply, ←linear_map.map_smul, prod.smul_mk, + smul_zero, smul_eq_mul, mul_one, ←linear_map.map_add, prod.mk_add_mk, zero_add, add_zero] + +lemma e0_mul_e0 : e0 Q * e0 Q = -1 := +(ι_sq_scalar _ _).trans $ by simp + +lemma v_sq_scalar (m : M) : v Q m * v Q m = algebra_map _ _ (Q m) := +(ι_sq_scalar _ _).trans $ by simp + +lemma neg_e0_mul_v (m : M) : -(e0 Q * v Q m) = v Q m * e0 Q := +begin + refine neg_eq_of_add_eq_zero_right ((ι_mul_ι_add_swap _ _).trans _), + dsimp [quadratic_form.polar], + simp only [add_zero, mul_zero, mul_one, zero_add, neg_zero, quadratic_form.map_zero, + add_sub_cancel, sub_self, map_zero, zero_sub], +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 +end + +@[simp] lemma e0_mul_v_mul_e0 (m : M) : e0 Q * v Q m * e0 Q = v Q m := +by rw [←neg_v_mul_e0, ←neg_mul, mul_assoc, e0_mul_e0, mul_neg_one, neg_neg] + +@[simp] lemma reverse_v (m : M) : reverse (v Q m) = v Q m := reverse_ι _ +@[simp] lemma involute_v (m : M) : involute (v Q m) = -v Q m := involute_ι _ + +@[simp] lemma reverse_e0 : reverse (e0 Q) = e0 Q := reverse_ι _ +@[simp] lemma involute_e0 : involute (e0 Q) = -e0 Q := involute_ι _ + +end equiv_even + +open equiv_even + +/-- The embedding from the smaller algebra into the new larger one. -/ +def to_even : clifford_algebra Q →ₐ[R] clifford_algebra.even (Q' Q) := +begin + refine clifford_algebra.lift Q ⟨_, λ m, _⟩, + { refine linear_map.cod_restrict _ _ (λ m, submodule.mem_supr_of_mem ⟨2, rfl⟩ _), + exact (algebra.lmul_left R $ e0 Q).comp (v Q), + rw [subtype.coe_mk, pow_two], + exact submodule.mul_mem_mul (linear_map.mem_range_self _ _) (linear_map.mem_range_self _ _), }, + { ext1, + dsimp only [subalgebra.coe_mul, linear_map.cod_restrict_apply, linear_map.comp_apply, + algebra.lmul_left_apply, linear_map.inl_apply, subalgebra.coe_algebra_map], + rw [←mul_assoc, e0_mul_v_mul_e0, v_sq_scalar] } +end + +@[simp] +lemma to_even_ι (m : M) : (to_even Q (ι Q m) : clifford_algebra (Q' Q)) = e0 Q * v Q m := +begin + rw [to_even, clifford_algebra.lift_ι_apply, linear_map.cod_restrict_apply], + refl, +end + +/-- The embedding from the even subalgebra with an extra dimension into the original algebra. -/ +def of_even : clifford_algebra.even (Q' Q) →ₐ[R] clifford_algebra Q := +begin + /- + Recall that we need: + * `f ⟨0,1⟩ ⟨x,0⟩ = ι x` + * `f ⟨x,0⟩ ⟨0,1⟩ = -ι x` + * `f ⟨x,0⟩ ⟨y,0⟩ = ι x * ι y` + * `f ⟨0,1⟩ ⟨0,1⟩ = -1` + -/ + let f : (M × R) →ₗ[R] (M × R) →ₗ[R] clifford_algebra Q := + ((algebra.lmul R (clifford_algebra Q)).to_linear_map.comp + $ ((ι Q).comp (linear_map.fst _ _ _)) + + (algebra.linear_map R _).comp (linear_map.snd _ _ _)).compl₂ + (((ι Q).comp (linear_map.fst _ _ _)) - (algebra.linear_map R _).comp (linear_map.snd _ _ _)), + have f_apply : + ∀ x y, f x y = (ι Q x.1 + algebra_map R _ x.2) * (ι Q y.1 - algebra_map R _ y.2) := + λ x y, rfl, + have hc : ∀ (r : R) (x : clifford_algebra Q), commute (algebra_map _ _ r) x := algebra.commutes, + have hm : ∀ m : M × R, + ι Q m.1 * ι Q m.1 - algebra_map R _ m.2 * algebra_map R _ m.2 = algebra_map R _ (Q' Q m), + { intro m, + rw [ι_sq_scalar, ←ring_hom.map_mul, ←ring_hom.map_sub, + sub_eq_add_neg, Q'_apply, sub_eq_add_neg] }, + refine even.lift (Q' Q) ⟨f, _, _⟩; simp_rw [f_apply], + { intro m, + rw [←(hc _ _).symm.mul_self_sub_mul_self_eq, hm] }, + { intros m₁ m₂ m₃, + rw [←mul_smul_comm, ←mul_assoc, mul_assoc(_ + _), ←(hc _ _).symm.mul_self_sub_mul_self_eq', + algebra.smul_def, ←mul_assoc, hm] }, +end + +lemma of_even_ι (x y : M × R) : + of_even Q ((even.ι _).bilin x y) = + (ι Q x.1 + algebra_map R _ x.2) * (ι Q y.1 - algebra_map R _ y.2) := +even.lift_ι _ _ _ _ + +lemma to_even_comp_of_even : (to_even Q).comp (of_even Q) = alg_hom.id R _ := +even.alg_hom_ext (Q' Q) $ even_hom.ext _ _ $ linear_map.ext $ λ m₁, linear_map.ext $ λ m₂, + subtype.ext $ + let ⟨m₁, r₁⟩ := m₁, ⟨m₂, r₂⟩ := m₂ in + calc ↑(to_even Q (of_even Q ((even.ι (Q' Q)).bilin (m₁, r₁) (m₂, r₂)))) + = (e0 Q * v Q m₁ + algebra_map R _ r₁) * (e0 Q * v Q m₂ - algebra_map R _ r₂) : + by rw [of_even_ι, alg_hom.map_mul, alg_hom.map_add, alg_hom.map_sub, alg_hom.commutes, + alg_hom.commutes, subalgebra.coe_mul, subalgebra.coe_add, subalgebra.coe_sub, + to_even_ι, to_even_ι, subalgebra.coe_algebra_map, subalgebra.coe_algebra_map] + ... = e0 Q * v Q m₁ * (e0 Q * v Q m₂) + r₁ • e0 Q * v Q m₂ - r₂ • e0 Q * v Q m₁ + - algebra_map R _ (r₁ * r₂) : + by rw [mul_sub, add_mul, add_mul, ←algebra.commutes, ←algebra.smul_def, ←map_mul, + ←algebra.smul_def, sub_add_eq_sub_sub, smul_mul_assoc, smul_mul_assoc] + ... = v Q m₁ * v Q m₂ + r₁ • e0 Q * v Q m₂ + v Q m₁ * r₂ • e0 Q + (r₁ • e0 Q) * r₂ • e0 Q : + have h1 : e0 Q * v Q m₁ * (e0 Q * v Q m₂) = v Q m₁ * v Q m₂, + by rw [←mul_assoc, e0_mul_v_mul_e0], + have h2 : -(r₂ • e0 Q * v Q m₁) = v Q m₁ * r₂ • e0 Q, + by rw [mul_smul_comm, smul_mul_assoc, ←smul_neg, neg_e0_mul_v], + have h3 : - algebra_map R _ (r₁ * r₂) = (r₁ • e0 Q) * r₂ • e0 Q, + by rw [algebra.algebra_map_eq_smul_one, smul_mul_smul, e0_mul_e0, smul_neg], + by rw [sub_eq_add_neg, sub_eq_add_neg, h1, h2, h3] + ... = ι _ (m₁, r₁) * ι _ (m₂, r₂) : + by rw [ι_eq_v_add_smul_e0, ι_eq_v_add_smul_e0, mul_add, add_mul, add_mul, add_assoc] + +lemma of_even_comp_to_even : + (of_even Q).comp (to_even Q) = alg_hom.id R _ := +clifford_algebra.hom_ext $ linear_map.ext $ λ m, + calc of_even Q (to_even Q (ι Q m)) + = of_even Q ⟨_, (to_even Q (ι Q m)).prop⟩ : by rw subtype.coe_eta + ... = (ι Q 0 + algebra_map R _ 1) * (ι Q m - algebra_map R _ 0) : begin + simp_rw to_even_ι, + exact of_even_ι Q _ _, + end + ... = ι Q m : by rw [map_one, map_zero, map_zero, sub_zero, zero_add, one_mul] + +/-- Any clifford algebra is isomorphic to the even subalgebra of a clifford algebra with an extra +dimension (that is, with vector space `M × R`), with a quadratic form evaluating to `-1` on that new +basis vector. -/ +@[simps] +def equiv_even : clifford_algebra Q ≃ₐ[R] clifford_algebra.even (Q' Q) := +alg_equiv.of_alg_hom + (to_even Q) + (of_even Q) + (to_even_comp_of_even Q) + (of_even_comp_to_even Q) + +/-- The representation of the clifford conjugate (i.e. the reverse of the involute) in the even +subalgebra is just the reverse of the representation. -/ +lemma coe_to_even_reverse_involute (x : clifford_algebra Q) : + ↑(to_even Q (reverse (involute x))) = reverse (to_even Q x : clifford_algebra (Q' Q)) := +begin + induction x using clifford_algebra.induction, + case h_grade0 : r { simp only [alg_hom.commutes, subalgebra.coe_algebra_map, reverse.commutes] }, + case h_grade1 : m + { simp only [involute_ι, subalgebra.coe_neg, to_even_ι, reverse.map_mul, + reverse_v, reverse_e0, reverse_ι, neg_e0_mul_v, map_neg] }, + case h_mul : x y hx hy { simp only [map_mul, subalgebra.coe_mul, reverse.map_mul, hx, hy] }, + case h_add : x y hx hy { simp only [map_add, subalgebra.coe_add, hx, hy] }, +end + +/-! ### Constructions needed for `clifford_algebra.even_equiv_even_neg` -/ + +/-- One direction of `clifford_algebra.even_equiv_even_neg` -/ +def even_to_neg (Q' : quadratic_form R M) (h : Q' = -Q) : + clifford_algebra.even Q →ₐ[R] clifford_algebra.even Q' := +even.lift Q + { bilin := -(even.ι Q' : _).bilin, + contract := λ m, by simp_rw [linear_map.neg_apply, even_hom.contract, h, + quadratic_form.neg_apply, map_neg, neg_neg], + contract_mid := λ m₁ m₂ m₃, + by simp_rw [linear_map.neg_apply, neg_mul_neg, even_hom.contract_mid, h, + quadratic_form.neg_apply, smul_neg, neg_smul] } + +@[simp] lemma even_to_neg_ι (Q' : quadratic_form R M) (h : Q' = -Q) (m₁ m₂ : M) : + even_to_neg Q Q' h ((even.ι Q).bilin m₁ m₂) = -(even.ι Q').bilin m₁ m₂ := +even.lift_ι _ _ m₁ m₂ + +lemma even_to_neg_comp_even_to_neg (Q' : quadratic_form R M) + (h : Q' = -Q) (h' : Q = -Q') : + (even_to_neg Q' Q h').comp (even_to_neg Q Q' h) = alg_hom.id R _ := +begin + ext m₁ m₂ : 4, + dsimp only [even_hom.compr₂_bilin, linear_map.compr₂_apply, alg_hom.to_linear_map_apply, + alg_hom.comp_apply, alg_hom.id_apply], + rw [even_to_neg_ι, map_neg, even_to_neg_ι, neg_neg] +end + +/-- The even subalgebras of the algebras with quadratic form `Q` and `-Q` are isomorphic. + +Stated another way, `𝒞ℓ⁺(p,q,r)` and `𝒞ℓ⁺(q,p,r)` are isomorphic. -/ +@[simps] +def even_equiv_even_neg : clifford_algebra.even Q ≃ₐ[R] clifford_algebra.even (-Q) := +alg_equiv.of_alg_hom + (even_to_neg Q _ rfl) + (even_to_neg (-Q) _ (neg_neg _).symm) + (even_to_neg_comp_even_to_neg _ _ _ _) + (even_to_neg_comp_even_to_neg _ _ _ _) + +end clifford_algebra From 59e69a26e90c12028e03430818380f9a94036ef2 Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Fri, 29 Jul 2022 08:40:24 +0000 Subject: [PATCH 135/173] feat(linear_algebra/matrix): inner prod of matrix (#14794) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A positive definite matrix `M` induces an inner product `⟪x, y⟫ = xᴴMy`. --- src/linear_algebra/matrix/pos_def.lean | 36 +++++++++++++++++++++++--- 1 file changed, 33 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index b4d9bb628cafa..036ba42ff1c3c 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -20,14 +20,16 @@ quadratic forms. namespace matrix -variables {R : Type*} [ordered_semiring R] [star_ring R] {n : Type*} [fintype n] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] open_locale matrix /-- A matrix `M : matrix n n R` is positive definite if it is hermitian and `xᴴMx` is greater than zero for all nonzero `x`. -/ -def pos_def (M : matrix n n R) := -M.is_hermitian ∧ ∀ x : n → R, x ≠ 0 → 0 < dot_product (star x) (M.mul_vec x) +def pos_def (M : matrix n n 𝕜) := +M.is_hermitian ∧ ∀ x : n → 𝕜, x ≠ 0 → 0 < is_R_or_C.re (dot_product (star x) (M.mul_vec x)) + +lemma pos_def.is_hermitian {M : matrix n n 𝕜} (hM : M.pos_def) : M.is_hermitian := hM.1 lemma pos_def_of_to_quadratic_form' [decidable_eq n] {M : matrix n n ℝ} (hM : M.is_symm) (hMq : M.to_quadratic_form'.pos_def) : @@ -71,3 +73,31 @@ begin end end quadratic_form + +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 +{ inner := λ x y, dot_product (star x) (M.mul_vec y), + conj_sym := λ 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, + begin + by_cases h : x = 0, + { simp [h] }, + { exact le_of_lt (hM.2 x h) } + end, + definite := λ x hx, + begin + by_contra' h, + simpa [hx, lt_self_iff_false] 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] } + +end matrix From dc1a3f446b6b553f1365b7f1e2272f24ca689229 Mon Sep 17 00:00:00 2001 From: Reid Barton Date: Fri, 29 Jul 2022 08:40:25 +0000 Subject: [PATCH 136/173] docs(topology/sets/compacts): fix typo (#15739) --- src/topology/sets/compacts.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/sets/compacts.lean b/src/topology/sets/compacts.lean index 196399df39db6..cf0a8b8e63cef 100644 --- a/src/topology/sets/compacts.lean +++ b/src/topology/sets/compacts.lean @@ -178,8 +178,8 @@ end nonempty_compacts /-! ### Positive compact sets -/ -/-- The type of compact sets nonempty interior of a topological space. See also `compacts` and -`nonempty_compacts` -/ +/-- The type of compact sets with nonempty interior of a topological space. +See also `compacts` and `nonempty_compacts`. -/ structure positive_compacts (α : Type*) [topological_space α] extends compacts α := (interior_nonempty' : (interior carrier).nonempty) From afe51c20eb8dcce880e2129747494c01a86ff363 Mon Sep 17 00:00:00 2001 From: Reid Barton Date: Fri, 29 Jul 2022 08:40:27 +0000 Subject: [PATCH 137/173] docs(measure_theory/measure/haar): Mention uniqueness in module docstring (#15740) --- src/measure_theory/measure/haar.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index b0dd4792f8e1b..7f81eef302559 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -9,8 +9,8 @@ import measure_theory.group.prod /-! # Haar measure -In this file we prove the existence of Haar measure for a locally compact Hausdorff topological -group. +In this file we prove the existence and uniqueness (up to scalar multiples) of Haar measure +for a locally compact Hausdorff topological group. For the construction, we follow the write-up by Jonathan Gleason, *Existence and Uniqueness of Haar Measure*. @@ -46,6 +46,8 @@ where `ᵒ` denotes the interior. it is invariant and gives finite mass to compact sets and positive mass to nonempty open sets. * `haar` : some choice of a Haar measure, on a locally compact Hausdorff group, constructed as `haar_measure K` where `K` is some arbitrary choice of a compact set with nonempty interior. +* `haar_measure_unique`: Every σ-finite left invariant measure on a locally compact Hausdorff group + is a scalar multiple of the Haar measure. ## References * Paul Halmos (1950), Measure Theory, §53 From b19dc71240814653284e7c98e31bea7e54db33c2 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Fri, 29 Jul 2022 11:05:19 +0000 Subject: [PATCH 138/173] feat(data/nat/factorization/basic): golf & move `card_multiples`, prove variant lemma `Ioc_filter_dvd_card_eq_div` (#15277) Golfs `card_multiples` ("exactly `n / p` naturals in `[1, n]` are multiples of `p`") and moves it from `data/nat/prime` to `data/nat/factorization/basic`. Also proves a slightly more convenient variant, `Ioc_filter_dvd_card_eq_div`. --- .../81_sum_of_prime_reciprocals_diverges.lean | 2 +- src/data/nat/factorization/basic.lean | 25 +++++++++++++++++++ src/data/nat/prime.lean | 12 --------- 3 files changed, 26 insertions(+), 13 deletions(-) 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 682063586ae0a..58b8eefafd121 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 @@ -126,7 +126,7 @@ begin ≤ ∑ p in P, card (N p) : by assumption_mod_cast ... ≤ ∑ p in P, x * (1 / p) : sum_le_sum (λ p hp, _) ... = x * ∑ p in P, 1 / p : mul_sum.symm, - simp only [mul_one_div, N, sep_def, filter_congr_decidable, card_multiples, nat.cast_div_le], + simp only [mul_one_div, N, sep_def, filter_congr_decidable, nat.card_multiples, nat.cast_div_le], end /-- diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 85ea29bc5c874..8846a22839aef 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -651,4 +651,29 @@ begin simp [factorization_def n (prime_of_mem_factorization hp)] } end +/-! ### Lemmas about factorizations of particular functions -/ + +-- TODO: Port lemmas from `data/nat/multiplicity` to here, re-written in terms of `factorization` + +/-- Exactly `n / p` naturals in `[1, n]` are multiples of `p`. -/ +lemma card_multiples (n p : ℕ) : card ((finset.range n).filter (λ e, p ∣ e + 1)) = n / p := +begin + induction n with n hn, { simp }, + simp [nat.succ_div, add_ite, add_zero, finset.range_succ, filter_insert, apply_ite card, + card_insert_of_not_mem, hn], +end + +/-- Exactly `n / p` naturals in `(0, n]` are multiples of `p`. -/ +lemma Ioc_filter_dvd_card_eq_div (n p : ℕ) : + ((Ioc 0 n).filter (λ x, p ∣ x)).card = n / p := +begin + induction n with n IH, { simp }, + -- TODO: Golf away `h1` after Yaël PRs a lemma asserting this + have h1 : Ioc 0 n.succ = insert n.succ (Ioc 0 n), + { rcases n.eq_zero_or_pos with rfl | hn, { simp }, + simp_rw [←Ico_succ_succ, Ico_insert_right (succ_le_succ hn.le), Ico_succ_right] }, + simp [nat.succ_div, add_ite, add_zero, h1, filter_insert, apply_ite card, + card_insert_eq_ite, IH, finset.mem_filter, mem_Ioc, not_le.2 (lt_add_one n)], +end + end nat diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index cc703fa8ced2c..dab7a6f614ecd 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -1195,15 +1195,3 @@ namespace int lemma prime_two : prime (2 : ℤ) := nat.prime_iff_prime_int.mp nat.prime_two lemma prime_three : prime (3 : ℤ) := nat.prime_iff_prime_int.mp nat.prime_three end int - -section -open finset -/-- Exactly `n / p` naturals in `[1, n]` are multiples of `p`. -/ -lemma card_multiples (n p : ℕ) : card ((range n).filter (λ e, p ∣ e + 1)) = n / p := -begin - induction n with n hn, - { rw [nat.zero_div, range_zero, filter_empty, card_empty] }, - { rw [nat.succ_div, add_ite, add_zero, range_succ, filter_insert, apply_ite card, - card_insert_of_not_mem (mem_filter.not.mpr (not_and_of_not_left _ not_mem_range_self)), hn] } -end -end From 1406534bcd66081dd9610a3891355c9e411ca0ea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 29 Jul 2022 12:50:09 +0000 Subject: [PATCH 139/173] feat(data/finset/pointwise): Singleton arithmetic (#15435) Lemmas about pointwise operations with a singleton. Also make `s` explicit in `finset.card_inv`. --- src/data/finset/n_ary.lean | 35 ++++++++++++++++++++++++ src/data/finset/pointwise.lean | 49 ++++++++++++++++++++++++++-------- 2 files changed, 73 insertions(+), 11 deletions(-) diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 43dc62905ad72..174012d5c2e35 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -140,6 +140,41 @@ begin mem_insert_self _ _, ha⟩, h.trans $ image₂_subset (subset_insert _ _) $ subset_insert _ _⟩⟩, end +variables (s t) + +lemma card_image₂_singleton_left (hf : injective (f a)) : (image₂ f {a} t).card = t.card := +by rw [image₂_singleton_left, card_image_of_injective _ hf] + +lemma card_image₂_singleton_right (hf : injective (λ a, f a b)) : (image₂ f s {b}).card = s.card := +by rw [image₂_singleton_right, card_image_of_injective _ hf] + +lemma image₂_singleton_inter [decidable_eq β] (t₁ t₂ : finset β) (hf : injective (f a)) : + image₂ f {a} (t₁ ∩ t₂) = image₂ f {a} t₁ ∩ image₂ f {a} t₂ := +by simp_rw [image₂_singleton_left, image_inter _ _ hf] + +lemma image₂_inter_singleton [decidable_eq α] (s₁ s₂ : finset α) (hf : injective (λ a, f a b)) : + image₂ f (s₁ ∩ s₂) {b} = image₂ f s₁ {b} ∩ image₂ f s₂ {b} := +by simp_rw [image₂_singleton_right, image_inter _ _ hf] + +lemma card_le_card_image₂_left {s : finset α} (hs : s.nonempty) (hf : ∀ a, injective (f a)) : + t.card ≤ (image₂ f s t).card := +begin + obtain ⟨a, ha⟩ := hs, + rw ←card_image₂_singleton_left _ (hf a), + exact card_le_of_subset (image₂_subset_right $ singleton_subset_iff.2 ha), +end + +lemma card_le_card_image₂_right {t : finset β} (ht : t.nonempty) + (hf : ∀ b, injective (λ a, f a b)) : + s.card ≤ (image₂ f s t).card := +begin + obtain ⟨b, hb⟩ := ht, + rw ←card_image₂_singleton_right _ (hf b), + exact card_le_of_subset (image₂_subset_left $ singleton_subset_iff.2 hb), +end + +variables {s t} + lemma bUnion_image_left : s.bUnion (λ a, t.image $ f a) = image₂ f s t := coe_injective $ by { push_cast, exact set.Union_image_left _ } diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 1e7b4eee1df73..cd6fed1ef1145 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -124,10 +124,10 @@ end has_inv open_locale pointwise section has_involutive_inv -variables [decidable_eq α] [has_involutive_inv α] {s t : finset α} +variables [decidable_eq α] [has_involutive_inv α] (s : finset α) @[simp, norm_cast, to_additive] -lemma coe_inv (s : finset α) : ↑(s⁻¹) = (s : set α)⁻¹ := coe_image.trans set.image_inv +lemma coe_inv : ↑(s⁻¹) = (s : set α)⁻¹ := coe_image.trans set.image_inv @[simp, to_additive] lemma card_inv : s⁻¹.card = s.card := card_image_of_injective _ inv_injective @@ -178,9 +178,8 @@ image₂_nonempty_iff @[to_additive] lemma nonempty.of_mul_left : (s * t).nonempty → s.nonempty := nonempty.of_image₂_left @[to_additive] lemma nonempty.of_mul_right : (s * t).nonempty → t.nonempty := nonempty.of_image₂_right -@[simp, to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right -@[simp, to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := -image₂_singleton_left +@[to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right +@[to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := image₂_singleton_left @[simp, to_additive] lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} := image₂_singleton @@ -643,9 +642,9 @@ image₂_nonempty_iff nonempty.of_image₂_left @[to_additive] lemma nonempty.of_smul_right : (s • t).nonempty → t.nonempty := nonempty.of_image₂_right -@[simp, to_additive] lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := +@[to_additive] lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := image₂_singleton_right -@[simp, to_additive] lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := +@[to_additive] lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := image₂_singleton_left @[simp, to_additive] lemma singleton_smul_singleton (a : α) (b : β) : ({a} : finset α) • ({b} : finset β) = {a • b} := @@ -709,8 +708,7 @@ lemma nonempty.of_vsub_left : (s -ᵥ t : finset α).nonempty → s.nonempty := lemma nonempty.of_vsub_right : (s -ᵥ t : finset α).nonempty → t.nonempty := nonempty.of_image₂_right @[simp] lemma vsub_singleton (b : β) : s -ᵥ ({b} : finset β) = s.image (-ᵥ b) := image₂_singleton_right -@[simp] lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) := -image₂_singleton_left +lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) := image₂_singleton_left @[simp] lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton @@ -888,11 +886,40 @@ coe_injective.no_zero_smul_divisors _ coe_zero coe_smul_finset end instances -@[to_additive] lemma pairwise_disjoint_smul_iff [decidable_eq α] [left_cancel_semigroup α] - {s : set α} {t : finset α} : +section left_cancel_semigroup +variables [left_cancel_semigroup α] [decidable_eq α] (s t : finset α) (a : α) + +@[to_additive] lemma pairwise_disjoint_smul_iff {s : set α} {t : finset α} : s.pairwise_disjoint (• t) ↔ ((s : set α) ×ˢ (t : set α) : set (α × α)).inj_on (λ p, p.1 * p.2) := by simp_rw [←pairwise_disjoint_coe, coe_smul_finset, set.pairwise_disjoint_smul_iff] +@[simp, to_additive] lemma card_singleton_mul : ({a} * t).card = t.card := +card_image₂_singleton_left _ $ mul_right_injective _ + +@[to_additive] lemma singleton_mul_inter : {a} * (s ∩ t) = ({a} * s) ∩ ({a} * t) := +image₂_singleton_inter _ _ $ mul_right_injective _ + +@[to_additive] lemma card_le_card_mul_left {s : finset α} (hs : s.nonempty) : + t.card ≤ (s * t).card := +card_le_card_image₂_left _ hs mul_right_injective + +end left_cancel_semigroup + +section +variables [right_cancel_semigroup α] [decidable_eq α] (s t : finset α) (a : α) + +@[simp, to_additive] lemma card_mul_singleton : (s * {a}).card = s.card := +card_image₂_singleton_right _ $ mul_left_injective _ + +@[to_additive] lemma inter_mul_singleton : (s ∩ t) * {a} = (s * {a}) ∩ (t * {a}) := +image₂_inter_singleton _ _ $ mul_left_injective _ + +@[to_additive] lemma card_le_card_mul_right {t : finset α} (ht : t.nonempty) : + s.card ≤ (s * t).card := +card_le_card_image₂_right _ ht mul_left_injective + +end + open_locale pointwise section group From 4ed946b338f018ee32a56690639ab045388913f8 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 29 Jul 2022 14:31:24 +0000 Subject: [PATCH 140/173] feat(algebra/direct_sum/decomposition): add an induction principle for `direct_sum.decomposition` class (#15654) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `direct_sum.decomposition M` and `p : M → Prop`, then to prove `p m` for an arbitrary `m`, it suffices to prove `p 0` and `p x` for homogeneous x and `p` being preserved by add. --- src/algebra/direct_sum/decomposition.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/algebra/direct_sum/decomposition.lean b/src/algebra/direct_sum/decomposition.lean index e302cc8279f8f..359f406bbb84d 100644 --- a/src/algebra/direct_sum/decomposition.lean +++ b/src/algebra/direct_sum/decomposition.lean @@ -72,6 +72,22 @@ def decompose : M ≃ ⨁ i, ℳ i := left_inv := decomposition.left_inv, right_inv := decomposition.right_inv } +protected lemma decomposition.induction_on {p : M → Prop} + (h_zero : p 0) (h_homogeneous : ∀ {i} (m : ℳ i), p (m : M)) + (h_add : ∀ (m m' : M), p m → p m' → p (m + m')) : ∀ m, p m := +begin + let ℳ' : ι → add_submonoid M := + λ i, (⟨ℳ i, λ _ _, add_mem_class.add_mem, zero_mem_class.zero_mem _⟩ : add_submonoid M), + haveI t : direct_sum.decomposition ℳ' := + { decompose' := direct_sum.decompose ℳ, + left_inv := λ _, (decompose ℳ).left_inv _, + right_inv := λ _, (decompose ℳ).right_inv _, }, + have mem : ∀ m, m ∈ supr ℳ' := + λ m, (direct_sum.is_internal.add_submonoid_supr_eq_top ℳ' + (decomposition.is_internal ℳ')).symm ▸ trivial, + exact λ m, add_submonoid.supr_induction ℳ' (mem m) (λ i m h, h_homogeneous ⟨m, h⟩) h_zero h_add, +end + @[simp] lemma decomposition.decompose'_eq : decomposition.decompose' = decompose ℳ := rfl @[simp] lemma decompose_symm_of {i : ι} (x : ℳ i) : From 7bc35f94d622121c00f35be1b238c9cf518afacb Mon Sep 17 00:00:00 2001 From: tb65536 Date: Fri, 29 Jul 2022 14:31:26 +0000 Subject: [PATCH 141/173] feat(field_theory/intermediate_field): Produce an intermediate field from a subalgebra satisfying `is_field`. (#15659) This construction came up when proving that if `p` splits in `L/K`, then `p.is_splitting_field K (adjoin K (p.root_set L))`. --- src/field_theory/intermediate_field.lean | 27 ++++++++++++++++++++---- 1 file changed, 23 insertions(+), 4 deletions(-) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index afa1a1ac2bf71..f2933d2f4e68f 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -190,11 +190,30 @@ def subalgebra.to_intermediate_field (S : subalgebra K L) (inv_mem : ∀ x ∈ S (S.to_intermediate_field inv_mem).to_subalgebra = S := by { ext, refl } -@[simp] lemma to_intermediate_field_to_subalgebra - (S : intermediate_field K L) (inv_mem : ∀ x ∈ S.to_subalgebra, x⁻¹ ∈ S) : - S.to_subalgebra.to_intermediate_field inv_mem = S := +@[simp] lemma to_intermediate_field_to_subalgebra (S : intermediate_field K L) : + S.to_subalgebra.to_intermediate_field (λ x, S.inv_mem) = S := by { ext, refl } +/-- Turn a subalgebra satisfying `is_field` into an intermediate_field -/ +def subalgebra.to_intermediate_field' (S : subalgebra K L) (hS : is_field S) : + intermediate_field K L := +S.to_intermediate_field $ λ x hx, begin + by_cases hx0 : x = 0, + { rw [hx0, inv_zero], + exact S.zero_mem }, + letI hS' := hS.to_field, + obtain ⟨y, hy⟩ := hS.mul_inv_cancel (show (⟨x, hx⟩ : S) ≠ 0, from subtype.ne_of_val_ne hx0), + rw [subtype.ext_iff, S.coe_mul, S.coe_one, subtype.coe_mk, mul_eq_one_iff_inv_eq₀ hx0] at hy, + exact hy.symm ▸ y.2, +end + +@[simp] lemma to_subalgebra_to_intermediate_field' (S : subalgebra K L) (hS : is_field S) : + (S.to_intermediate_field' hS).to_subalgebra = S := +by { ext, refl } + +@[simp] lemma to_intermediate_field'_to_subalgebra (S : intermediate_field K L) : + S.to_subalgebra.to_intermediate_field' (field.to_is_field S) = S := +by { ext, refl } /-- Turn a subfield of `L` containing the image of `K` into an intermediate field -/ def subfield.to_intermediate_field (S : subfield L) @@ -476,7 +495,7 @@ def subalgebra_equiv_intermediate_field (alg : algebra.is_algebraic K L) : { to_fun := λ S, S.to_intermediate_field (λ x hx, S.inv_mem_of_algebraic (alg (⟨x, hx⟩ : S))), inv_fun := λ S, S.to_subalgebra, left_inv := λ S, to_subalgebra_to_intermediate_field _ _, - right_inv := λ S, to_intermediate_field_to_subalgebra _ _, + right_inv := to_intermediate_field_to_subalgebra, map_rel_iff' := λ S S', iff.rfl } @[simp] lemma mem_subalgebra_equiv_intermediate_field (alg : algebra.is_algebraic K L) From dd55a7b2d778ff61b87db729b9b92ff0b58bc99b Mon Sep 17 00:00:00 2001 From: tb65536 Date: Fri, 29 Jul 2022 14:31:27 +0000 Subject: [PATCH 142/173] feat(field_theory/adjoin): A compositum of algebraic extensions is algebraic (#15724) This PR proves that a compositum of algebraic extensions is algebraic. Coming soon to a mathlib near you: A compositum of normal extensions is normal. --- src/field_theory/adjoin.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index aa443be1eba38..d890a38fb2cde 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -455,6 +455,11 @@ begin exact this hx, end +lemma exists_finset_of_mem_supr' {ι : Type*} {f : ι → intermediate_field F E} + {x : E} (hx : x ∈ ⨆ i, f i) : ∃ s : finset (Σ i, f i), x ∈ ⨆ i ∈ s, F⟮(i.2 : E)⟯ := +exists_finset_of_mem_supr (set_like.le_def.mp (supr_le + (λ i x h, set_like.le_def.mp (le_supr_of_le ⟨i, x, h⟩ le_rfl) (mem_adjoin_simple_self F x))) hx) + end adjoin_simple end adjoin_def @@ -945,6 +950,18 @@ begin exact this.symm ▸ intermediate_field.finite_dimensional_supr_of_finite, end +lemma is_algebraic_supr {ι : Type*} {f : ι → intermediate_field K L} + (h : ∀ i, algebra.is_algebraic K (f i)) : + algebra.is_algebraic K (⨆ i, f i : intermediate_field K L) := +begin + rintros ⟨x, hx⟩, + obtain ⟨s, hx⟩ := exists_finset_of_mem_supr' hx, + rw [algebraic_iff, subtype.coe_mk, ←subtype.coe_mk x hx, ←algebraic_iff], + haveI : ∀ i : (Σ i, f i), finite_dimensional K K⟮(i.2 : L)⟯ := + λ ⟨i, x⟩, adjoin.finite_dimensional (is_algebraic_iff_is_integral.mp (algebraic_iff.mp (h i x))), + apply algebra.is_algebraic_of_finite, +end + end supremum end intermediate_field From 0fbd1ebfff681bad1fb4dffa9cb228fb2a768121 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 29 Jul 2022 15:19:12 +0000 Subject: [PATCH 143/173] fix(linear_algebra/prod): add missing `of_dual`s in lemma statements (#15750) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Without these the lemmas contain nonsense statements of the form `(x : α) ≤ (y : αᵒᵈ)`. --- src/linear_algebra/prod.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index cfb0e6a6e6183..65aadf802d91f 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -686,7 +686,7 @@ Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of all isomorphic to `M`. -/ def tunnel (f : M × N →ₗ[R] M) (i : injective f) : ℕ →o (submodule R M)ᵒᵈ := -⟨λ n, (tunnel' f i n).1, monotone_nat_of_le_succ (λ n, begin +⟨λ n, order_dual.to_dual (tunnel' f i n).1, monotone_nat_of_le_succ (λ n, begin dsimp [tunnel', tunnel_aux], rw [submodule.map_comp, submodule.map_comp], apply submodule.map_subtype_le, @@ -705,7 +705,7 @@ def tailing_linear_equiv (f : M × N →ₗ[R] M) (i : injective f) (n : ℕ) : (tunnel_aux_injective f i (tunnel' f i n))).symm.trans (submodule.snd_equiv R M N) lemma tailing_le_tunnel (f : M × N →ₗ[R] M) (i : injective f) (n : ℕ) : - tailing f i n ≤ tunnel f i n := + tailing f i n ≤ (tunnel f i n).of_dual := begin dsimp [tailing, tunnel_aux], rw [submodule.map_comp, submodule.map_comp], @@ -713,7 +713,7 @@ begin end lemma tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : injective f) (n : ℕ) : - disjoint (tailing f i n) (tunnel f i (n+1)) := + disjoint (tailing f i n) (tunnel f i (n+1)).of_dual := begin rw disjoint_iff, dsimp [tailing, tunnel, tunnel'], @@ -723,7 +723,7 @@ begin end lemma tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : injective f) (n : ℕ) : - tailing f i n ⊔ tunnel f i (n+1) ≤ tunnel f i n := + tailing f i n ⊔ (tunnel f i (n+1)).of_dual ≤ (tunnel f i n).of_dual := begin dsimp [tailing, tunnel, tunnel', tunnel_aux], rw [←submodule.map_sup, sup_comm, submodule.fst_sup_snd, submodule.map_comp, submodule.map_comp], @@ -743,7 +743,7 @@ by simp [tailings] by simp [tailings] lemma tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : injective f) (n : ℕ) : - disjoint (tailings f i n) (tunnel f i (n+1)) := + disjoint (tailings f i n) (tunnel f i (n+1)).of_dual := begin induction n with n ih, { simp only [tailings_zero], From 883a3e3ef83426d44bc523108f82a2c84882f657 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 29 Jul 2022 18:28:34 +0000 Subject: [PATCH 144/173] refactor(algebra/homology): better defeqs, less case splitting (#15690) --- src/algebra/homology/additive.lean | 2 - src/algebra/homology/complex_shape.lean | 29 +- src/algebra/homology/homological_complex.lean | 201 ++++-------- src/algebra/homology/homology.lean | 13 +- src/algebra/homology/homotopy.lean | 289 ++++++------------ src/algebra/homology/homotopy_category.lean | 6 +- src/category_theory/functor/left_derived.lean | 7 +- 7 files changed, 180 insertions(+), 367 deletions(-) diff --git a/src/algebra/homology/additive.lean b/src/algebra/homology/additive.lean index 19b104a07db11..072555a1eabce 100644 --- a/src/algebra/homology/additive.lean +++ b/src/algebra/homology/additive.lean @@ -67,8 +67,6 @@ namespace homological_complex instance eval_additive (i : ι) : (eval V c i).additive := {} -variables [has_zero_object V] - instance cycles_additive [has_equalizers V] : (cycles_functor V c i).additive := {} variables [has_images V] [has_image_maps V] diff --git a/src/algebra/homology/complex_shape.lean b/src/algebra/homology/complex_shape.lean index 1a31f498d7581..a2287f28e970d 100644 --- a/src/algebra/homology/complex_shape.lean +++ b/src/algebra/homology/complex_shape.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Scott Morrison -/ import algebra.group.defs -import data.option.basic import logic.relation /-! @@ -23,8 +22,8 @@ we only allow nonzero differentials `d i j` from `i` to `j` if `c.rel i j`. Further, we require that `{ j // c.rel i j }` and `{ i // c.rel i j }` are subsingletons. This means that the shape consists of some union of lines, rays, intervals, and circles. -Convenience functions `c.next` and `c.prev` provide, as an `option`, these related elements -when they exist. +Convenience functions `c.next` and `c.prev` provide these related elements +when they exist, and return their input otherwise. This design aims to avoid certain problems arising from dependent type theory. In particular we never have to ensure morphisms `d i : X i ⟶ X (succ i)` compose as @@ -54,7 +53,7 @@ and we will only allow a non-zero differential from `i` to `j` when `rel i j`. There are axioms which imply `{ j // c.rel i j }` and `{ i // c.rel i j }` are subsingletons. This means that the shape consists of some union of lines, rays, intervals, and circles. -Below we define `c.next` and `c.prev` which provide, as an `option`, these related elements. +Below we define `c.next` and `c.prev` which provide these related elements. -/ @[ext, nolint has_inhabited_instance] structure complex_shape (ι : Type*) := @@ -131,22 +130,24 @@ begin end /-- -An option-valued arbitary choice of index `j` such that `rel i j`, if such exists. +An arbitary choice of index `j` such that `rel i j`, if such exists. +Returns `i` otherwise. -/ -def next (c : complex_shape ι) (i : ι) : option { j // c.rel i j } := -option.choice _ +def next (c : complex_shape ι) (i : ι) : ι := +if h : ∃ j, c.rel i j then h.some else i /-- -An option-valued arbitary choice of index `i` such that `rel i j`, if such exists. +An arbitary choice of index `i` such that `rel i j`, if such exists. +Returns `j` otherwise. -/ -def prev (c : complex_shape ι) (j : ι) : option { i // c.rel i j } := -option.choice _ +def prev (c : complex_shape ι) (j : ι) : ι := +if h : ∃ i, c.rel i j then h.some else j -lemma next_eq_some (c : complex_shape ι) {i j : ι} (h : c.rel i j) : c.next i = some ⟨j, h⟩ := -option.choice_eq _ +lemma next_eq' (c : complex_shape ι) {i j : ι} (h : c.rel i j) : c.next i = j := +by { apply c.next_eq _ h, dsimp only [next], rw dif_pos, exact Exists.some_spec ⟨j, h⟩, } -lemma prev_eq_some (c : complex_shape ι) {i j : ι} (h : c.rel i j) : c.prev j = some ⟨i, h⟩ := -option.choice_eq _ +lemma prev_eq' (c : complex_shape ι) {i j : ι} (h : c.rel i j) : c.prev j = i := +by { apply c.prev_eq _ h, dsimp only [prev], rw dif_pos, exact Exists.some_spec ⟨i, h⟩, } /-- The `complex_shape` allowing differentials from `X i` to `X (i+a)`. diff --git a/src/algebra/homology/homological_complex.lean b/src/algebra/homology/homological_complex.lean index 7c4851e060cc3..8e81c26ccc202 100644 --- a/src/algebra/homology/homological_complex.lean +++ b/src/algebra/homology/homological_complex.lean @@ -108,40 +108,40 @@ homological_complex V (complex_shape.up α) namespace chain_complex @[simp] lemma prev (α : Type*) [add_right_cancel_semigroup α] [has_one α] (i : α) : - (complex_shape.down α).prev i = some ⟨i+1, rfl⟩ := -option.choice_eq _ + (complex_shape.down α).prev i = i+1 := +(complex_shape.down α).prev_eq' rfl @[simp] lemma next (α : Type*) [add_group α] [has_one α] (i : α) : - (complex_shape.down α).next i = some ⟨i-1, sub_add_cancel i 1⟩ := -option.choice_eq _ + (complex_shape.down α).next i = i-1 := +(complex_shape.down α).next_eq' $ sub_add_cancel _ _ @[simp] lemma next_nat_zero : - (complex_shape.down ℕ).next 0 = none := -@option.choice_eq_none _ ⟨by rintro ⟨j, ⟨⟩⟩⟩ + (complex_shape.down ℕ).next 0 = 0 := +by { classical, refine dif_neg _, push_neg, intro, apply nat.no_confusion } @[simp] lemma next_nat_succ (i : ℕ) : - (complex_shape.down ℕ).next (i+1) = some ⟨i, rfl⟩ := -option.choice_eq _ + (complex_shape.down ℕ).next (i+1) = i := +(complex_shape.down ℕ).next_eq' rfl end chain_complex namespace cochain_complex @[simp] lemma prev (α : Type*) [add_group α] [has_one α] (i : α) : - (complex_shape.up α).prev i = some ⟨i-1, sub_add_cancel i 1⟩ := -option.choice_eq _ + (complex_shape.up α).prev i = i-1 := +(complex_shape.up α).prev_eq' $ sub_add_cancel _ _ @[simp] lemma next (α : Type*) [add_right_cancel_semigroup α] [has_one α] (i : α) : - (complex_shape.up α).next i = some ⟨i+1, rfl⟩ := -option.choice_eq _ + (complex_shape.up α).next i = i+1 := +(complex_shape.up α).next_eq' rfl @[simp] lemma prev_nat_zero : - (complex_shape.up ℕ).prev 0 = none := -@option.choice_eq_none _ ⟨by rintro ⟨j, ⟨⟩⟩⟩ + (complex_shape.up ℕ).prev 0 = 0 := +by { classical, refine dif_neg _, push_neg, intro, apply nat.no_confusion } @[simp] lemma prev_nat_succ (i : ℕ) : - (complex_shape.up ℕ).prev (i+1) = some ⟨i, rfl⟩ := -option.choice_eq _ + (complex_shape.up ℕ).prev (i+1) = i := +(complex_shape.up ℕ).prev_eq' rfl end cochain_complex @@ -290,114 +290,81 @@ end section -variables [has_zero_object V] -open_locale zero_object - -/-- Either `C.X i`, if there is some `i` with `c.rel i j`, or the zero object. -/ -def X_prev (j : ι) : V := -match c.prev j with -| none := 0 -| (some ⟨i,_⟩) := C.X i -end +/-- Either `C.X i`, if there is some `i` with `c.rel i j`, or `C.X j`. -/ +abbreviation X_prev (j : ι) : V := C.X (c.prev j) /-- If `c.rel i j`, then `C.X_prev j` is isomorphic to `C.X i`. -/ def X_prev_iso {i j : ι} (r : c.rel i j) : C.X_prev j ≅ C.X i := -eq_to_iso begin - dsimp [X_prev], - rw c.prev_eq_some r, - refl, -end +eq_to_iso $ by rw ← c.prev_eq' r -/-- If there is no `i` so `c.rel i j`, then `C.X_prev j` is isomorphic to `0`. -/ -def X_prev_iso_zero {j : ι} (h : c.prev j = none) : - C.X_prev j ≅ 0 := -eq_to_iso begin - dsimp [X_prev], - rw h, - refl, +/-- If there is no `i` so `c.rel i j`, then `C.X_prev j` is isomorphic to `C.X j`. -/ +def X_prev_iso_zero {j : ι} (h : ¬c.rel (c.prev j) j) : + C.X_prev j ≅ C.X j := +eq_to_iso $ congr_arg C.X begin + dsimp [complex_shape.prev], + rw dif_neg, push_neg, intros i hi, + have : c.prev j = i := c.prev_eq' hi, + rw this at h, contradiction, end -/-- Either `C.X j`, if there is some `j` with `c.rel i j`, or the zero object. -/ -def X_next (i : ι) : V := -match c.next i with -| none := 0 -| (some ⟨j,_⟩) := C.X j -end +/-- Either `C.X j`, if there is some `j` with `c.rel i j`, or `C.X j`. -/ +abbreviation X_next (i : ι) : V := C.X (c.next i) /-- If `c.rel i j`, then `C.X_next i` is isomorphic to `C.X j`. -/ def X_next_iso {i j : ι} (r : c.rel i j) : C.X_next i ≅ C.X j := -eq_to_iso begin - dsimp [X_next], - rw c.next_eq_some r, - refl, -end +eq_to_iso $ by rw ← c.next_eq' r /-- If there is no `j` so `c.rel i j`, then `C.X_next i` is isomorphic to `0`. -/ -def X_next_iso_zero {i : ι} (h : c.next i = none) : - C.X_next i ≅ 0 := -eq_to_iso begin - dsimp [X_next], - rw h, - refl, +def X_next_iso_zero {i : ι} (h : ¬c.rel i (c.next i)) : + C.X_next i ≅ C.X i := +eq_to_iso $ congr_arg C.X begin + dsimp [complex_shape.next], + rw dif_neg, rintro ⟨j, hj⟩, + have : c.next i = j := c.next_eq' hj, + rw this at h, contradiction, end /-- The differential mapping into `C.X j`, or zero if there isn't one. -/ -def d_to (j : ι) : C.X_prev j ⟶ C.X j := -match c.prev j with -| none := (0 : C.X_prev j ⟶ C.X j) -| (some ⟨i, w⟩) := (C.X_prev_iso w).hom ≫ C.d i j -end +abbreviation d_to (j : ι) : C.X_prev j ⟶ C.X j := C.d (c.prev j) j /-- The differential mapping out of `C.X i`, or zero if there isn't one. -/ -def d_from (i : ι) : C.X i ⟶ C.X_next i := -match c.next i with -| none := (0 : C.X i ⟶ C.X_next i) -| (some ⟨j, w⟩) := C.d i j ≫ (C.X_next_iso w).inv -end +abbreviation d_from (i : ι) : C.X i ⟶ C.X_next i := C.d i (c.next i) lemma d_to_eq {i j : ι} (r : c.rel i j) : C.d_to j = (C.X_prev_iso r).hom ≫ C.d i j := begin - dsimp [d_to, X_prev_iso], - rw c.prev_eq_some r, - refl, + obtain rfl := c.prev_eq' r, + exact (category.id_comp _).symm, end @[simp] -lemma d_to_eq_zero {j : ι} (h : c.prev j = none) : +lemma d_to_eq_zero {j : ι} (h : ¬c.rel (c.prev j) j) : C.d_to j = 0 := -begin - dsimp [d_to], - rw h, refl, -end +C.shape _ _ h lemma d_from_eq {i j : ι} (r : c.rel i j) : C.d_from i = C.d i j ≫ (C.X_next_iso r).inv := begin - dsimp [d_from, X_next_iso], - rw c.next_eq_some r, - refl, + obtain rfl := c.next_eq' r, + exact (category.comp_id _).symm, end @[simp] -lemma d_from_eq_zero {i : ι} (h : c.next i = none) : +lemma d_from_eq_zero {i : ι} (h : ¬c.rel i (c.next i)) : C.d_from i = 0 := -begin - dsimp [d_from], - rw h, refl, -end +C.shape _ _ h @[simp, reassoc] lemma X_prev_iso_comp_d_to {i j : ι} (r : c.rel i j) : (C.X_prev_iso r).inv ≫ C.d_to j = C.d i j := by simp [C.d_to_eq r] -@[simp, reassoc] lemma X_prev_iso_zero_comp_d_to {j : ι} (h : c.prev j = none) : +@[simp, reassoc] lemma X_prev_iso_zero_comp_d_to {j : ι} (h : ¬c.rel (c.prev j) j) : (C.X_prev_iso_zero h).inv ≫ C.d_to j = 0 := by simp [h] @@ -405,20 +372,13 @@ by simp [h] C.d_from i ≫ (C.X_next_iso r).hom = C.d i j := by simp [C.d_from_eq r] -@[simp, reassoc] lemma d_from_comp_X_next_iso_zero {i : ι} (h : c.next i = none) : +@[simp, reassoc] lemma d_from_comp_X_next_iso_zero {i : ι} (h : ¬c.rel i (c.next i)) : C.d_from i ≫ (C.X_next_iso_zero h).hom = 0 := by simp [h] @[simp] lemma d_to_comp_d_from (j : ι) : C.d_to j ≫ C.d_from j = 0 := -begin - rcases h₁ : c.next j with _ | ⟨k,w₁⟩, - { rw [d_from_eq_zero _ h₁], simp }, - { rw [d_from_eq _ w₁], - rcases h₂ : c.prev j with _ | ⟨i,w₂⟩, - { rw [d_to_eq_zero _ h₂], simp }, - { rw [d_to_eq _ w₂], simp } } -end +C.d_comp_d _ _ _ lemma kernel_from_eq_kernel [has_kernels V] {i j : ι} (r : c.rel i j) : kernel_subobject (C.d_from i) = kernel_subobject (C.d i j) := @@ -468,58 +428,37 @@ def iso_of_components (f : Π i, C₁.X i ≅ C₂.X i) iso_app (iso_of_components f hf) i = f i := by { ext, simp, } -variables [has_zero_object V] -open_locale zero_object - /-! Lemmas relating chain maps and `d_to`/`d_from`. -/ /-- `f.prev j` is `f.f i` if there is some `r i j`, and zero otherwise. -/ -def prev (f : hom C₁ C₂) (j : ι) : C₁.X_prev j ⟶ C₂.X_prev j := -match c.prev j with -| none := 0 -| some ⟨i,w⟩ := (C₁.X_prev_iso w).hom ≫ f.f i ≫ (C₂.X_prev_iso w).inv -end +abbreviation prev (f : hom C₁ C₂) (j : ι) : C₁.X_prev j ⟶ C₂.X_prev j := f.f _ lemma prev_eq (f : hom C₁ C₂) {i j : ι} (w : c.rel i j) : f.prev j = (C₁.X_prev_iso w).hom ≫ f.f i ≫ (C₂.X_prev_iso w).inv := begin - dsimp [prev], - rw c.prev_eq_some w, - refl, + obtain rfl := c.prev_eq' w, + simp only [X_prev_iso, eq_to_iso_refl, iso.refl_hom, iso.refl_inv, id_comp, comp_id], end /-- `f.next i` is `f.f j` if there is some `r i j`, and zero otherwise. -/ -def next (f : hom C₁ C₂) (i : ι) : C₁.X_next i ⟶ C₂.X_next i := -match c.next i with -| none := 0 -| some ⟨j,w⟩ := (C₁.X_next_iso w).hom ≫ f.f j ≫ (C₂.X_next_iso w).inv -end +abbreviation next (f : hom C₁ C₂) (i : ι) : C₁.X_next i ⟶ C₂.X_next i := f.f _ lemma next_eq (f : hom C₁ C₂) {i j : ι} (w : c.rel i j) : f.next i = (C₁.X_next_iso w).hom ≫ f.f j ≫ (C₂.X_next_iso w).inv := begin - dsimp [next], - rw c.next_eq_some w, - refl, + obtain rfl := c.next_eq' w, + simp only [X_next_iso, eq_to_iso_refl, iso.refl_hom, iso.refl_inv, id_comp, comp_id], end @[simp, reassoc, elementwise] lemma comm_from (f : hom C₁ C₂) (i : ι) : f.f i ≫ C₂.d_from i = C₁.d_from i ≫ f.next i := -begin - rcases h : c.next i with _ | ⟨j,w⟩, - { simp [h] }, - { simp [d_from_eq _ w, next_eq _ w] } -end +f.comm _ _ @[simp, reassoc, elementwise] lemma comm_to (f : hom C₁ C₂) (j : ι) : f.prev j ≫ C₂.d_to j = C₁.d_to j ≫ f.f j := -begin - rcases h : c.prev j with _ | ⟨j,w⟩, - { simp [h] }, - { simp [d_to_eq _ w, prev_eq _ w] } -end +f.comm _ _ /-- A morphism of chain complexes @@ -531,30 +470,10 @@ arrow.hom_mk (f.comm_from i) @[simp] lemma sq_from_left (f : hom C₁ C₂) (i : ι) : (f.sq_from i).left = f.f i := rfl @[simp] lemma sq_from_right (f : hom C₁ C₂) (i : ι) : (f.sq_from i).right = f.next i := rfl -@[simp] lemma sq_from_id (C₁ : homological_complex V c) (i : ι) : sq_from (𝟙 C₁) i = 𝟙 _ := -begin - rcases h : c.next i with _ | ⟨j,w⟩, - { ext, - { refl }, - { dsimp, simp only [next, h], - symmetry, - apply zero_of_target_iso_zero, - exact X_next_iso_zero _ h } }, - { ext, refl, dsimp, simp [next, h] } -end +@[simp] lemma sq_from_id (C₁ : homological_complex V c) (i : ι) : sq_from (𝟙 C₁) i = 𝟙 _ := rfl @[simp] lemma sq_from_comp (f : C₁ ⟶ C₂) (g : C₂ ⟶ C₃) (i : ι) : - sq_from (f ≫ g) i = sq_from f i ≫ sq_from g i := -begin - rcases h : c.next i with _ | ⟨j,w⟩, - { ext, - { refl }, - { dsimp, simp only [next, h], - symmetry, - apply zero_of_target_iso_zero, - exact X_next_iso_zero _ h } }, - { ext, refl, dsimp, simp [next, h] } -end + sq_from (f ≫ g) i = sq_from f i ≫ sq_from g i := rfl /-- A morphism of chain complexes diff --git a/src/algebra/homology/homology.lean b/src/algebra/homology/homology.lean index 786a74f5d3a9c..e53bd485edeac 100644 --- a/src/algebra/homology/homology.lean +++ b/src/algebra/homology/homology.lean @@ -33,8 +33,6 @@ noncomputable theory namespace homological_complex -variables [has_zero_object V] - section cycles variables [has_kernels V] @@ -55,7 +53,7 @@ def cycles_iso_kernel {i j : ι} (r : c.rel i j) : subobject.iso_of_eq _ _ (C.cycles_eq_kernel_subobject r) ≪≫ kernel_subobject_iso (C.d i j) -lemma cycles_eq_top {i} (h : c.next i = none) : C.cycles i = ⊤ := +lemma cycles_eq_top {i} (h : ¬c.rel i (c.next i)) : C.cycles i = ⊤ := begin rw eq_top_iff, apply le_kernel_subobject, @@ -84,7 +82,8 @@ def boundaries_iso_image [has_equalizers V] {i j : ι} (r : c.rel i j) : subobject.iso_of_eq _ _ (C.boundaries_eq_image_subobject r) ≪≫ image_subobject_iso (C.d i j) -lemma boundaries_eq_bot {j} (h : c.prev j = none) : C.boundaries j = ⊥ := +lemma boundaries_eq_bot [has_zero_object V] {j} (h : ¬c.rel (c.prev j) j) : + C.boundaries j = ⊥ := begin rw eq_bot_iff, refine image_subobject_le _ 0 _, @@ -128,7 +127,7 @@ open homological_complex /-! Computing the cycles is functorial. -/ section -variables [has_zero_object V] [has_kernels V] +variables [has_kernels V] variables {C₁ C₂ C₃ : homological_complex V c} (f : C₁ ⟶ C₂) /-- @@ -161,7 +160,7 @@ end /-! Computing the boundaries is functorial. -/ section -variables [has_zero_object V] [has_images V] [has_image_maps V] +variables [has_images V] [has_image_maps V] variables {C₁ C₂ C₃ : homological_complex V c} (f : C₁ ⟶ C₂) /-- @@ -183,7 +182,7 @@ end section /-! The `boundaries_to_cycles` morphisms are natural. -/ -variables [has_zero_object V] [has_equalizers V] [has_images V] [has_image_maps V] +variables [has_equalizers V] [has_images V] [has_image_maps V] variables {C₁ C₂ : homological_complex V c} (f : C₁ ⟶ C₂) @[simp, reassoc] diff --git a/src/algebra/homology/homotopy.lean b/src/algebra/homology/homotopy.lean index 5d37a03874969..3ad0100890e27 100644 --- a/src/algebra/homology/homotopy.lean +++ b/src/algebra/homology/homotopy.lean @@ -30,157 +30,83 @@ section /-- The composition of `C.d i i' ≫ f i' i` if there is some `i'` coming after `i`, and `0` otherwise. -/ def d_next (i : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X i ⟶ D.X i) := -add_monoid_hom.mk' (λ f, match c.next i with -| none := 0 -| some ⟨i',w⟩ := C.d i i' ≫ f i' i -end) -begin - intros f g, - rcases c.next i with _|⟨i',w⟩, - exact (zero_add _).symm, - exact preadditive.comp_add _ _ _ _ _ _, -end +add_monoid_hom.mk' (λ f, C.d i (c.next i) ≫ f (c.next i) i) $ +λ f g, preadditive.comp_add _ _ _ _ _ _ /-- `f i' i` if `i'` comes after `i`, and 0 if there's no such `i'`. Hopefully there won't be much need for this, except in `d_next_eq_d_from_from_next` to see that `d_next` factors through `C.d_from i`. -/ -def from_next [has_zero_object V] (i : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X_next i ⟶ D.X i) := -add_monoid_hom.mk' (λ f, match c.next i with -| none := 0 -| some ⟨i',w⟩ := (C.X_next_iso w).hom ≫ f i' i -end) -begin - intros f g, - rcases c.next i with _|⟨i',w⟩, - exact (zero_add _).symm, - exact preadditive.comp_add _ _ _ _ _ _, -end +def from_next (i : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X_next i ⟶ D.X i) := +add_monoid_hom.mk' (λ f, f (c.next i) i) $ λ f g, rfl @[simp] -lemma d_next_eq_d_from_from_next [has_zero_object V] (f : Π i j, C.X i ⟶ D.X j) (i : ι) : - d_next i f = C.d_from i ≫ from_next i f := -begin - dsimp [d_next, from_next], - rcases c.next i with ⟨⟩|⟨⟨i', w⟩⟩; - { dsimp [d_next, from_next], simp }, -end +lemma d_next_eq_d_from_from_next (f : Π i j, C.X i ⟶ D.X j) (i : ι) : + d_next i f = C.d_from i ≫ from_next i f := rfl lemma d_next_eq (f : Π i j, C.X i ⟶ D.X j) {i i' : ι} (w : c.rel i i') : d_next i f = C.d i i' ≫ f i' i := -begin - dsimp [d_next], - rw c.next_eq_some w, - refl, -end +by { obtain rfl := c.next_eq' w, refl } @[simp] lemma d_next_comp_left (f : C ⟶ D) (g : Π i j, D.X i ⟶ E.X j) (i : ι) : d_next i (λ i j, f.f i ≫ g i j) = f.f i ≫ d_next i g := -begin - dsimp [d_next], - rcases c.next i with _|⟨i',w⟩, - { exact comp_zero.symm, }, - { dsimp [d_next], - simp, }, -end +(f.comm_assoc _ _ _).symm @[simp] lemma d_next_comp_right (f : Π i j, C.X i ⟶ D.X j) (g : D ⟶ E) (i : ι) : d_next i (λ i j, f i j ≫ g.f j) = d_next i f ≫ g.f i := -begin - dsimp [d_next], - rcases c.next i with _|⟨i',w⟩, - { exact zero_comp.symm, }, - { dsimp [d_next], - simp, }, -end +(category.assoc _ _ _).symm /-- The composition of `f j j' ≫ D.d j' j` if there is some `j'` coming before `j`, and `0` otherwise. -/ def prev_d (j : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X j ⟶ D.X j) := -add_monoid_hom.mk' (λ f, match c.prev j with -| none := 0 -| some ⟨j',w⟩ := f j j' ≫ D.d j' j -end) -begin - intros f g, - rcases c.prev j with _|⟨j',w⟩, - exact (zero_add _).symm, - exact preadditive.add_comp _ _ _ _ _ _, -end +add_monoid_hom.mk' (λ f, f j (c.prev j) ≫ D.d (c.prev j) j) $ +λ f g, preadditive.add_comp _ _ _ _ _ _ /-- `f j j'` if `j'` comes after `j`, and 0 if there's no such `j'`. Hopefully there won't be much need for this, except in `d_next_eq_d_from_from_next` to see that `d_next` factors through `C.d_from i`. -/ -def to_prev [has_zero_object V] (j : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X j ⟶ D.X_prev j) := -add_monoid_hom.mk' (λ f, match c.prev j with -| none := 0 -| some ⟨j',w⟩ := f j j' ≫ (D.X_prev_iso w).inv -end) -begin - intros f g, - rcases c.prev j with _|⟨j',w⟩, - exact (zero_add _).symm, - exact preadditive.add_comp _ _ _ _ _ _, -end +def to_prev (j : ι) : (Π i j, C.X i ⟶ D.X j) →+ (C.X j ⟶ D.X_prev j) := +add_monoid_hom.mk' (λ f, f j (c.prev j)) $ λ f g, rfl @[simp] -lemma prev_d_eq_to_prev_d_to [has_zero_object V] (f : Π i j, C.X i ⟶ D.X j) (j : ι) : - prev_d j f = to_prev j f ≫ D.d_to j := -begin - dsimp [prev_d, to_prev], - rcases c.prev j with ⟨⟩|⟨⟨j', w⟩⟩; - { dsimp [prev_d, to_prev], simp }, -end +lemma prev_d_eq_to_prev_d_to (f : Π i j, C.X i ⟶ D.X j) (j : ι) : + prev_d j f = to_prev j f ≫ D.d_to j := rfl lemma prev_d_eq (f : Π i j, C.X i ⟶ D.X j) {j j' : ι} (w : c.rel j' j) : prev_d j f = f j j' ≫ D.d j' j := -begin - dsimp [prev_d], - rw c.prev_eq_some w, - refl, -end +by { obtain rfl := c.prev_eq' w, refl } @[simp] lemma prev_d_comp_left (f : C ⟶ D) (g : Π i j, D.X i ⟶ E.X j) (j : ι) : prev_d j (λ i j, f.f i ≫ g i j) = f.f j ≫ prev_d j g := -begin - dsimp [prev_d], - rcases c.prev j with _|⟨j',w⟩, - { exact comp_zero.symm, }, - { dsimp [prev_d, hom.prev], - simp, }, -end +category.assoc _ _ _ @[simp] lemma prev_d_comp_right (f : Π i j, C.X i ⟶ D.X j) (g : D ⟶ E) (j : ι) : prev_d j (λ i j, f i j ≫ g.f j) = prev_d j f ≫ g.f j := -begin - dsimp [prev_d], - rcases c.prev j with _|⟨j',w⟩, - { exact zero_comp.symm, }, - { dsimp [prev_d], - simp, }, -end +by { dsimp [prev_d], simp only [category.assoc, g.comm] } lemma d_next_nat (C D : chain_complex V ℕ) (i : ℕ) (f : Π i j, C.X i ⟶ D.X j) : d_next i f = C.d i (i-1) ≫ f (i-1) i := begin + dsimp [d_next], cases i, - { dsimp [d_next], - rcases (complex_shape.down ℕ).next 0 with _|⟨j,hj⟩; - dsimp [d_next], - { rw [C.shape, zero_comp], dsimp, dec_trivial }, - { dsimp at hj, exact (nat.succ_ne_zero _ hj).elim } }, - rw d_next_eq, dsimp, refl + { simp only [shape, chain_complex.next_nat_zero, complex_shape.down_rel, + nat.one_ne_zero, not_false_iff, zero_comp], }, + { dsimp only [nat.succ_eq_add_one], + have : (complex_shape.down ℕ).next (i + 1) = i + 1 - 1, + { rw chain_complex.next_nat_succ, refl }, + congr' 2, } end lemma prev_d_nat (C D : cochain_complex V ℕ) (i : ℕ) (f : Π i j, C.X i ⟶ D.X j) : prev_d i f = f i (i-1) ≫ D.d (i-1) i := begin + dsimp [prev_d], cases i, - { dsimp [prev_d], - rcases (complex_shape.up ℕ).prev 0 with _|⟨j,hj⟩; - dsimp [prev_d], - { rw [D.shape, comp_zero], dsimp, dec_trivial }, - { dsimp at hj, exact (nat.succ_ne_zero _ hj).elim } }, - rw prev_d_eq, dsimp, refl + { simp only [shape, cochain_complex.prev_nat_zero, complex_shape.up_rel, + nat.one_ne_zero, not_false_iff, comp_zero]}, + { dsimp only [nat.succ_eq_add_one], + have : (complex_shape.up ℕ).prev (i + 1) = i + 1 - 1, + { rw cochain_complex.prev_nat_succ, refl }, + congr' 2, }, end /-- @@ -300,13 +226,9 @@ def null_homotopic_map (hom : Π i j, C.X i ⟶ D.X j) : C ⟶ D := comm' := λ i j hij, begin have eq1 : prev_d i hom ≫ D.d i j = 0, - { rcases h : c.prev i with _|⟨i',w⟩, - { dsimp [prev_d], rw h, erw zero_comp, }, - { rw [prev_d_eq hom w, category.assoc, D.d_comp_d' i' i j w hij, comp_zero], }, }, + { simp only [prev_d, add_monoid_hom.mk'_apply, category.assoc, d_comp_d, comp_zero], }, have eq2 : C.d i j ≫ d_next j hom = 0, - { rcases h : c.next j with _|⟨j',w⟩, - { dsimp [d_next], rw h, erw comp_zero, }, - { rw [d_next_eq hom w, ← category.assoc, C.d_comp_d' i j j' hij w, zero_comp], }, }, + { simp only [d_next, add_monoid_hom.mk'_apply, d_comp_d_assoc, zero_comp], }, rw [d_next_eq hom hij, prev_d_eq hom hij, preadditive.comp_add, preadditive.add_comp, eq1, eq2, add_zero, zero_add, category.assoc], end } @@ -322,8 +244,8 @@ lemma null_homotopic_map_comp (hom : Π i j, C.X i ⟶ D.X j) (g : D ⟶ E) : null_homotopic_map hom ≫ g = null_homotopic_map (λ i j, hom i j ≫ g.f j) := begin ext n, - dsimp [null_homotopic_map], - simp only [preadditive.add_comp, d_next_comp_right, prev_d_comp_right], + dsimp [null_homotopic_map, from_next, to_prev, add_monoid_hom.mk'_apply], + simp only [preadditive.add_comp, category.assoc, g.comm], end /-- Compatibility of `null_homotopic_map'` with the postcomposition by a morphism @@ -346,8 +268,8 @@ lemma comp_null_homotopic_map (f : C ⟶ D) (hom : Π i j, D.X i ⟶ E.X j) : f ≫ null_homotopic_map hom = null_homotopic_map (λ i j, f.f i ≫ hom i j) := begin ext n, - dsimp [null_homotopic_map], - simp only [preadditive.comp_add, d_next_comp_left, prev_d_comp_left], + dsimp [null_homotopic_map, from_next, to_prev, add_monoid_hom.mk'_apply], + simp only [preadditive.comp_add, category.assoc, f.comm_assoc], end /-- Compatibility of `null_homotopic_map'` with the precomposition by a morphism @@ -372,10 +294,7 @@ lemma map_null_homotopic_map {W : Type*} [category W] [preadditive W] begin ext i, dsimp [null_homotopic_map, d_next, prev_d], - rcases c.next i with _|⟨inext,wn⟩; - rcases c.prev i with _|⟨iprev,wp⟩; - dsimp [d_next, prev_d]; - simp only [G.map_comp, functor.map_zero, functor.map_add], + simp only [G.map_comp, functor.map_add], end /-- Compatibility of `null_homotopic_map'` with the application of additive functors -/ @@ -423,7 +342,7 @@ with `null_homotopic_map` or `null_homotopic_map'` -/ lemma null_homotopic_map_f {k₂ k₁ k₀ : ι} (r₂₁ : c.rel k₂ k₁) (r₁₀ : c.rel k₁ k₀) (hom : Π i j, C.X i ⟶ D.X j) : (null_homotopic_map hom).f k₁ = C.d k₁ k₀ ≫ hom k₀ k₁ + hom k₁ k₂ ≫ D.d k₂ k₁ := -by { dsimp [null_homotopic_map], rw [d_next_eq hom r₁₀, prev_d_eq hom r₂₁], } +by { dsimp only [null_homotopic_map], rw [d_next_eq hom r₁₀, prev_d_eq hom r₂₁], } @[simp] lemma null_homotopic_map'_f {k₂ k₁ k₀ : ι} (r₂₁ : c.rel k₂ k₁) (r₁₀ : c.rel k₁ k₀) @@ -443,10 +362,9 @@ lemma null_homotopic_map_f_of_not_rel_left {k₁ k₀ : ι} (r₁₀ : c.rel k (hom : Π i j, C.X i ⟶ D.X j) : (null_homotopic_map hom).f k₀ = hom k₀ k₁ ≫ D.d k₁ k₀ := begin - dsimp [null_homotopic_map], - rw prev_d_eq hom r₁₀, - rcases h : c.next k₀ with _|⟨l,w⟩, swap, exfalso, exact hk₀ l w, - dsimp [d_next], rw h, erw zero_add, + dsimp only [null_homotopic_map], + rw [prev_d_eq hom r₁₀, d_next, add_monoid_hom.mk'_apply, C.shape, zero_comp, zero_add], + exact hk₀ _ end @[simp] @@ -468,10 +386,9 @@ lemma null_homotopic_map_f_of_not_rel_right {k₁ k₀ : ι} (r₁₀ : c.rel k (hom : Π i j, C.X i ⟶ D.X j) : (null_homotopic_map hom).f k₁ = C.d k₁ k₀ ≫ hom k₀ k₁ := begin - dsimp [null_homotopic_map], - rw d_next_eq hom r₁₀, - rcases h : c.prev k₁ with _|⟨l,w⟩, swap, exfalso, exact hk₁ l w, - dsimp [prev_d], rw h, erw add_zero, + dsimp only [null_homotopic_map], + rw [d_next_eq hom r₁₀, prev_d, add_monoid_hom.mk'_apply, D.shape, comp_zero, add_zero], + exact hk₁ _, end @[simp] @@ -493,13 +410,8 @@ lemma null_homotopic_map_f_eq_zero {k₀ : ι} (hom : Π i j, C.X i ⟶ D.X j) : (null_homotopic_map hom).f k₀ = 0 := begin - dsimp [null_homotopic_map], - rcases h1 : c.next k₀ with _|⟨l,w⟩, swap, exfalso, exact hk₀ l w, - rcases h2 : c.prev k₀ with _|⟨l,w⟩, swap, exfalso, exact hk₀' l w, - dsimp [d_next, prev_d], - rw [h1, h2], - erw zero_add, - refl, + dsimp [null_homotopic_map, d_next, prev_d], + rw [C.shape, D.shape, zero_comp, comp_zero, add_zero]; apply_assumption, end @[simp] @@ -534,24 +446,24 @@ variables {P Q : chain_complex V ℕ} prev_d j f = f j (j+1) ≫ Q.d _ _ := begin dsimp [prev_d], - simp only [chain_complex.prev], - refl, + have : (complex_shape.down ℕ).prev j = j + 1 := chain_complex.prev ℕ j, + congr' 2, end @[simp] lemma d_next_succ_chain_complex (f : Π i j, P.X i ⟶ Q.X j) (i : ℕ) : d_next (i+1) f = P.d _ _ ≫ f i (i+1) := begin dsimp [d_next], - simp only [chain_complex.next_nat_succ], - refl, + have : (complex_shape.down ℕ).next (i + 1) = i := chain_complex.next_nat_succ _, + congr' 2, end @[simp] lemma d_next_zero_chain_complex (f : Π i j, P.X i ⟶ Q.X j) : d_next 0 f = 0 := begin dsimp [d_next], - simp only [chain_complex.next_nat_zero], - refl, + rw [P.shape, zero_comp], + rw chain_complex.next_nat_zero, dsimp, dec_trivial, end variables (e : P ⟶ Q) @@ -591,8 +503,6 @@ def mk_inductive_aux₁ : section -variable [has_zero_object V] - /-- An auxiliary construction for `mk_inductive`. -/ @@ -603,10 +513,10 @@ def mk_inductive_aux₂ : | (n+1) := let I := mk_inductive_aux₁ e zero comm_zero one comm_one succ n in ⟨(P.X_next_iso rfl).hom ≫ I.1, I.2.1 ≫ (Q.X_prev_iso rfl).inv, by simpa using I.2.2⟩ -lemma mk_inductive_aux₃ (i : ℕ) : - (mk_inductive_aux₂ e zero comm_zero one comm_one succ i).2.1 ≫ (Q.X_prev_iso rfl).hom - = (P.X_next_iso rfl).inv ≫ (mk_inductive_aux₂ e zero comm_zero one comm_one succ (i+1)).1 := -by rcases i with (_|_|i); { dsimp, simp, } +lemma mk_inductive_aux₃ (i j : ℕ) (h : i+1 = j) : + (mk_inductive_aux₂ e zero comm_zero one comm_one succ i).2.1 ≫ (Q.X_prev_iso h).hom + = (P.X_next_iso h).inv ≫ (mk_inductive_aux₂ e zero comm_zero one comm_one succ j).1 := +by subst j; rcases i with (_|_|i); { dsimp, simp, } /-- A constructor for a `homotopy e 0`, for `e` a chain map between `ℕ`-indexed chain complexes, @@ -627,20 +537,16 @@ def mk_inductive : homotopy e 0 := comm := λ i, begin dsimp, simp only [add_zero], convert (mk_inductive_aux₂ e zero comm_zero one comm_one succ i).2.2, - { rcases i with (_|_|_|i), - { dsimp, - simp only [d_next_zero_chain_complex, d_from_eq_zero, limits.comp_zero], }, - all_goals - { simp only [d_next_succ_chain_complex], - dsimp, - simp only [category.comp_id, category.assoc, iso.inv_hom_id, d_from_comp_X_next_iso_assoc, - dite_eq_ite, if_true, eq_self_iff_true]}, }, { cases i, - all_goals - { simp only [prev_d_chain_complex], - dsimp, - simp only [category.comp_id, category.assoc, iso.inv_hom_id, X_prev_iso_comp_d_to, - dite_eq_ite, if_true, eq_self_iff_true], }, }, + { dsimp [from_next], rw dif_neg, + simp only [chain_complex.next_nat_zero, nat.one_ne_zero, not_false_iff], }, + { dsimp [from_next], rw dif_pos, swap, { simp only [chain_complex.next_nat_succ] }, + have aux : (complex_shape.down ℕ).next i.succ = i := chain_complex.next_nat_succ i, + rw mk_inductive_aux₃ e zero comm_zero one comm_one succ + ((complex_shape.down ℕ).next i.succ) (i+1) (by rw aux), + dsimp [X_next_iso], erw category.id_comp, } }, + { dsimp [to_prev], rw dif_pos, swap, { simp only [chain_complex.prev] }, + dsimp [X_prev_iso], erw category.comp_id, }, end, } end @@ -660,24 +566,24 @@ variables {P Q : cochain_complex V ℕ} d_next j f = P.d _ _ ≫ f (j+1) j := begin dsimp [d_next], - simp only [cochain_complex.next], - refl, + have : (complex_shape.up ℕ).next j = j + 1 := cochain_complex.next ℕ j, + congr' 2, end @[simp] lemma prev_d_succ_cochain_complex (f : Π i j, P.X i ⟶ Q.X j) (i : ℕ) : prev_d (i+1) f = f (i+1) _ ≫ Q.d i (i+1) := begin dsimp [prev_d], - simp [cochain_complex.prev_nat_succ], - refl, + have : (complex_shape.up ℕ).prev (i+1) = i := cochain_complex.prev_nat_succ i, + congr' 2, end @[simp] lemma prev_d_zero_cochain_complex (f : Π i j, P.X i ⟶ Q.X j) : prev_d 0 f = 0 := begin dsimp [prev_d], - simp only [cochain_complex.prev_nat_zero], - refl, + rw [Q.shape, comp_zero], + rw [cochain_complex.prev_nat_zero], dsimp, dec_trivial, end variables (e : P ⟶ Q) @@ -717,8 +623,6 @@ def mk_coinductive_aux₁ : section -variable [has_zero_object V] - /-- An auxiliary construction for `mk_inductive`. -/ @@ -730,10 +634,10 @@ def mk_coinductive_aux₂ : | (n+1) := let I := mk_coinductive_aux₁ e zero comm_zero one comm_one succ n in ⟨I.1 ≫ (Q.X_prev_iso rfl).inv, (P.X_next_iso rfl).hom ≫ I.2.1, by simpa using I.2.2⟩ -lemma mk_coinductive_aux₃ (i : ℕ) : - (mk_coinductive_aux₂ e zero comm_zero one comm_one succ i).2.1 ≫ (Q.X_prev_iso rfl).inv - = (P.X_next_iso rfl).hom ≫ (mk_coinductive_aux₂ e zero comm_zero one comm_one succ (i+1)).1 := -by rcases i with (_|_|i); { dsimp, simp, } +lemma mk_coinductive_aux₃ (i j : ℕ) (h : i + 1 = j) : + (P.X_next_iso h).inv ≫ (mk_coinductive_aux₂ e zero comm_zero one comm_one succ i).2.1 + = (mk_coinductive_aux₂ e zero comm_zero one comm_one succ j).1 ≫ (Q.X_prev_iso h).hom := +by subst j; rcases i with (_|_|i); { dsimp, simp, } /-- A constructor for a `homotopy e 0`, for `e` a chain map between `ℕ`-indexed cochain complexes, @@ -755,26 +659,16 @@ def mk_coinductive : homotopy e 0 := dsimp, rw [add_zero, add_comm], convert (mk_coinductive_aux₂ e zero comm_zero one comm_one succ i).2.2 using 2, - { rcases i with (_|_|_|i), - { simp only [mk_coinductive_aux₂, prev_d_zero_cochain_complex, zero_comp] }, - all_goals - { simp only [prev_d_succ_cochain_complex], - dsimp, - simp only [eq_self_iff_true, iso.inv_hom_id_assoc, dite_eq_ite, - if_true, category.assoc, X_prev_iso_comp_d_to], }, }, { cases i, - { dsimp, - simp only [eq_self_iff_true, d_next_cochain_complex, dif_pos, - d_from_comp_X_next_iso_assoc, ←comm_zero], - rw mk_coinductive_aux₂, - dsimp, - convert comm_zero.symm, - simp only [iso.inv_hom_id_assoc], }, - { dsimp, - simp only [eq_self_iff_true, d_next_cochain_complex, dif_pos, d_from_comp_X_next_iso_assoc], - rw mk_coinductive_aux₂, - dsimp only, - simp only [iso.inv_hom_id_assoc], }, }, + { dsimp [to_prev], rw dif_neg, + simp only [cochain_complex.prev_nat_zero, nat.one_ne_zero, not_false_iff], }, + { dsimp [to_prev], rw dif_pos, swap, { simp only [cochain_complex.prev_nat_succ] }, + have aux : (complex_shape.up ℕ).prev i.succ = i := cochain_complex.prev_nat_succ i, + rw mk_coinductive_aux₃ e zero comm_zero one comm_one succ + ((complex_shape.up ℕ).prev i.succ) (i+1) (by rw aux), + dsimp [X_prev_iso], erw category.comp_id, } }, + { dsimp [from_next], rw dif_pos, swap, { simp only [cochain_complex.next] }, + dsimp [X_next_iso], erw category.id_comp, }, end } end @@ -829,7 +723,7 @@ instance : inhabited (homotopy_equiv C C) := ⟨refl C⟩ end homotopy_equiv -variables [has_equalizers V] [has_cokernels V] [has_images V] [has_image_maps V] [has_zero_object V] +variables [has_equalizers V] [has_cokernels V] [has_images V] [has_image_maps V] /-- Homotopic maps induce the same map on homology. @@ -880,13 +774,10 @@ def functor.map_homotopy (F : V ⥤ W) [F.additive] {f g : C ⟶ D} (h : homotop { hom := λ i j, F.map (h.hom i j), zero' := λ i j w, by { rw [h.zero i j w, F.map_zero], }, comm := λ i, begin - have := h.comm i, dsimp [d_next, prev_d] at *, - rcases c.next i with _|⟨inext,wn⟩; - rcases c.prev i with _|⟨iprev,wp⟩; - dsimp [d_next, prev_d] at *; - { intro h, - simp [h] }, + rw h.comm i, + simp only [F.map_add, ← F.map_comp], + refl end, } /-- An additive functor preserves homotopy equivalences. -/ diff --git a/src/algebra/homology/homotopy_category.lean b/src/algebra/homology/homotopy_category.lean index 373657a40e757..d9c581898935e 100644 --- a/src/algebra/homology/homotopy_category.lean +++ b/src/algebra/homology/homotopy_category.lean @@ -50,9 +50,9 @@ namespace homotopy_category def quotient : homological_complex V c ⥤ homotopy_category V c := category_theory.quotient.functor _ -local attribute [instance] has_zero_object.has_zero +open_locale zero_object --- TODO upgrade this is to `has_zero_object`, presumably for any `quotient`. +-- TODO upgrade this to `has_zero_object`, presumably for any `quotient`. instance [has_zero_object V] : inhabited (homotopy_category V c) := ⟨(quotient V c).obj 0⟩ variables {V c} @@ -113,7 +113,7 @@ def homotopy_equiv_of_iso homotopy_hom_inv_id := homotopy_of_eq _ _ (by { simp, refl, }), homotopy_inv_hom_id := homotopy_of_eq _ _ (by { simp, refl, }), } -variables (V c) [has_zero_object V] [has_equalizers V] [has_images V] [has_image_maps V] +variables (V c) [has_equalizers V] [has_images V] [has_image_maps V] [has_cokernels V] /-- The `i`-th homology, as a functor from the homotopy category. -/ diff --git a/src/category_theory/functor/left_derived.lean b/src/category_theory/functor/left_derived.lean index 6f54af43a5929..b8df704a33603 100644 --- a/src/category_theory/functor/left_derived.lean +++ b/src/category_theory/functor/left_derived.lean @@ -52,7 +52,7 @@ variables {C : Type u} [category.{v} C] {D : Type*} [category D] -- `[abelian C] [enough_projectives C] [abelian D]` suffices to acquire all the following: variables [preadditive C] [has_zero_object C] [has_equalizers C] [has_images C] [has_projective_resolutions C] -variables [preadditive D] [has_zero_object D] [has_equalizers D] [has_cokernels D] +variables [preadditive D] [has_equalizers D] [has_cokernels D] [has_images D] [has_image_maps D] /-- The left derived functors of an additive functor. -/ @@ -72,6 +72,9 @@ def functor.left_derived_obj_iso (F : C ⥤ D) [F.additive] (n : ℕ) (F.map_homotopy_equiv (ProjectiveResolution.homotopy_equiv _ P))) ≪≫ (homotopy_category.homology_factors D _ n).app _ +section +variables [has_zero_object D] + /-- The 0-th derived functor of `F` on a projective object `X` is just `F.obj X`. -/ @[simps] def functor.left_derived_obj_projective_zero (F : C ⥤ D) [F.additive] @@ -93,6 +96,8 @@ F.left_derived_obj_iso (n+1) (ProjectiveResolution.self X) ≪≫ (chain_complex.homology_functor_succ_single₀ D n).app (F.obj X) ≪≫ (functor.zero_obj _).iso_zero +end + /-- We can compute a left derived functor on a morphism using a lift of that morphism to a chain map between chosen projective resolutions. From d52945123a872d653a560a3c6cf0f07e0cd8e564 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sat, 30 Jul 2022 00:53:43 +0000 Subject: [PATCH 145/173] feat(group_theory/p_group): Groups of order p^2 are commutative (#8632) Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> --- src/group_theory/p_group.lean | 98 ++++++++++++++++++++++++++++------- 1 file changed, 79 insertions(+), 19 deletions(-) diff --git a/src/group_theory/p_group.lean b/src/group_theory/p_group.lean index 465062b3829f7..ffc7e3ae2b48e 100644 --- a/src/group_theory/p_group.lean +++ b/src/group_theory/p_group.lean @@ -9,6 +9,8 @@ import group_theory.index import group_theory.group_action.conj_act import group_theory.group_action.quotient import group_theory.perm.cycle.type +import group_theory.specific_groups.cyclic +import tactic.interval_cases /-! # p-groups @@ -93,15 +95,22 @@ variables [hp : fact p.prime] include hp -lemma index (H : subgroup G) [fintype (G ⧸ H)] : +lemma index (H : subgroup G) [finite (G ⧸ H)] : ∃ n : ℕ, H.index = p ^ n := begin + casesI nonempty_fintype (G ⧸ H), obtain ⟨n, hn⟩ := iff_card.mp (hG.to_quotient H.normal_core), obtain ⟨k, hk1, hk2⟩ := (nat.dvd_prime_pow hp.out).mp ((congr_arg _ (H.normal_core.index_eq_card.trans hn)).mp (subgroup.index_dvd_of_le H.normal_core_le)), exact ⟨k, hk2⟩, end +lemma nontrivial_iff_card [fintype G] : nontrivial G ↔ ∃ n > 0, card G = p ^ n := +⟨λ hGnt, let ⟨k, hk⟩ := iff_card.1 hG in ⟨k, nat.pos_of_ne_zero $ λ hk0, + by rw [hk0, pow_zero] at hk; exactI fintype.one_lt_card.ne' hk, hk⟩, +λ ⟨k, hk0, hk⟩, one_lt_card_iff_nontrivial.1 $ hk.symm ▸ + one_lt_pow (fact.out p.prime).one_lt (ne_of_gt hk0)⟩ + variables {α : Type*} [mul_action G α] lemma card_orbit (a : α) [fintype (orbit G a)] : @@ -113,11 +122,12 @@ begin exact hG.index (stabilizer G a), end -variables (α) [fintype α] [fintype (fixed_points G α)] +variables (α) [fintype α] /-- 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 `α` -/ -lemma card_modeq_card_fixed_points : card α ≡ card (fixed_points G α) [MOD p] := +lemma card_modeq_card_fixed_points [fintype (fixed_points G α)] : + card α ≡ card (fixed_points G α) [MOD p] := begin classical, calc card α = card (Σ y : quotient (orbit_rel G α), {x // quotient.mk' x = y}) : @@ -142,10 +152,12 @@ end /-- If a p-group acts on `α` and the cardinality of `α` is not a multiple of `p` then the action has a fixed point. -/ -lemma nonempty_fixed_point_of_prime_not_dvd_card (hpα : ¬ p ∣ card α) : +lemma nonempty_fixed_point_of_prime_not_dvd_card (hpα : ¬ p ∣ card α) + [finite (fixed_points G α)] : (fixed_points G α).nonempty := @set.nonempty_of_nonempty_subtype _ _ begin -rw [←card_pos_iff, pos_iff_ne_zero], + casesI nonempty_fintype (fixed_points G α), + rw [←card_pos_iff, pos_iff_ne_zero], contrapose! hpα, rw [←nat.modeq_zero_iff_dvd, ←hpα], exact hG.card_modeq_card_fixed_points α, @@ -156,30 +168,32 @@ end lemma exists_fixed_point_of_prime_dvd_card_of_fixed_point (hpα : p ∣ card α) {a : α} (ha : a ∈ fixed_points G α) : ∃ b, b ∈ fixed_points G α ∧ a ≠ b := -have hpf : p ∣ card (fixed_points G α) := - nat.modeq_zero_iff_dvd.mp ((hG.card_modeq_card_fixed_points α).symm.trans hpα.modeq_zero_nat), -have hα : 1 < card (fixed_points G α) := - (fact.out p.prime).one_lt.trans_le (nat.le_of_dvd (card_pos_iff.2 ⟨⟨a, ha⟩⟩) hpf), -let ⟨⟨b, hb⟩, hba⟩ := exists_ne_of_one_lt_card hα ⟨a, ha⟩ in -⟨b, hb, λ hab, hba (by simp_rw [hab])⟩ - -lemma center_nontrivial [nontrivial G] [fintype G] : nontrivial (subgroup.center G) := +begin + casesI nonempty_fintype (fixed_points G α), + have hpf : p ∣ card (fixed_points G α) := + nat.modeq_zero_iff_dvd.mp ((hG.card_modeq_card_fixed_points α).symm.trans hpα.modeq_zero_nat), + have hα : 1 < card (fixed_points G α) := + (fact.out p.prime).one_lt.trans_le (nat.le_of_dvd (card_pos_iff.2 ⟨⟨a, ha⟩⟩) hpf), + exact let ⟨⟨b, hb⟩, hba⟩ := exists_ne_of_one_lt_card hα ⟨a, ha⟩ in + ⟨b, hb, λ hab, hba (by simp_rw [hab])⟩ +end + +lemma center_nontrivial [nontrivial G] [finite G] : nontrivial (subgroup.center G) := begin classical, + casesI nonempty_fintype G, have := (hG.of_equiv conj_act.to_conj_act).exists_fixed_point_of_prime_dvd_card_of_fixed_point G, rw conj_act.fixed_points_eq_center at this, obtain ⟨g, hg⟩ := this _ (subgroup.center G).one_mem, { exact ⟨⟨1, ⟨g, hg.1⟩, mt subtype.ext_iff.mp hg.2⟩⟩ }, - { obtain ⟨n, hn⟩ := is_p_group.iff_card.mp hG, - rw hn, - apply dvd_pow_self, - rintro rfl, - exact (fintype.one_lt_card).ne' hn }, + { obtain ⟨n, hn0, hn⟩ := hG.nontrivial_iff_card.mp infer_instance, + exact hn.symm ▸ dvd_pow_self _ (ne_of_gt hn0) }, end -lemma bot_lt_center [nontrivial G] [fintype G] : ⊥ < subgroup.center G := +lemma bot_lt_center [nontrivial G] [finite G] : ⊥ < subgroup.center G := begin haveI := center_nontrivial hG, + casesI nonempty_fintype G, classical, exact bot_lt_iff_ne_bot.mpr ((subgroup.center G).one_lt_card_iff_ne_bot.mp fintype.one_lt_card), end @@ -280,4 +294,50 @@ begin simpa [this] using hn₁, end +section p2comm + +variables [fintype G] [fact p.prime] {n : ℕ} (hGpn : card G = p ^ n) +include hGpn +open subgroup + +/-- The cardinality of the `center` of a `p`-group is `p ^ k` where `k` is positive. -/ +lemma card_center_eq_prime_pow (hn : 0 < n) [fintype (center G)] : + ∃ k > 0, card (center G) = p ^ k := +begin + have hcG := to_subgroup (of_card hGpn) (center G), + rcases iff_card.1 hcG with ⟨k, hk⟩, + haveI : nontrivial G := (nontrivial_iff_card $ of_card hGpn).2 ⟨n, hn, hGpn⟩, + exact (nontrivial_iff_card hcG).mp (center_nontrivial (of_card hGpn)), +end + +omit hGpn + +/-- The quotient by the center of a group of cardinality `p ^ 2` is cyclic. -/ +lemma cyclic_center_quotient_of_card_eq_prime_sq (hG : card G = p ^ 2) : + is_cyclic (G ⧸ (center G)) := +begin + classical, + rcases card_center_eq_prime_pow hG zero_lt_two with ⟨k, hk0, hk⟩, + rw [card_eq_card_quotient_mul_card_subgroup (center G), mul_comm, hk] at hG, + have hk2 := (nat.pow_dvd_pow_iff_le_right (fact.out p.prime).one_lt).1 ⟨_, hG.symm⟩, + interval_cases k, + { rw [sq, pow_one, nat.mul_right_inj (fact.out p.prime).pos] at hG, + exact is_cyclic_of_prime_card hG }, + { exact @is_cyclic_of_subsingleton _ _ ⟨fintype.card_le_one_iff.1 ((nat.mul_right_inj + (pow_pos (fact.out p.prime).pos 2)).1 (hG.trans (mul_one (p ^ 2)).symm)).le⟩ }, +end + +/-- A group of order `p ^ 2` is commutative. See also `is_p_group.comm_group_of_card_eq_prime_sq` +for the `comm_group` instance. -/ +def comm_group_of_card_eq_prime_sq (hG : card G = p ^ 2) : comm_group G := +@comm_group_of_cycle_center_quotient _ _ _ _ (cyclic_center_quotient_of_card_eq_prime_sq hG) _ + (quotient_group.ker_mk (center G)).le + +/-- A group of order `p ^ 2` is commutative. See also `is_p_group.commutative_of_card_eq_prime_sq` +for just the proof that `∀ a b, a * b = b * a` -/ +lemma commutative_of_card_eq_prime_sq (hG : card G = p ^ 2) : ∀ a b : G, a * b = b * a := +(comm_group_of_card_eq_prime_sq hG).mul_comm + +end p2comm + end is_p_group From 8af57a7502f75688bf1701d03058a39f32d77219 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 30 Jul 2022 04:28:46 +0000 Subject: [PATCH 146/173] feat(algebra/big_operators): add lemmas about sum of disjoint multiset (#15556) Some easy lemmas about sum of disjoint ```multiset``` are proved. --- src/algebra/big_operators/basic.lean | 69 ++++++++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 9636c53e79a5f..eda04d9787bcb 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1631,8 +1631,77 @@ end list namespace multiset +lemma disjoint_list_sum_left {a : multiset α} {l : list (multiset α)} : + multiset.disjoint l.sum a ↔ ∀ b ∈ l, multiset.disjoint b a := +begin + induction l with b bs ih, + { simp only [zero_disjoint, list.not_mem_nil, is_empty.forall_iff, forall_const, list.sum_nil], }, + { simp_rw [list.sum_cons, disjoint_add_left, list.mem_cons_iff, forall_eq_or_imp], + simp [and.congr_left_iff, iff_self, ih], }, +end + +lemma disjoint_list_sum_right {a : multiset α} {l : list (multiset α)} : + multiset.disjoint a l.sum ↔ ∀ b ∈ l, multiset.disjoint a b := +by simpa only [disjoint_comm] using disjoint_list_sum_left + +lemma disjoint_sum_left {a : multiset α} {i : multiset (multiset α)} : + multiset.disjoint i.sum a ↔ ∀ b ∈ i, multiset.disjoint b a := +quotient.induction_on i $ λ l, begin + rw [quot_mk_to_coe, multiset.coe_sum], + exact disjoint_list_sum_left, +end + +lemma disjoint_sum_right {a : multiset α} {i : multiset (multiset α)} : + multiset.disjoint a i.sum ↔ ∀ b ∈ i, multiset.disjoint a b := +by simpa only [disjoint_comm] using disjoint_sum_left + +lemma disjoint_finset_sum_left {β : Type*} {i : finset β} {f : β → multiset α} {a : 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], +end + +lemma disjoint_finset_sum_right {β : Type*} {i : finset β} {f : β → multiset α} {a : multiset α} : + multiset.disjoint a (i.sum f) ↔ ∀ b ∈ i, multiset.disjoint a (f b) := +by simpa only [disjoint_comm] using disjoint_finset_sum_left + variables [decidable_eq α] +lemma add_eq_union_left_of_le {x y z : multiset α} (h : y ≤ x) : + z + x = z ∪ y ↔ z.disjoint x ∧ x = y := +begin + rw ←add_eq_union_iff_disjoint, + split, + { intro h0, + rw and_iff_right_of_imp, + { exact (le_of_add_le_add_left $ h0.trans_le $ union_le_add z y).antisymm h, }, + { rintro rfl, + exact h0, } }, + { rintro ⟨h0, rfl⟩, + exact h0, } +end + +lemma add_eq_union_right_of_le {x y z : multiset α} (h : z ≤ y) : + x + y = x ∪ z ↔ y = z ∧ x.disjoint y := +by simpa only [and_comm] using add_eq_union_left_of_le h + +lemma finset_sum_eq_sup_iff_disjoint {β : Type*} {i : finset β} {f : β → multiset α} : + i.sum f = i.sup f ↔ ∀ x y ∈ i, x ≠ y → multiset.disjoint (f x) (f y) := +begin + induction i using finset.cons_induction_on with z i hz hr, + { simp only [finset.not_mem_empty, is_empty.forall_iff, implies_true_iff, + finset.sum_empty, finset.sup_empty, bot_eq_zero, eq_self_iff_true], }, + { simp_rw [finset.sum_cons hz, finset.sup_cons, finset.mem_cons, multiset.sup_eq_union, + forall_eq_or_imp, ne.def, eq_self_iff_true, not_true, is_empty.forall_iff, true_and, + imp_and_distrib, forall_and_distrib, ←hr, @eq_comm _ z], + have := λ x ∈ i, ne_of_mem_of_not_mem H hz, + simp only [this, not_false_iff, true_implies_iff] {contextual := tt}, + simp_rw [←disjoint_finset_sum_left, ←disjoint_finset_sum_right, disjoint_comm, ←and_assoc, + and_self], + exact add_eq_union_left_of_le (finset.sup_le (λ x hx, le_sum_of_mem (mem_map_of_mem f hx))), }, +end + @[simp] lemma to_finset_sum_count_eq (s : multiset α) : (∑ a in s.to_finset, s.count a) = s.card := multiset.induction_on s rfl From dd3fac27b7ec7733daadb7dbf8068d999a22c310 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 30 Jul 2022 06:36:36 +0000 Subject: [PATCH 147/173] refactor(tactic/lint/type_classes): change inhabited linter to nonempty linter (#15542) The `inhabited` typeclass is intended to be a default value (and computable, if one is writing programs) but many times it is filled in with a junk (and perhaps noncomputable) value to satisfy the `has_inhabited_instance` linter. This commit switches to a `has_nonempty_instance` linter to push contributors toward supplying a `nonempty` instance instead. The linter still accepts `inhabited` and `unique` instances, which should be preferred over `nonempty` when appropriate. --- scripts/nolints.txt | 22 ++++++------- src/algebra/algebra/basic.lean | 4 +-- src/algebra/category/Group/epi_mono.lean | 2 +- src/algebra/category/Module/adjunctions.lean | 2 +- src/algebra/hom/group_action.lean | 4 +-- src/algebra/homology/complex_shape.lean | 2 +- src/algebra/homology/homological_complex.lean | 4 +-- src/algebra/homology/homotopy.lean | 2 +- .../homology/short_exact/preadditive.lean | 2 +- src/algebra/module/basic.lean | 2 +- src/algebra/module/equiv.lean | 2 +- src/algebra/module/localized_module.lean | 2 +- src/algebra/order/group.lean | 4 +-- src/algebra/order/ring.lean | 4 +-- src/algebra/quandle.lean | 4 +-- src/algebra/quotient.lean | 4 +-- src/algebra/star/chsh.lean | 2 +- src/algebraic_geometry/gluing.lean | 2 +- .../locally_ringed_space.lean | 2 +- .../presheafed_space/gluing.lean | 6 ++-- .../prime_spectrum/basic.lean | 2 +- .../projective_spectrum/topology.lean | 2 +- src/algebraic_topology/simplex_category.lean | 2 +- src/algebraic_topology/simplicial_object.lean | 12 +++---- src/analysis/calculus/implicit.lean | 2 +- src/analysis/inner_product_space/basic.lean | 2 +- src/analysis/von_neumann_algebra/basic.lean | 2 +- src/category_theory/adjunction/basic.lean | 4 +-- src/category_theory/bicategory/free.lean | 2 +- src/category_theory/bicategory/functor.lean | 2 +- src/category_theory/category/PartialFun.lean | 2 +- .../concrete_category/bundled.lean | 2 +- src/category_theory/core.lean | 2 +- src/category_theory/differential_object.lean | 4 +-- src/category_theory/elements.lean | 2 +- src/category_theory/enriched/basic.lean | 6 ++-- src/category_theory/essentially_small.lean | 4 +-- src/category_theory/full_subcategory.lean | 2 +- src/category_theory/glue_data.lean | 2 +- src/category_theory/grothendieck.lean | 2 +- src/category_theory/idempotents/karoubi.lean | 2 +- src/category_theory/limits/has_limits.lean | 4 +-- src/category_theory/limits/is_limit.lean | 4 +-- .../limits/shapes/biproducts.lean | 12 +++---- .../limits/shapes/multiequalizer.lean | 8 ++--- src/category_theory/limits/shapes/types.lean | 2 +- src/category_theory/limits/types.lean | 2 +- .../localization/construction.lean | 4 +-- src/category_theory/monad/algebra.lean | 4 +-- src/category_theory/monoidal/center.lean | 6 ++-- src/category_theory/monoidal/free/basic.lean | 2 +- .../monoidal/free/coherence.lean | 2 +- .../monoidal/of_chosen_finite_products.lean | 2 +- src/category_theory/monoidal/opposite.lean | 2 +- src/category_theory/preadditive/Mat.lean | 2 +- .../preadditive/additive_functor.lean | 2 +- .../preadditive/injective.lean | 2 +- .../preadditive/injective_resolution.lean | 2 +- .../preadditive/projective.lean | 2 +- .../preadditive/projective_resolution.lean | 2 +- src/category_theory/shift.lean | 2 +- src/category_theory/single_obj.lean | 2 +- src/category_theory/sites/cover_lifting.lean | 2 +- .../sites/cover_preserving.lean | 4 +-- src/category_theory/sites/dense_subsite.lean | 2 +- src/category_theory/sites/grothendieck.lean | 4 +-- src/category_theory/sites/sheafification.lean | 2 +- src/category_theory/sites/sieves.lean | 2 +- src/category_theory/structured_arrow.lean | 4 +-- src/category_theory/with_terminal.lean | 4 +-- src/combinatorics/quiver/basic.lean | 2 +- .../quiver/connected_component.lean | 2 +- src/combinatorics/quiver/subquiver.lean | 4 +-- src/data/bundle.lean | 2 +- src/data/multiset/fintype.lean | 2 +- src/data/pfunctor/multivariate/W.lean | 2 +- src/data/pfunctor/univariate/basic.lean | 2 +- .../qpf/multivariate/constructions/fix.lean | 2 +- src/data/qpf/univariate/basic.lean | 2 +- src/geometry/manifold/charted_space.lean | 4 +-- src/geometry/manifold/diffeomorph.lean | 2 +- src/geometry/manifold/partition_of_unity.lean | 2 +- .../smooth_manifold_with_corners.lean | 2 +- src/geometry/manifold/tangent_bundle.lean | 2 +- src/group_theory/free_product.lean | 2 +- src/group_theory/monoid_localization.lean | 6 ++-- src/group_theory/nielsen_schreier.lean | 2 +- .../affine_space/affine_equiv.lean | 2 +- src/linear_algebra/dual.lean | 2 +- src/linear_algebra/free_module/basic.lean | 2 +- src/linear_algebra/free_module/pid.lean | 2 +- src/linear_algebra/matrix/transvection.lean | 2 +- .../projective_space/basic.lean | 2 +- .../quadratic_form/isometry.lean | 2 +- src/linear_algebra/ray.lean | 4 +-- src/logic/embedding.lean | 2 +- src/logic/small.lean | 2 +- .../covering/vitali_family.lean | 2 +- src/measure_theory/measure/measure_space.lean | 2 +- src/model_theory/basic.lean | 4 +-- .../legendre_symbol/add_character.lean | 2 +- src/order/galois_connection.lean | 4 +-- src/order/initial_seg.lean | 2 +- src/order/prime_ideal.lean | 2 +- src/order/rel_iso.lean | 2 +- src/ring_theory/dedekind_domain/ideal.lean | 2 +- .../homogeneous_localization.lean | 4 +-- src/ring_theory/perfection.lean | 8 ++--- src/ring_theory/power_basis.lean | 2 +- src/ring_theory/valuation/basic.lean | 4 +-- src/ring_theory/witt_vector/isocrystal.lean | 6 ++-- src/tactic/derive_fintype.lean | 2 +- src/tactic/lint/default.lean | 3 +- src/tactic/lint/type_classes.lean | 33 +++++++++++-------- src/topology/algebra/module/basic.lean | 2 +- src/topology/algebra/module/weak_dual.lean | 4 +-- src/topology/fiber_bundle.lean | 12 +++---- src/topology/gluing.lean | 4 +-- src/topology/homeomorph.lean | 2 +- src/topology/local_homeomorph.lean | 2 +- .../metric_space/gromov_hausdorff.lean | 2 +- .../gromov_hausdorff_realized.lean | 2 +- src/topology/metric_space/isometry.lean | 2 +- src/topology/metric_space/polish.lean | 4 +-- src/topology/path_connected.lean | 2 +- src/topology/sheaves/presheaf.lean | 2 +- src/topology/shrinking_lemma.lean | 2 +- src/topology/uniform_space/equiv.lean | 2 +- src/topology/vector_bundle/basic.lean | 10 +++--- test/has_scalar_comp_loop.lean | 2 +- 130 files changed, 222 insertions(+), 216 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 2331c52d28ab1..3c203a8305e45 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -206,7 +206,7 @@ apply_nolint traversable.mfoldr doc_blame -- control/monad/cont.lean apply_nolint cont doc_blame -apply_nolint cont_t has_inhabited_instance doc_blame +apply_nolint cont_t has_nonempty_instance doc_blame apply_nolint cont_t.map doc_blame apply_nolint cont_t.monad_lift doc_blame apply_nolint cont_t.run doc_blame @@ -216,7 +216,7 @@ apply_nolint except_t.mk_label doc_blame apply_nolint is_lawful_monad_cont doc_blame apply_nolint monad_cont doc_blame apply_nolint monad_cont.goto doc_blame -apply_nolint monad_cont.label has_inhabited_instance doc_blame +apply_nolint monad_cont.label has_nonempty_instance doc_blame apply_nolint option_t.call_cc doc_blame apply_nolint option_t.mk_label doc_blame apply_nolint reader_t.call_cc doc_blame @@ -232,7 +232,7 @@ apply_nolint except_t.pass_aux doc_blame apply_nolint option_t.pass_aux doc_blame apply_nolint swap_right doc_blame apply_nolint writer doc_blame -apply_nolint writer_t has_inhabited_instance doc_blame +apply_nolint writer_t has_nonempty_instance doc_blame apply_nolint writer_t.adapt doc_blame apply_nolint writer_t.bind doc_blame apply_nolint writer_t.ext unused_arguments @@ -270,22 +270,22 @@ apply_nolint tactic.interactive.traverse_field unused_arguments apply_nolint tactic.interactive.with_prefix doc_blame -- data/analysis/filter.lean -apply_nolint cfilter has_inhabited_instance +apply_nolint cfilter has_nonempty_instance apply_nolint cfilter.to_realizer doc_blame -apply_nolint filter.realizer has_inhabited_instance +apply_nolint filter.realizer has_nonempty_instance apply_nolint filter.realizer.of_eq doc_blame -- data/analysis/topology.lean -apply_nolint compact.realizer has_inhabited_instance doc_blame unused_arguments -apply_nolint ctop has_inhabited_instance -apply_nolint ctop.realizer has_inhabited_instance +apply_nolint compact.realizer has_nonempty_instance doc_blame unused_arguments +apply_nolint ctop has_nonempty_instance +apply_nolint ctop.realizer has_nonempty_instance apply_nolint ctop.realizer.id doc_blame apply_nolint ctop.realizer.nhds doc_blame apply_nolint ctop.realizer.nhds_F unused_arguments apply_nolint ctop.realizer.nhds_σ unused_arguments apply_nolint ctop.realizer.of_equiv doc_blame apply_nolint ctop.to_realizer doc_blame -apply_nolint locally_finite.realizer has_inhabited_instance doc_blame +apply_nolint locally_finite.realizer has_nonempty_instance doc_blame -- data/finset/noncomm_prod.lean apply_nolint add_monoid_hom.pi_ext fintype_finite @@ -374,7 +374,7 @@ apply_nolint int.shift2 doc_blame -- data/holor.lean apply_nolint holor.assoc_left doc_blame apply_nolint holor.assoc_right doc_blame -apply_nolint holor_index has_inhabited_instance +apply_nolint holor_index has_nonempty_instance apply_nolint holor_index.assoc_left doc_blame apply_nolint holor_index.assoc_right doc_blame apply_nolint holor_index.drop doc_blame @@ -1629,4 +1629,4 @@ apply_nolint Cauchy.gen doc_blame apply_nolint uniform_space.completion.completion_separation_quotient_equiv doc_blame apply_nolint uniform_space.completion.cpkg doc_blame apply_nolint uniform_space.completion.extension₂ doc_blame -apply_nolint uniform_space.completion.map₂ doc_blame \ No newline at end of file +apply_nolint uniform_space.completion.map₂ doc_blame diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 510a7b85e31b3..78f89406d1bb4 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -112,7 +112,7 @@ An associative unital `R`-algebra is a semiring `A` equipped with a map into its See the implementation notes in this file for discussion of the details of this definition. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] class algebra (R : Type u) (A : Type v) [comm_semiring R] [semiring A] extends has_smul R A, R →+* A := (commutes' : ∀ r x, to_fun r * x = x * to_fun r) @@ -474,7 +474,7 @@ end module set_option old_structure_cmd true /-- Defining the homomorphism in the category R-Alg. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure alg_hom (R : Type u) (A : Type v) (B : Type w) [comm_semiring R] [semiring A] [semiring B] [algebra R A] [algebra R B] extends ring_hom A B := (commutes' : ∀ r : R, to_fun (algebra_map R A r) = algebra_map R B r) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index acc4cac5ed5dd..63b00f01b4fa9 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -61,7 +61,7 @@ local notation `X` := set.range (function.swap left_coset f.range.carrier) /-- Define `X'` to be the set of all left cosets with an extra point at "infinity". -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] inductive X_with_infinity | from_coset : set.range (function.swap left_coset f.range.carrier) → X_with_infinity | infinity : X_with_infinity diff --git a/src/algebra/category/Module/adjunctions.lean b/src/algebra/category/Module/adjunctions.lean index d810bdd2231af..e3d6d7980eb0e 100644 --- a/src/algebra/category/Module/adjunctions.lean +++ b/src/algebra/category/Module/adjunctions.lean @@ -156,7 +156,7 @@ universes v u we will equip with a category structure where the morphisms are formal `R`-linear combinations of the morphisms in `C`. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def Free (R : Type*) (C : Type u) := C /-- diff --git a/src/algebra/hom/group_action.lean b/src/algebra/hom/group_action.lean index 787da40609da9..cd9109774eeb2 100644 --- a/src/algebra/hom/group_action.lean +++ b/src/algebra/hom/group_action.lean @@ -54,7 +54,7 @@ variables (G : Type*) [group G] (H : subgroup G) set_option old_structure_cmd true /-- Equivariant functions. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure mul_action_hom := (to_fun : X → Y) (map_smul' : ∀ (m : M') (x : X), to_fun (m • x) = m • to_fun x) @@ -258,7 +258,7 @@ end semiring end distrib_mul_action_hom /-- Equivariant ring homomorphisms. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure mul_semiring_action_hom extends R →+[M] S, R →+* S. /-- Reinterpret an equivariant ring homomorphism as a ring homomorphism. -/ diff --git a/src/algebra/homology/complex_shape.lean b/src/algebra/homology/complex_shape.lean index a2287f28e970d..640e7ce12ade7 100644 --- a/src/algebra/homology/complex_shape.lean +++ b/src/algebra/homology/complex_shape.lean @@ -55,7 +55,7 @@ This means that the shape consists of some union of lines, rays, intervals, and Below we define `c.next` and `c.prev` which provide these related elements. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure complex_shape (ι : Type*) := (rel : ι → ι → Prop) (next_eq : ∀ {i j j'}, rel i j → rel i j' → j = j') diff --git a/src/algebra/homology/homological_complex.lean b/src/algebra/homology/homological_complex.lean index 8e81c26ccc202..7ae35d627c867 100644 --- a/src/algebra/homology/homological_complex.lean +++ b/src/algebra/homology/homological_complex.lean @@ -552,7 +552,7 @@ Auxiliary structure for setting up the recursion in `mk`. This is purely an implementation detail: for some reason just using the dependent 6-tuple directly results in `mk_aux` taking much longer (well over the `-T100000` limit) to elaborate. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure mk_struct := (X₀ X₁ X₂ : V) (d₀ : X₁ ⟶ X₀) @@ -745,7 +745,7 @@ Auxiliary structure for setting up the recursion in `mk`. This is purely an implementation detail: for some reason just using the dependent 6-tuple directly results in `mk_aux` taking much longer (well over the `-T100000` limit) to elaborate. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure mk_struct := (X₀ X₁ X₂ : V) (d₀ : X₀ ⟶ X₁) diff --git a/src/algebra/homology/homotopy.lean b/src/algebra/homology/homotopy.lean index 3ad0100890e27..b733ad12c7c10 100644 --- a/src/algebra/homology/homotopy.lean +++ b/src/algebra/homology/homotopy.lean @@ -113,7 +113,7 @@ end A homotopy `h` between chain maps `f` and `g` consists of components `h i j : C.X i ⟶ D.X j` which are zero unless `c.rel j i`, satisfying the homotopy condition. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure homotopy (f g : C ⟶ D) := (hom : Π i j, C.X i ⟶ D.X j) (zero' : ∀ i j, ¬ c.rel j i → hom i j = 0 . obviously) diff --git a/src/algebra/homology/short_exact/preadditive.lean b/src/algebra/homology/short_exact/preadditive.lean index 5fc1314c04831..aea22167cbedc 100644 --- a/src/algebra/homology/short_exact/preadditive.lean +++ b/src/algebra/homology/short_exact/preadditive.lean @@ -167,7 +167,7 @@ end preadditive /-- A *splitting* of a sequence `A -f⟶ B -g⟶ C` is an isomorphism to the short exact sequence `0 ⟶ A ⟶ A ⊞ C ⟶ C ⟶ 0` such that the vertical maps on the left and the right are the identity. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure splitting [has_zero_morphisms 𝒜] [has_binary_biproducts 𝒜] := (iso : B ≅ A ⊞ C) (comp_iso_eq_inl : f ≫ iso.hom = biprod.inl) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 73c64a199c293..8e8ca9dc67b6b 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -201,7 +201,7 @@ instance add_comm_group.int_module : module ℤ M := and `smul_zero`. As these fields can be deduced from the other ones when `M` is an `add_comm_group`, this provides a way to construct a module structure by checking less properties, in `module.of_core`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure module.core extends has_smul R M := (smul_add : ∀(r : R) (x y : M), r • (x + y) = r • x + r • y) (add_smul : ∀(r s : R) (x : M), (r + s) • x = r • x + s • x) diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index 70322a8407a86..9c33929572c84 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -45,7 +45,7 @@ section set_option old_structure_cmd true /-- A linear equivalence is an invertible linear map. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure linear_equiv {R : Type*} {S : Type*} [semiring R] [semiring S] (σ : R →+* S) {σ' : S →+* R} [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ] (M : Type*) (M₂ : Type*) diff --git a/src/algebra/module/localized_module.lean b/src/algebra/module/localized_module.lean index 7cfcf5cd01fe6..870e117a2ef2a 100644 --- a/src/algebra/module/localized_module.lean +++ b/src/algebra/module/localized_module.lean @@ -70,7 +70,7 @@ instance r.setoid : setoid (M × S) := If `S` is a multiplicative subset of a ring `R` and `M` an `R`-module, then we can localize `M` by `S`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def _root_.localized_module : Type (max u v) := quotient (r.setoid S M) section diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index 8106e0424248c..55a25bb9208bf 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -1258,7 +1258,7 @@ namespace add_comm_group /-- A collection of elements in an `add_comm_group` designated as "non-negative". This is useful for constructing an `ordered_add_commm_group` by choosing a positive cone in an exisiting `add_comm_group`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure positive_cone (α : Type*) [add_comm_group α] := (nonneg : α → Prop) (pos : α → Prop := λ a, nonneg a ∧ ¬ nonneg (-a)) @@ -1269,7 +1269,7 @@ structure positive_cone (α : Type*) [add_comm_group α] := /-- A positive cone in an `add_comm_group` induces a linear order if for every `a`, either `a` or `-a` is non-negative. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure total_positive_cone (α : Type*) [add_comm_group α] extends positive_cone α := (nonneg_decidable : decidable_pred nonneg) (nonneg_total : ∀ a : α, nonneg a ∨ nonneg (-a)) diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index cc7b63bda7912..625adcf7dcb99 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -1435,7 +1435,7 @@ namespace ring /-- A positive cone in a ring consists of a positive cone in underlying `add_comm_group`, which contains `1` and such that the positive elements are closed under multiplication. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure positive_cone (α : Type*) [ring α] extends add_comm_group.positive_cone α := (one_nonneg : nonneg 1) (mul_pos : ∀ (a b), pos a → pos b → pos (a * b)) @@ -1444,7 +1444,7 @@ structure positive_cone (α : Type*) [ring α] extends add_comm_group.positive_c add_decl_doc positive_cone.to_positive_cone /-- A positive cone in a ring induces a linear order if `1` is a positive element. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure total_positive_cone (α : Type*) [ring α] extends positive_cone α, add_comm_group.total_positive_cone α := (one_pos : pos 1) diff --git a/src/algebra/quandle.lean b/src/algebra/quandle.lean index 866d9c0752fe8..75534effef233 100644 --- a/src/algebra/quandle.lean +++ b/src/algebra/quandle.lean @@ -298,7 +298,7 @@ instance opposite_quandle : quandle Qᵐᵒᵖ := The conjugation quandle of a group. Each element of the group acts by the corresponding inner automorphism. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def conj (G : Type*) := G instance conj.quandle (G : Type*) [group G] : quandle (conj G) := @@ -338,7 +338,7 @@ The dihedral quandle. This is the conjugation quandle of the dihedral group rest Used for Fox n-colorings of knots. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def dihedral (n : ℕ) := zmod n /-- diff --git a/src/algebra/quotient.lean b/src/algebra/quotient.lean index 82c1ccb8eae7d..0f31cb354613c 100644 --- a/src/algebra/quotient.lean +++ b/src/algebra/quotient.lean @@ -21,7 +21,7 @@ The actual quotient structures are defined in the following files: The following notation is introduced: -* `G ⧸ H` stands for the quotient of the type `G` by some term `H` +* `G ⧸ H` stands for the quotient of the type `G` by some term `H` (for example, `H` can be a normal subgroup of `G`). To implement this notation for other quotients, you should provide a `has_quotient` instance. Note that since `G` can usually be inferred from `H`, `_ ⧸ H` can also be used, @@ -49,7 +49,7 @@ class has_quotient (A : out_param $ Type u) (B : Type v) := This differs from `has_quotient.quotient'` in that the `A` argument is explicit, which is necessary to make Lean show the notation in the goal state. -/ -@[reducible, nolint has_inhabited_instance] -- Will be provided by e.g. `ideal.quotient.inhabited` +@[reducible, nolint has_nonempty_instance] -- Will be provided by e.g. `ideal.quotient.inhabited` def has_quotient.quotient (A : out_param $ Type u) {B : Type v} [has_quotient A B] (b : B) : Type (max u v) := has_quotient.quotient' b diff --git a/src/algebra/star/chsh.lean b/src/algebra/star/chsh.lean index dc060bed34440..e6a99191a1fd3 100644 --- a/src/algebra/star/chsh.lean +++ b/src/algebra/star/chsh.lean @@ -79,7 +79,7 @@ the `Aᵢ` commute with the `Bⱼ`. The physical interpretation is that `A₀` and `A₁` are a pair of boolean observables which are spacelike separated from another pair `B₀` and `B₁` of boolean observables. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure is_CHSH_tuple {R} [monoid R] [star_semigroup R] (A₀ A₁ B₀ B₁ : R) := (A₀_inv : A₀^2 = 1) (A₁_inv : A₁^2 = 1) (B₀_inv : B₀^2 = 1) (B₁_inv : B₁^2 = 1) (A₀_sa : star A₀ = A₀) (A₁_sa : star A₁ = A₁) (B₀_sa : star B₀ = B₀) (B₁_sa : star B₁ = B₁) diff --git a/src/algebraic_geometry/gluing.lean b/src/algebraic_geometry/gluing.lean index 8e7a89c701c99..9fc37f88a00ca 100644 --- a/src/algebraic_geometry/gluing.lean +++ b/src/algebraic_geometry/gluing.lean @@ -76,7 +76,7 @@ such that We can then glue the schemes `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subschemes of the glued space. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data extends category_theory.glue_data Scheme := (f_open : ∀ i j, is_open_immersion (f i j)) diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index 10096613620f6..ab03f1d82675c 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -31,7 +31,7 @@ such that all the stalks are local rings. A morphism of locally ringed spaces is a morphism of ringed spaces such that the morphisms induced on stalks are local ring homomorphisms. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure LocallyRingedSpace extends SheafedSpace CommRing := (local_ring : ∀ x, local_ring (presheaf.stalk x)) diff --git a/src/algebraic_geometry/presheafed_space/gluing.lean b/src/algebraic_geometry/presheafed_space/gluing.lean index 762d2b64a9440..94ef63d65c728 100644 --- a/src/algebraic_geometry/presheafed_space/gluing.lean +++ b/src/algebraic_geometry/presheafed_space/gluing.lean @@ -87,7 +87,7 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data extends glue_data (PresheafedSpace.{v} C) := (f_open : ∀ i j, is_open_immersion (f i j)) @@ -488,7 +488,7 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data extends glue_data (SheafedSpace.{v} C) := (f_open : ∀ i j, SheafedSpace.is_open_immersion (f i j)) @@ -560,7 +560,7 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data extends glue_data LocallyRingedSpace := (f_open : ∀ i j, LocallyRingedSpace.is_open_immersion (f i j)) diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index f93eb904aa8c6..f3ec6ef8b5dce 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -57,7 +57,7 @@ is the type of all prime ideals of `R`. It is naturally endowed with a topology (the Zariski topology), and a sheaf of commutative rings (see `algebraic_geometry.structure_sheaf`). It is a fundamental building block in algebraic geometry. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def prime_spectrum := {I : ideal R // I.is_prime} variable {R} diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index 4425efc12d9de..a0e04d1a10277 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -47,7 +47,7 @@ variables (𝒜 : ℕ → submodule R A) [graded_algebra 𝒜] 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. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def projective_spectrum := {I : homogeneous_ideal 𝒜 // I.to_ideal.is_prime ∧ ¬(homogeneous_ideal.irrelevant 𝒜 ≤ I)} diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index 36815db57f061..0dc093321ccff 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -62,7 +62,7 @@ def len (n : simplex_category) : ℕ := n @[simp] lemma mk_len (n : simplex_category) : [n.len] = n := rfl /-- Morphisms in the simplex_category. -/ -@[irreducible, nolint has_inhabited_instance] +@[irreducible, nolint has_nonempty_instance] protected def hom (a b : simplex_category) := fin (a.len + 1) →o fin (b.len + 1) namespace hom diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index 2c68362906215..be384bc4a9b86 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -31,7 +31,7 @@ variables (C : Type u) [category.{v} C] /-- The category of simplicial objects valued in a category `C`. This is the category of contravariant functors from `simplex_category` to `C`. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def simplicial_object := simplex_categoryᵒᵖ ⥤ C namespace simplicial_object @@ -119,7 +119,7 @@ def whiskering (D : Type*) [category D] : whiskering_right _ _ _ /-- Truncated simplicial objects. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def truncated (n : ℕ) := (simplex_category.truncated n)ᵒᵖ ⥤ C variable {C} @@ -163,7 +163,7 @@ variable (C) abbreviation const : C ⥤ simplicial_object C := category_theory.functor.const _ /-- The category of augmented simplicial objects, defined as a comma category. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def augmented := comma (𝟭 (simplicial_object C)) (const C) variable {C} @@ -261,7 +261,7 @@ by { dsimp, rw [simplex_category.hom_zero_zero ([0].const 0), op_id, X.map_id, c end simplicial_object /-- Cosimplicial objects. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def cosimplicial_object := simplex_category ⥤ C namespace cosimplicial_object @@ -349,7 +349,7 @@ def whiskering (D : Type*) [category D] : whiskering_right _ _ _ /-- Truncated cosimplicial objects. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def truncated (n : ℕ) := simplex_category.truncated n ⥤ C variable {C} @@ -394,7 +394,7 @@ variable (C) abbreviation const : C ⥤ cosimplicial_object C := category_theory.functor.const _ /-- Augmented cosimplicial objects. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def augmented := comma (const C) (𝟭 (cosimplicial_object C)) variable {C} diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 47aad7899c0a5..622726d32387c 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -87,7 +87,7 @@ such that * both functions are strictly differentiable at `a`; * the derivatives are surjective; * the kernels of the derivatives are complementary subspaces of `E`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure implicit_function_data (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] [complete_space E] (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] [complete_space F] diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 66177b2cc0ef1..b905e5c6b55cf 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -128,7 +128,7 @@ 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_inhabited_instance] +@[nolint has_nonempty_instance] structure inner_product_space.core (𝕜 : Type*) (F : Type*) [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] := diff --git a/src/analysis/von_neumann_algebra/basic.lean b/src/analysis/von_neumann_algebra/basic.lean index 720b1dd2b0401..e6c427a9f1301 100644 --- a/src/analysis/von_neumann_algebra/basic.lean +++ b/src/analysis/von_neumann_algebra/basic.lean @@ -64,7 +64,7 @@ Thus we can't say that the bounded operators `H →L[ℂ] H` form a `von_neumann (although we will later construct the instance `wstar_algebra (H →L[ℂ] H)`), and instead will use `⊤ : von_neumann_algebra H`. -/ -@[nolint has_inhabited_instance] +@[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) diff --git a/src/category_theory/adjunction/basic.lean b/src/category_theory/adjunction/basic.lean index 5bdfc3a91b088..0df5be9436e53 100644 --- a/src/category_theory/adjunction/basic.lean +++ b/src/category_theory/adjunction/basic.lean @@ -170,7 +170,7 @@ This is an auxiliary data structure useful for constructing adjunctions. See `adjunction.mk_of_hom_equiv`. This structure won't typically be used anywhere else. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure core_hom_equiv (F : C ⥤ D) (G : D ⥤ C) := (hom_equiv : Π (X Y), (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y)) (hom_equiv_naturality_left_symm' : Π {X' X Y} (f : X' ⟶ X) (g : X ⟶ G.obj Y), @@ -201,7 +201,7 @@ This is an auxiliary data structure useful for constructing adjunctions. See `adjunction.mk_of_unit_counit`. This structure won't typically be used anywhere else. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure core_unit_counit (F : C ⥤ D) (G : D ⥤ C) := (unit : 𝟭 C ⟶ F.comp G) (counit : G.comp F ⟶ 𝟭 D) diff --git a/src/category_theory/bicategory/free.lean b/src/category_theory/bicategory/free.lean index 6b78babfdbf1c..6528cd11f2bf2 100644 --- a/src/category_theory/bicategory/free.lean +++ b/src/category_theory/bicategory/free.lean @@ -45,7 +45,7 @@ inductive hom : B → B → Type (max u v) instance (a b : B) [inhabited (a ⟶ b)] : inhabited (hom a b) := ⟨hom.of default⟩ /-- Representatives of 2-morphisms in the free bicategory. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] inductive hom₂ : Π {a b : B}, hom a b → hom a b → Type (max u v) | id {a b} (f : hom a b) : hom₂ f f | vcomp {a b} {f g h : hom a b} (η : hom₂ f g) (θ : hom₂ g h) : hom₂ f h diff --git a/src/category_theory/bicategory/functor.lean b/src/category_theory/bicategory/functor.lean index e2697618d3fec..ad5757a9bb35b 100644 --- a/src/category_theory/bicategory/functor.lean +++ b/src/category_theory/bicategory/functor.lean @@ -240,7 +240,7 @@ def comp (F : oplax_functor B C) (G : oplax_functor C D) : oplax_functor B D := A structure on an oplax functor that promotes an oplax functor to a pseudofunctor. See `pseudofunctor.mk_of_oplax`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure pseudo_core (F : oplax_functor B C) := (map_id_iso (a : B) : F.map (𝟙 a) ≅ 𝟙 (F.obj a)) (map_comp_iso {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : F.map (f ≫ g) ≅ F.map f ≫ F.map g) diff --git a/src/category_theory/category/PartialFun.lean b/src/category_theory/category/PartialFun.lean index 2f1349a549b6c..bef9af2a87470 100644 --- a/src/category_theory/category/PartialFun.lean +++ b/src/category_theory/category/PartialFun.lean @@ -39,7 +39,7 @@ namespace PartialFun instance : has_coe_to_sort PartialFun Type* := ⟨id⟩ /-- Turns a type into a `PartialFun`. -/ -@[nolint has_inhabited_instance] def of : Type* → PartialFun := id +@[nolint has_nonempty_instance] def of : Type* → PartialFun := id @[simp] lemma coe_of (X : Type*) : ↥(of X) = X := rfl diff --git a/src/category_theory/concrete_category/bundled.lean b/src/category_theory/concrete_category/bundled.lean index 419621b09a5bf..de675d9604c40 100644 --- a/src/category_theory/concrete_category/bundled.lean +++ b/src/category_theory/concrete_category/bundled.lean @@ -22,7 +22,7 @@ variables {c d : Type u → Type v} {α : Type u} /-- `bundled` is a type bundled with a type class instance for that type. Only the type class is exposed as a parameter. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure bundled (c : Type u → Type v) : Type (max (u+1) v) := (α : Type u) (str : c α . tactic.apply_instance) diff --git a/src/category_theory/core.lean b/src/category_theory/core.lean index 5f583ecb9e1f6..8c55b96351640 100644 --- a/src/category_theory/core.lean +++ b/src/category_theory/core.lean @@ -26,7 +26,7 @@ universes v₁ v₂ u₁ u₂ -- morphism levels before object levels. See note /-- The core of a category C is the groupoid whose morphisms are all the isomorphisms of C. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def core (C : Type u₁) := C variables {C : Type u₁} [category.{v₁} C] diff --git a/src/category_theory/differential_object.lean b/src/category_theory/differential_object.lean index 268eb4e9ff21a..e6a656e86a25c 100644 --- a/src/category_theory/differential_object.lean +++ b/src/category_theory/differential_object.lean @@ -35,7 +35,7 @@ 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`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure differential_object := (X : C) (d : X ⟶ X⟦1⟧) @@ -51,7 +51,7 @@ namespace differential_object /-- A morphism of differential objects is a morphism commuting with the differentials. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure hom (X Y : differential_object C) := (f : X.X ⟶ Y.X) (comm' : X.d ≫ f⟦1⟧' = f ≫ Y.d . obviously) diff --git a/src/category_theory/elements.lean b/src/category_theory/elements.lean index 9a8cb9b3f8d14..2229113de827e 100644 --- a/src/category_theory/elements.lean +++ b/src/category_theory/elements.lean @@ -40,7 +40,7 @@ variables {C : Type u} [category.{v} C] The type of objects for the category of elements of a functor `F : C ⥤ Type` is a pair `(X : C, x : F.obj X)`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def functor.elements (F : C ⥤ Type w) := (Σ c : C, F.obj c) /-- The category structure on `F.elements`, for `F : C ⥤ Type`. diff --git a/src/category_theory/enriched/basic.lean b/src/category_theory/enriched/basic.lean index 1627a7ae4dd6a..dcb15e7119b03 100644 --- a/src/category_theory/enriched/basic.lean +++ b/src/category_theory/enriched/basic.lean @@ -93,7 +93,7 @@ A type synonym for `C`, which should come equipped with a `V`-enriched category In a moment we will equip this with the `W`-enriched category structure obtained by applying the functor `F : lax_monoidal_functor V W` to each hom object. -/ -@[nolint has_inhabited_instance unused_arguments] +@[nolint has_nonempty_instance unused_arguments] def transport_enrichment (F : lax_monoidal_functor V W) (C : Type u₁) := C instance (F : lax_monoidal_functor V W) : @@ -180,7 +180,7 @@ When `V` is any of `Type`, `Top`, `AddCommGroup`, or `Module R`, For `V = Algebra R`, the usual forgetful functor is coyoneda of `polynomial R`, not of `R`. (Perhaps we should have a typeclass for this situation: `concrete_monoidal`?) -/ -@[nolint has_inhabited_instance unused_arguments] +@[nolint has_nonempty_instance unused_arguments] def forget_enrichment (W : Type (v+1)) [category.{v} W] [monoidal_category W] (C : Type u₁) [enriched_category W C] := C @@ -354,7 +354,7 @@ coming from the ambient braiding on `V`.) The type of `A`-graded natural transformations between `V`-functors `F` and `G`. This is the type of morphisms in `V` from `A` to the `V`-object of natural transformations. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure graded_nat_trans (A : center V) (F G : enriched_functor V C D) := (app : Π (X : C), A.1 ⟶ (F.obj X ⟶[V] G.obj X)) (naturality : diff --git a/src/category_theory/essentially_small.lean b/src/category_theory/essentially_small.lean index 82779c2c980b9..66c86072e308b 100644 --- a/src/category_theory/essentially_small.lean +++ b/src/category_theory/essentially_small.lean @@ -39,7 +39,7 @@ lemma essentially_small.mk' {C : Type u} [category.{v} C] {S : Type w} [small_ca /-- An arbitrarily chosen small model for an essentially small category. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def small_model (C : Type u) [category.{v} C] [essentially_small.{w} C] : Type w := classical.some (@essentially_small.equiv_small_category C _ _) @@ -115,7 +115,7 @@ instance locally_small_of_essentially_small We define a type alias `shrink_homs C` for `C`. When we have `locally_small.{w} C`, we'll put a `category.{w}` instance on `shrink_homs C`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def shrink_homs (C : Type u) := C namespace shrink_homs diff --git a/src/category_theory/full_subcategory.lean b/src/category_theory/full_subcategory.lean index 8bbbfda80519a..e34c71dcd3e33 100644 --- a/src/category_theory/full_subcategory.lean +++ b/src/category_theory/full_subcategory.lean @@ -47,7 +47,7 @@ include F which provides a category structure so that the morphisms `X ⟶ Y` are the morphisms in `D` from `F X` to `F Y`. -/ -@[nolint has_inhabited_instance unused_arguments] +@[nolint has_nonempty_instance unused_arguments] def induced_category : Type u₁ := C variables {D} diff --git a/src/category_theory/glue_data.lean b/src/category_theory/glue_data.lean index bd2220c9a6b92..15b95f35be6b7 100644 --- a/src/category_theory/glue_data.lean +++ b/src/category_theory/glue_data.lean @@ -42,7 +42,7 @@ such that `t' : V i j ×[U i] V i k ⟶ V j k ×[U j] V j i`. 10. `t' i j k ≫ t' j k i ≫ t' k i j = 𝟙 _`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data := (J : Type v) (U : J → C) diff --git a/src/category_theory/grothendieck.lean b/src/category_theory/grothendieck.lean index 12d8105a96533..c304cd2f89a7f 100644 --- a/src/category_theory/grothendieck.lean +++ b/src/category_theory/grothendieck.lean @@ -48,7 +48,7 @@ gives a category whose `base : X.base ⟶ Y.base` and `f.fiber : (F.map base).obj X.fiber ⟶ Y.fiber` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure grothendieck := (base : C) (fiber : F.obj base) diff --git a/src/category_theory/idempotents/karoubi.lean b/src/category_theory/idempotents/karoubi.lean index 40f9de580a7cc..45d790af537da 100644 --- a/src/category_theory/idempotents/karoubi.lean +++ b/src/category_theory/idempotents/karoubi.lean @@ -41,7 +41,7 @@ obvious idempotent `X ⟶ P ⟶ X` which is the projection onto `P` with kernel one may define a formal direct factor of an object `X : C` : it consists of an idempotent `p : X ⟶ X` which is thought as the "formal image" of `p`. The type `karoubi C` shall be the type of the objects of the karoubi enveloppe of `C`. It makes sense for any category `C`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure karoubi := (X : C) (p : X ⟶ X) (idem : p ≫ p = p) namespace karoubi diff --git a/src/category_theory/limits/has_limits.lean b/src/category_theory/limits/has_limits.lean index 37f6c18894d68..d2d87212c65e4 100644 --- a/src/category_theory/limits/has_limits.lean +++ b/src/category_theory/limits/has_limits.lean @@ -69,7 +69,7 @@ variables {F : J ⥤ C} section limit /-- `limit_cone F` contains a cone over `F` together with the information that it is a limit. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure limit_cone (F : J ⥤ C) := (cone : cone F) (is_limit : is_limit cone) @@ -519,7 +519,7 @@ section colimit /-- `colimit_cocone F` contains a cocone over `F` together with the information that it is a colimit. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure colimit_cocone (F : J ⥤ C) := (cocone : cocone F) (is_colimit : is_colimit cocone) diff --git a/src/category_theory/limits/is_limit.lean b/src/category_theory/limits/is_limit.lean index 9b220c75b965f..56ce006946f2a 100644 --- a/src/category_theory/limits/is_limit.lean +++ b/src/category_theory/limits/is_limit.lean @@ -50,7 +50,7 @@ cone morphism to `t`. See . -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure is_limit (t : cone F) := (lift : Π (s : cone F), s.X ⟶ t.X) (fac' : ∀ (s : cone F) (j : J), lift s ≫ t.π.app j = s.π.app j . obviously) @@ -504,7 +504,7 @@ cocone morphism from `t`. See . -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure is_colimit (t : cocone F) := (desc : Π (s : cocone F), t.X ⟶ s.X) (fac' : ∀ (s : cocone F) (j : J), t.ι.app j ≫ desc s = s.ι.app j . obviously) diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index 6cec24a59572b..7613d92dc0690 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -74,7 +74,7 @@ A `c : bicone F` is: * morphisms `π j : X ⟶ F j` and `ι j : F j ⟶ X` for each `j`, * such that `ι j ≫ π j'` is the identity when `j = j'` and zero otherwise. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure bicone (F : J → C) := (X : C) (π : Π j, X ⟶ F j) @@ -145,7 +145,7 @@ lemma π_of_is_colimit {f : J → C} {t : bicone f} (ht : is_colimit t.to_cocone ht.hom_ext (λ j', by { rw ht.fac, discrete_cases, simp [t.ι_π] }) /-- Structure witnessing that a bicone is both a limit cone and a colimit cocone. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure is_bilimit {F : J → C} (B : bicone F) := (is_limit : is_limit B.to_cone) (is_colimit : is_colimit B.to_cocone) @@ -212,7 +212,7 @@ end bicone /-- A bicone over `F : J → C`, which is both a limit cone and a colimit cocone. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure limit_bicone (F : J → C) := (bicone : bicone F) (is_bilimit : bicone.is_bilimit) @@ -740,7 +740,7 @@ A binary bicone for a pair of objects `P Q : C` consists of the cone point `X`, maps from `X` to both `P` and `Q`, and maps from both `P` and `Q` to `X`, so that `inl ≫ fst = 𝟙 P`, `inl ≫ snd = 0`, `inr ≫ fst = 0`, and `inr ≫ snd = 𝟙 Q` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure binary_bicone (P Q : C) := (X : C) (fst : X ⟶ P) @@ -861,7 +861,7 @@ is_colimit.equiv_iso_colimit $ cocones.ext (iso.refl _) (λ j, by { rcases j wit end bicone /-- Structure witnessing that a binary bicone is a limit cone and a limit cocone. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure binary_bicone.is_bilimit {P Q : C} (b : binary_bicone P Q) := (is_limit : is_limit b.to_cone) (is_colimit : is_colimit b.to_cocone) @@ -888,7 +888,7 @@ def bicone.to_binary_bicone_is_bilimit {X Y : C} (b : bicone (pair_function X Y) /-- A bicone over `P Q : C`, which is both a limit cone and a colimit cocone. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure binary_biproduct_data (P Q : C) := (bicone : binary_bicone P Q) (is_bilimit : bicone.is_bilimit) diff --git a/src/category_theory/limits/shapes/multiequalizer.lean b/src/category_theory/limits/shapes/multiequalizer.lean index 7292893ebb255..5cc5edaf285db 100644 --- a/src/category_theory/limits/shapes/multiequalizer.lean +++ b/src/category_theory/limits/shapes/multiequalizer.lean @@ -112,7 +112,7 @@ instance : small_category (walking_multispan fst snd) := end walking_multispan /-- This is a structure encapsulating the data necessary to define a `multicospan`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure multicospan_index (C : Type u) [category.{v} C] := (L R : Type w) (fst_to snd_to : R → L) @@ -122,7 +122,7 @@ structure multicospan_index (C : Type u) [category.{v} C] := (snd : Π b, left (snd_to b) ⟶ right b) /-- This is a structure encapsulating the data necessary to define a `multispan`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure multispan_index (C : Type u) [category.{v} C] := (L R : Type w) (fst_from snd_from : L → R) @@ -252,11 +252,11 @@ end multispan_index variables {C : Type u} [category.{v} C] /-- A multifork is a cone over a multicospan. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] abbreviation multifork (I : multicospan_index C) := cone I.multicospan /-- A multicofork is a cocone over a multispan. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] abbreviation multicofork (I : multispan_index C) := cocone I.multispan namespace multifork diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index 480590a151641..1fb3086376ac7 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -388,7 +388,7 @@ variables (f : X ⟶ Z) (g : Y ⟶ Z) The usual explicit pullback in the category of types, as a subtype of the product. The full `limit_cone` data is bundled as `pullback_limit_cone f g`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] abbreviation pullback_obj : Type u := { p : X × Y // f p.1 = g p.2 } -- `pullback_obj f g` comes with a coercion to the product type `X × Y`. diff --git a/src/category_theory/limits/types.lean b/src/category_theory/limits/types.lean index adf4696e1c695..f02267b171972 100644 --- a/src/category_theory/limits/types.lean +++ b/src/category_theory/limits/types.lean @@ -194,7 +194,7 @@ A quotient type implementing the colimit of a functor `F : J ⥤ Type u`, as pairs `⟨j, x⟩` where `x : F.obj j`, modulo the equivalence relation generated by `⟨j, x⟩ ~ ⟨j', x'⟩` whenever there is a morphism `f : j ⟶ j'` so `F.map f x = x'`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def quot (F : J ⥤ Type (max v u)) : Type (max v u) := @quot (Σ j, F.obj j) (quot.rel F) diff --git a/src/category_theory/localization/construction.lean b/src/category_theory/localization/construction.lean index 112084a4bf505..a49d234daee08 100644 --- a/src/category_theory/localization/construction.lean +++ b/src/category_theory/localization/construction.lean @@ -75,7 +75,7 @@ namespace construction /-- If `W : morphism_property C`, `loc_quiver W` is a quiver with the same objects as `C`, and whose morphisms are those in `C` and placeholders for formal inverses of the morphisms in `W`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure loc_quiver (W : morphism_property C) := (obj : C) instance : quiver (loc_quiver W) := @@ -115,7 +115,7 @@ open localization.construction /-- The localized category obtained by formally inverting the morphisms in `W : morphism_property C` -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def localization := category_theory.quotient (localization.construction.relations W) /-- The obvious functor `C ⥤ W.localization` -/ diff --git a/src/category_theory/monad/algebra.lean b/src/category_theory/monad/algebra.lean index 0c62df9578e4e..198f5f83f93e2 100644 --- a/src/category_theory/monad/algebra.lean +++ b/src/category_theory/monad/algebra.lean @@ -241,7 +241,7 @@ end monad namespace comonad /-- An Eilenberg-Moore coalgebra for a comonad `T`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure coalgebra (G : comonad C) : Type (max u₁ v₁) := (A : C) (a : A ⟶ (G : C ⥤ C).obj A) @@ -256,7 +256,7 @@ namespace coalgebra variables {G : comonad C} /-- A morphism of Eilenberg-Moore coalgebras for the comonad `G`. -/ -@[ext, nolint has_inhabited_instance] structure hom (A B : coalgebra G) := +@[ext, nolint has_nonempty_instance] structure hom (A B : coalgebra G) := (f : A.A ⟶ B.A) (h' : A.a ≫ (G : C ⥤ C).map f = f ≫ B.a . obviously) diff --git a/src/category_theory/monoidal/center.lean b/src/category_theory/monoidal/center.lean index 4ffea5d7850a4..4c6eb069a90e3 100644 --- a/src/category_theory/monoidal/center.lean +++ b/src/category_theory/monoidal/center.lean @@ -47,7 +47,7 @@ Thinking of `C` as a 2-category with a single `0`-morphism, these are the same a transformations (in the pseudo- sense) of the identity 2-functor on `C`, which send the unique `0`-morphism to `X`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure half_braiding (X : C) := (β : Π U, X ⊗ U ≅ U ⊗ X) (monoidal' : ∀ U U', (β (U ⊗ U')).hom = @@ -65,7 +65,7 @@ variables (C) The Drinfeld center of a monoidal category `C` has as objects pairs `⟨X, b⟩`, where `X : C` and `b` is a half-braiding on `X`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def center := Σ X : C, half_braiding X namespace center @@ -73,7 +73,7 @@ namespace center variables {C} /-- A morphism in the Drinfeld center of `C`. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure hom (X Y : center C) := (f : X.1 ⟶ Y.1) (comm' : ∀ U, (f ⊗ 𝟙 U) ≫ (Y.2.β U).hom = (X.2.β U).hom ≫ (𝟙 U ⊗ f) . obviously) diff --git a/src/category_theory/monoidal/free/basic.lean b/src/category_theory/monoidal/free/basic.lean index 4b94068049a12..dd133c7f02b1b 100644 --- a/src/category_theory/monoidal/free/basic.lean +++ b/src/category_theory/monoidal/free/basic.lean @@ -52,7 +52,7 @@ namespace free_monoidal_category /-- Formal compositions and tensor products of identities, unitors and associators. The morphisms of the free monoidal category are obtained as a quotient of these formal morphisms by the relations defining a monoidal category. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] inductive hom : F C → F C → Type u | id (X) : hom X X | α_hom (X Y Z : F C) : hom ((X.tensor Y).tensor Z) (X.tensor (Y.tensor Z)) diff --git a/src/category_theory/monoidal/free/coherence.lean b/src/category_theory/monoidal/free/coherence.lean index 2d28aae09bc78..a1314d9cb820a 100644 --- a/src/category_theory/monoidal/free/coherence.lean +++ b/src/category_theory/monoidal/free/coherence.lean @@ -49,7 +49,7 @@ variables (C) /-- We say an object in the free monoidal category is in normal form if it is of the form `(((𝟙_ C) ⊗ X₁) ⊗ X₂) ⊗ ⋯`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] inductive normal_monoidal_object : Type u | unit : normal_monoidal_object | tensor : normal_monoidal_object → C → normal_monoidal_object diff --git a/src/category_theory/monoidal/of_chosen_finite_products.lean b/src/category_theory/monoidal/of_chosen_finite_products.lean index 9e822085b515c..9191730e2f057 100644 --- a/src/category_theory/monoidal/of_chosen_finite_products.lean +++ b/src/category_theory/monoidal/of_chosen_finite_products.lean @@ -321,7 +321,7 @@ a fixed choice of limit data for the empty functor, and for `pair X Y` for every This is an implementation detail for `symmetric_of_chosen_finite_products`. -/ -@[derive category, nolint unused_arguments has_inhabited_instance] +@[derive category, nolint unused_arguments has_nonempty_instance] def monoidal_of_chosen_finite_products_synonym (𝒯 : limit_cone (functor.empty.{v} C)) (ℬ : Π (X Y : C), limit_cone (pair X Y)):= C diff --git a/src/category_theory/monoidal/opposite.lean b/src/category_theory/monoidal/opposite.lean index 28f98b1404a0d..fcb952d416984 100644 --- a/src/category_theory/monoidal/opposite.lean +++ b/src/category_theory/monoidal/opposite.lean @@ -21,7 +21,7 @@ namespace category_theory open category_theory.monoidal_category /-- A type synonym for the monoidal opposite. Use the notation `Cᴹᵒᵖ`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def monoidal_opposite (C : Type u₁) := C namespace monoidal_opposite diff --git a/src/category_theory/preadditive/Mat.lean b/src/category_theory/preadditive/Mat.lean index 6e0685e0c24c7..e3a1f83a30298 100644 --- a/src/category_theory/preadditive/Mat.lean +++ b/src/category_theory/preadditive/Mat.lean @@ -71,7 +71,7 @@ namespace Mat_ variables {C} /-- A morphism in `Mat_ C` is a dependently typed matrix of morphisms. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def hom (M N : Mat_ C) : Type v₁ := dmatrix M.ι N.ι (λ i j, M.X i ⟶ N.X j) namespace hom diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index 78590a36a8d41..bf80733103044 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -145,7 +145,7 @@ section variables (C D : Type*) [category C] [category D] [preadditive C] [preadditive D] /-- Bundled additive functors. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def AdditiveFunctor := { F : C ⥤ D // functor.additive F } diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index 69c7841196cbc..f36beaf2db668 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -37,7 +37,7 @@ section An injective presentation of an object `X` consists of a monomorphism `f : X ⟶ J` to some injective object `J`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure injective_presentation (X : C) := (J : C) (injective : injective J . tactic.apply_instance) diff --git a/src/category_theory/preadditive/injective_resolution.lean b/src/category_theory/preadditive/injective_resolution.lean index 1ef24b230c5a9..d002bb53a9b1e 100644 --- a/src/category_theory/preadditive/injective_resolution.lean +++ b/src/category_theory/preadditive/injective_resolution.lean @@ -49,7 +49,7 @@ you will not typically need to use this bundled object, and will instead use `injective_resolution Z` (all the components are equipped with `mono` instances, and when the category is `abelian` we will show `ι` is a quasi-iso). -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure InjectiveResolution (Z : C) := (cocomplex : cochain_complex C ℕ) (ι: ((cochain_complex.single₀ C).obj Z) ⟶ cocomplex) diff --git a/src/category_theory/preadditive/projective.lean b/src/category_theory/preadditive/projective.lean index aa8ecaf35dcc0..5fb7122a8b522 100644 --- a/src/category_theory/preadditive/projective.lean +++ b/src/category_theory/preadditive/projective.lean @@ -45,7 +45,7 @@ section A projective presentation of an object `X` consists of an epimorphism `f : P ⟶ X` from some projective object `P`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure projective_presentation (X : C) := (P : C) (projective : projective P . tactic.apply_instance) diff --git a/src/category_theory/preadditive/projective_resolution.lean b/src/category_theory/preadditive/projective_resolution.lean index c61f01a3c8343..d9f1d8c7636c5 100644 --- a/src/category_theory/preadditive/projective_resolution.lean +++ b/src/category_theory/preadditive/projective_resolution.lean @@ -66,7 +66,7 @@ you will not typically need to use this bundled object, and will instead use `(single C _ 0).obj Z` (all the components are equipped with `epi` instances, and when the category is `abelian` we will show `π` is a quasi-iso). -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure ProjectiveResolution (Z : C) := (complex : chain_complex C ℕ) (π : homological_complex.hom complex ((chain_complex.single₀ C).obj Z)) diff --git a/src/category_theory/shift.lean b/src/category_theory/shift.lean index c61713108faab..51d671b5b7f8c 100644 --- a/src/category_theory/shift.lean +++ b/src/category_theory/shift.lean @@ -77,7 +77,7 @@ 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_inhabited_instance] +@[nolint has_nonempty_instance] structure shift_mk_core := (F : A → (C ⥤ C)) (ε : 𝟭 C ≅ F 0) diff --git a/src/category_theory/single_obj.lean b/src/category_theory/single_obj.lean index 033642514465a..4e7c97bd837d8 100644 --- a/src/category_theory/single_obj.lean +++ b/src/category_theory/single_obj.lean @@ -38,7 +38,7 @@ universes u v w namespace category_theory /-- Type tag on `unit` used to define single-object categories and groupoids. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def single_obj (α : Type u) : Type := unit namespace single_obj diff --git a/src/category_theory/sites/cover_lifting.lean b/src/category_theory/sites/cover_lifting.lean index 6e3eb8f840867..deca91fbd087b 100644 --- a/src/category_theory/sites/cover_lifting.lean +++ b/src/category_theory/sites/cover_lifting.lean @@ -60,7 +60,7 @@ variables {L : grothendieck_topology E} A functor `G : (C, J) ⥤ (D, K)` between sites is called to have the cover-lifting property if for all covering sieves `R` in `D`, `R.pullback G` is a covering sieve in `C`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure cover_lifting (G : C ⥤ D) : Prop := (cover_lift : ∀ {U : C} {S : sieve (G.obj U)} (hS : S ∈ K (G.obj U)), S.functor_pullback G ∈ J U) diff --git a/src/category_theory/sites/cover_preserving.lean b/src/category_theory/sites/cover_preserving.lean index 753a7a0ad6e21..c0d61ab099427 100644 --- a/src/category_theory/sites/cover_preserving.lean +++ b/src/category_theory/sites/cover_preserving.lean @@ -62,7 +62,7 @@ variables {L : grothendieck_topology A} A functor `G : (C, J) ⥤ (D, K)` between sites is *cover-preserving* if for all covering sieves `R` in `C`, `R.pushforward_functor G` is a covering sieve in `D`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure cover_preserving (G : C ⥤ D) : Prop := (cover_preserve : ∀ {U : C} {S : sieve U} (hS : S ∈ J U), S.functor_pushforward G ∈ K (G.obj U)) @@ -86,7 +86,7 @@ compatible family of elements at `C` and valued in `G.op ⋙ ℱ`, and each comm This is actually stronger than merely preserving compatible families because of the definition of `functor_pushforward` used. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure compatible_preserving (K : grothendieck_topology D) (G : C ⥤ D) : Prop := (compatible : ∀ (ℱ : SheafOfTypes.{w} K) {Z} {T : presieve Z} diff --git a/src/category_theory/sites/dense_subsite.lean b/src/category_theory/sites/dense_subsite.lean index d7498c315fffa..ceac0e5660ff7 100644 --- a/src/category_theory/sites/dense_subsite.lean +++ b/src/category_theory/sites/dense_subsite.lean @@ -51,7 +51,7 @@ variables {L : grothendieck_topology E} /-- An auxiliary structure that witnesses the fact that `f` factors through an image object of `G`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure presieve.cover_by_image_structure (G : C ⥤ D) {V U : D} (f : V ⟶ U) := (obj : C) (lift : V ⟶ G.obj obj) diff --git a/src/category_theory/sites/grothendieck.lean b/src/category_theory/sites/grothendieck.lean index ebba468fb34db..658de4af3c30f 100644 --- a/src/category_theory/sites/grothendieck.lean +++ b/src/category_theory/sites/grothendieck.lean @@ -406,14 +406,14 @@ instance : semilattice_inf (J.cover X) := instance : inhabited (J.cover X) := ⟨⊤⟩ /-- An auxiliary structure, used to define `S.index` in `plus.lean`. -/ -@[nolint has_inhabited_instance, ext] +@[nolint has_nonempty_instance, ext] structure arrow (S : J.cover X) := (Y : C) (f : Y ⟶ X) (hf : S f) /-- An auxiliary structure, used to define `S.index` in `plus.lean`. -/ -@[nolint has_inhabited_instance, ext] +@[nolint has_nonempty_instance, ext] structure relation (S : J.cover X) := (Y₁ Y₂ Z : C) (g₁ : Z ⟶ Y₁) diff --git a/src/category_theory/sites/sheafification.lean b/src/category_theory/sites/sheafification.lean index fc5fb5a349270..08aeb21ebdf07 100644 --- a/src/category_theory/sites/sheafification.lean +++ b/src/category_theory/sites/sheafification.lean @@ -34,7 +34,7 @@ local attribute [instance] concrete_category.has_coe_to_fun /-- A concrete version of the multiequalizer, to be used below. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def meq {X : C} (P : Cᵒᵖ ⥤ D) (S : J.cover X) := { x : Π (I : S.arrow), P.obj (op I.Y) // ∀ (I : S.relation), P.map I.g₁.op (x I.fst) = P.map I.g₂.op (x I.snd) } diff --git a/src/category_theory/sites/sieves.lean b/src/category_theory/sites/sieves.lean index 41522c7061f28..0eaba6acc1718 100644 --- a/src/category_theory/sites/sieves.lean +++ b/src/category_theory/sites/sieves.lean @@ -170,7 +170,7 @@ def functor_pushforward (S : presieve X) : presieve (F.obj X) := /-- An auxillary definition in order to fix the choice of the preimages between various definitions. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure functor_pushforward_structure (S : presieve X) {Y} (f : Y ⟶ F.obj X) := (preobj : C) (premap : preobj ⟶ X) (lift : Y ⟶ F.obj preobj) (cover : S premap) (fac : f = lift ≫ F.map premap) diff --git a/src/category_theory/structured_arrow.lean b/src/category_theory/structured_arrow.lean index 5b765d7c15d3b..4290ae69a837f 100644 --- a/src/category_theory/structured_arrow.lean +++ b/src/category_theory/structured_arrow.lean @@ -29,7 +29,7 @@ The category of `T`-structured arrows with domain `S : D` (here `T : C ⥤ D`), has as its objects `D`-morphisms of the form `S ⟶ T Y`, for some `Y : C`, and morphisms `C`-morphisms `Y ⟶ Y'` making the obvious triangle commute. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def structured_arrow (S : D) (T : C ⥤ D) := comma (functor.from_punit S) T namespace structured_arrow @@ -145,7 +145,7 @@ The category of `S`-costructured arrows with target `T : D` (here `S : C ⥤ D`) has as its objects `D`-morphisms of the form `S Y ⟶ T`, for some `Y : C`, and morphisms `C`-morphisms `Y ⟶ Y'` making the obvious triangle commute. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def costructured_arrow (S : C ⥤ D) (T : D) := comma S (functor.from_punit T) namespace costructured_arrow diff --git a/src/category_theory/with_terminal.lean b/src/category_theory/with_terminal.lean index b3e3142e36c5c..28b0f21f32763 100644 --- a/src/category_theory/with_terminal.lean +++ b/src/category_theory/with_terminal.lean @@ -56,7 +56,7 @@ local attribute [tidy] tactic.case_bash variable {C} /-- Morphisms for `with_terminal C`. -/ -@[simp, nolint has_inhabited_instance] +@[simp, nolint has_nonempty_instance] def hom : with_terminal C → with_terminal C → Type v | (of X) (of Y) := X ⟶ Y | star (of X) := pempty @@ -213,7 +213,7 @@ local attribute [tidy] tactic.case_bash variable {C} /-- Morphisms for `with_initial C`. -/ -@[simp, nolint has_inhabited_instance] +@[simp, nolint has_nonempty_instance] def hom : with_initial C → with_initial C → Type v | (of X) (of Y) := X ⟶ Y | (of X) _ := pempty diff --git a/src/combinatorics/quiver/basic.lean b/src/combinatorics/quiver/basic.lean index c0f0e947c10b3..37c9b5a9eb945 100644 --- a/src/combinatorics/quiver/basic.lean +++ b/src/combinatorics/quiver/basic.lean @@ -92,7 +92,7 @@ def hom.unop {V} [quiver V] {X Y : Vᵒᵖ} (f : X ⟶ Y) : unop Y ⟶ unop X := attribute [irreducible] quiver.opposite /-- A type synonym for a quiver with no arrows. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def empty (V) : Type u := V instance empty_quiver (V : Type u) : quiver.{u} (empty V) := ⟨λ a b, pempty⟩ diff --git a/src/combinatorics/quiver/connected_component.lean b/src/combinatorics/quiver/connected_component.lean index b43d6b4b177bc..6ce1276737e98 100644 --- a/src/combinatorics/quiver/connected_component.lean +++ b/src/combinatorics/quiver/connected_component.lean @@ -23,7 +23,7 @@ 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_inhabited_instance] +@[nolint has_nonempty_instance] def symmetrify (V) : Type u := V instance symmetrify_quiver (V : Type u) [quiver V] : quiver (symmetrify V) := diff --git a/src/combinatorics/quiver/subquiver.lean b/src/combinatorics/quiver/subquiver.lean index 812fff2a73b8e..551a2a2eaad7e 100644 --- a/src/combinatorics/quiver/subquiver.lean +++ b/src/combinatorics/quiver/subquiver.lean @@ -25,7 +25,7 @@ def wide_subquiver (V) [quiver.{v+1} V] := /-- A type synonym for `V`, when thought of as a quiver having only the arrows from some `wide_subquiver`. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def wide_subquiver.to_Type (V) [quiver V] (H : wide_subquiver V) : Type u := V instance wide_subquiver_has_coe_to_sort {V} [quiver V] : @@ -44,7 +44,7 @@ instance {V} [quiver V] : inhabited (wide_subquiver V) := ⟨⊤⟩ /-- `total V` is the type of _all_ arrows of `V`. -/ -- TODO Unify with `category_theory.arrow`? (The fields have been named to match.) -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure total (V : Type u) [quiver.{v} V] : Sort (max (u+1) v) := (left : V) (right : V) diff --git a/src/data/bundle.lean b/src/data/bundle.lean index 5f9810be8eca6..946d03f62ac00 100644 --- a/src/data/bundle.lean +++ b/src/data/bundle.lean @@ -89,7 +89,7 @@ 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_inhabited_instance] def pullback (f : B' → B) (E : B → Type*) := λ x, E (f x) +@[nolint has_nonempty_instance] def pullback (f : B' → B) (E : B → Type*) := λ x, E (f x) notation f ` *ᵖ ` E := pullback f E diff --git a/src/data/multiset/fintype.lean b/src/data/multiset/fintype.lean index f9e8388c856ba..70830824d1bb2 100644 --- a/src/data/multiset/fintype.lean +++ b/src/data/multiset/fintype.lean @@ -37,7 +37,7 @@ variables {α : Type*} [decidable_eq α] {m : multiset α} /-- Auxiliary definition for the `has_coe_to_sort` instance. This prevents the `has_coe m α` instance from inadverently applying to other sigma types. One should not use this definition directly. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def multiset.to_type (m : multiset α) : Type* := Σ (x : α), fin (m.count x) /-- Create a type that has the same number of elements as the multiset. diff --git a/src/data/pfunctor/multivariate/W.lean b/src/data/pfunctor/multivariate/W.lean index dbb2b2165b607..3923df2b0c67f 100644 --- a/src/data/pfunctor/multivariate/W.lean +++ b/src/data/pfunctor/multivariate/W.lean @@ -114,7 +114,7 @@ def Wp : mvpfunctor n := { A := P.last.W, B := P.W_path } /-- W-type of `P` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def W (α : typevec n) : Type* := P.Wp.obj α instance mvfunctor_W : mvfunctor P.W := by delta mvpfunctor.W; apply_instance diff --git a/src/data/pfunctor/univariate/basic.lean b/src/data/pfunctor/univariate/basic.lean index e4aeca9a721ca..ba63040d31e2b 100644 --- a/src/data/pfunctor/univariate/basic.lean +++ b/src/data/pfunctor/univariate/basic.lean @@ -65,7 +65,7 @@ def W := _root_.W_type P.B /- inhabitants of W types is awkward to encode as an instance assumption because there needs to be a value `a : P.A` such that `P.B a` is empty to yield a finite tree -/ -attribute [nolint has_inhabited_instance] W +attribute [nolint has_nonempty_instance] W variables {P} /-- root element of a W tree -/ diff --git a/src/data/qpf/multivariate/constructions/fix.lean b/src/data/qpf/multivariate/constructions/fix.lean index 2d6733aa52772..ef1470eae2ebb 100644 --- a/src/data/qpf/multivariate/constructions/fix.lean +++ b/src/data/qpf/multivariate/constructions/fix.lean @@ -184,7 +184,7 @@ fix F a b = F a b (fix F a b) def fix {n : ℕ} (F : typevec (n+1) → Type*) [mvfunctor F] [q : mvqpf F] (α : typevec n) := quotient (W_setoid α : setoid (q.P.W α)) -attribute [nolint has_inhabited_instance] fix +attribute [nolint has_nonempty_instance] fix /-- `fix F` is a functor -/ def fix.map {α β : typevec n} (g : α ⟹ β) : fix F α → fix F β := diff --git a/src/data/qpf/univariate/basic.lean b/src/data/qpf/univariate/basic.lean index 271577b624e48..94c474a2a2055 100644 --- a/src/data/qpf/univariate/basic.lean +++ b/src/data/qpf/univariate/basic.lean @@ -213,7 +213,7 @@ def W_setoid : setoid q.P.W := local attribute [instance] W_setoid /-- inductive type defined as initial algebra of a Quotient of Polynomial Functor -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def fix (F : Type u → Type u) [functor F] [q : qpf F] := quotient (W_setoid : setoid q.P.W) /-- recursor of a type defined by a qpf -/ diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index e57ae70149fcb..b54cb95cf588c 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -648,7 +648,7 @@ end charted_space have a topological structure, where the topology would come from the charts. For this, one needs charts that are only local equivs, and continuity properties for their composition. This is formalised in `charted_space_core`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure charted_space_core (H : Type*) [topological_space H] (M : Type*) := (atlas : set (local_equiv M H)) (chart_at : M → local_equiv M H) @@ -939,7 +939,7 @@ end topological_space.opens /-- A `G`-diffeomorphism between two charted spaces is a homeomorphism which, when read in the charts, belongs to `G`. We avoid the word diffeomorph as it is too related to the smooth category, and use structomorph instead. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure structomorph (G : structure_groupoid H) (M : Type*) (M' : Type*) [topological_space M] [topological_space M'] [charted_space H M] [charted_space H M'] extends homeomorph M M' := diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 72a80945f79fb..a1b654d4ce246 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -70,7 +70,7 @@ variables (I I' M M' n) /-- `n`-times continuously differentiable diffeomorphism between `M` and `M'` with respect to I and I' -/ -@[protect_proj, nolint has_inhabited_instance] +@[protect_proj, nolint has_nonempty_instance] structure diffeomorph extends M ≃ M' := (cont_mdiff_to_fun : cont_mdiff I I' n to_equiv) (cont_mdiff_inv_fun : cont_mdiff I' I n to_equiv.symm) diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index 778490f7c8900..d5c795f33cc34 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -97,7 +97,7 @@ subordinate to `U`, see `smooth_bump_covering.exists_is_subordinate`. This covering can be used, e.g., to construct a partition of unity and to prove the weak Whitney embedding theorem. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure smooth_bump_covering (s : set M := univ) := (c : ι → M) (to_fun : Π i, smooth_bump_function I (c i)) diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index abc1771348f2c..d6f4a9de3d345 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -125,7 +125,7 @@ section model_with_corners model vector space `E` over the field `𝕜`. This is all what is needed to define a smooth manifold with model space `H`, and model vector space `E`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure model_with_corners (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] extends local_equiv H E := diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean index a0771fdfc9d5d..b87cefe50a854 100644 --- a/src/geometry/manifold/tangent_bundle.lean +++ b/src/geometry/manifold/tangent_bundle.lean @@ -473,7 +473,7 @@ 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_inhabited_instance, reducible] -- is empty if the base manifold is empty +@[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 diff --git a/src/group_theory/free_product.lean b/src/group_theory/free_product.lean index beb73b8706d7e..a20b6ee73241f 100644 --- a/src/group_theory/free_product.lean +++ b/src/group_theory/free_product.lean @@ -380,7 +380,7 @@ variable (M) /-- A `neword M i j` is a representation of a non-empty reduced words where the first letter comes from `M i` and the last letter comes from `M j`. It can be constructed from singletons and via concatentation, and thus provides a useful induction principle. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] inductive neword : ι → ι → Type (max u_1 u_2) | singleton : ∀ {i} (x : M i) (hne1 : x ≠ 1), neword i i | append : ∀ {i j k l} (w₁ : neword i j) (hne : j ≠ k) (w₂ : neword k l), neword i l diff --git a/src/group_theory/monoid_localization.lean b/src/group_theory/monoid_localization.lean index 3b6394173eb0e..b797ad903101e 100644 --- a/src/group_theory/monoid_localization.lean +++ b/src/group_theory/monoid_localization.lean @@ -68,7 +68,7 @@ variables {M : Type*} [add_comm_monoid M] (S : add_submonoid M) (N : Type*) [add /-- The type of add_monoid homomorphisms satisfying the characteristic predicate: if `f : M →+ N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ -@[nolint has_inhabited_instance] structure localization_map +@[nolint has_nonempty_instance] structure localization_map 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) @@ -88,7 +88,7 @@ namespace submonoid /-- The type of monoid homomorphisms satisfying the characteristic predicate: if `f : M →* N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ -@[nolint has_inhabited_instance] structure localization_map +@[nolint has_nonempty_instance] structure localization_map 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) @@ -1334,7 +1334,7 @@ namespace submonoid /-- The type of homomorphisms between monoids with zero satisfying the characteristic predicate: if `f : M →*₀ N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ -@[nolint has_inhabited_instance] structure localization_with_zero_map +@[nolint has_nonempty_instance] structure localization_with_zero_map extends localization_map S N := (map_zero' : to_fun 0 = 0) diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index 6d863349658e4..c36283d3646de 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -56,7 +56,7 @@ open category_theory category_theory.action_category category_theory.single_obj /-- `is_free_groupoid.generators G` is a type synonym for `G`. We think of this as the vertices of the generating quiver of `G` when `G` is free. We can't use `G` directly, since `G` already has a quiver instance from being a groupoid. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def is_free_groupoid.generators (G) [groupoid G] := G /-- A groupoid `G` is free when we have the following data: diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 7177d57104ed5..646f3b099e444 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -39,7 +39,7 @@ and inverse maps are affine. We define it using an `equiv` for the map and a `linear_equiv` for the linear part in order to allow affine equivalences with good definitional equalities. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure affine_equiv (k P₁ P₂ : Type*) {V₁ V₂ : Type*} [ring k] [add_comm_group V₁] [module k V₁] [add_torsor V₁ P₁] [add_comm_group V₂] [module k V₂] [add_torsor V₂ P₂] extends P₁ ≃ P₂ := diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 96ef0d5e4fdbd..854ecd539f682 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -354,7 +354,7 @@ variables {R M ι : Type*} variables [comm_semiring R] [add_comm_monoid M] [module R M] [decidable_eq ι] /-- `e` and `ε` have characteristic properties of a basis and its dual -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure dual_pair (e : ι → M) (ε : ι → (dual R M)) := (eval : ∀ i j : ι, ε i (e j) = if i = j then 1 else 0) (total : ∀ {m : M}, (∀ i, ε i m = 0) → m = 0) diff --git a/src/linear_algebra/free_module/basic.lean b/src/linear_algebra/free_module/basic.lean index f90e7bc658b6a..8374d02967877 100644 --- a/src/linear_algebra/free_module/basic.lean +++ b/src/linear_algebra/free_module/basic.lean @@ -67,7 +67,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_inhabited_instance] def choose_basis_index := (exists_basis R M).some.1 +@[nolint has_nonempty_instance] 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`. -/ diff --git a/src/linear_algebra/free_module/pid.lean b/src/linear_algebra/free_module/pid.lean index 5f1726359916b..73df194265e59 100644 --- a/src/linear_algebra/free_module/pid.lean +++ b/src/linear_algebra/free_module/pid.lean @@ -405,7 +405,7 @@ section smith_normal /-- A Smith normal form basis for a submodule `N` of a module `M` consists of bases for `M` and `N` such that the inclusion map `N → M` can be written as a (rectangular) matrix with `a` along the diagonal: in Smith normal form. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure basis.smith_normal_form (N : submodule R M) (ι : Type*) (n : ℕ) := (bM : basis ι R M) (bN : basis (fin n) R N) diff --git a/src/linear_algebra/matrix/transvection.lean b/src/linear_algebra/matrix/transvection.lean index 2e130f81624ba..1265eb6799ad4 100644 --- a/src/linear_algebra/matrix/transvection.lean +++ b/src/linear_algebra/matrix/transvection.lean @@ -132,7 +132,7 @@ variables (R n) /-- A structure containing all the information from which one can build a nontrivial transvection. This structure is easier to manipulate than transvections as one has a direct access to all the relevant fields. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure transvection_struct := (i j : n) (hij : i ≠ j) diff --git a/src/linear_algebra/projective_space/basic.lean b/src/linear_algebra/projective_space/basic.lean index 4be3ed2e0e4a0..cde97cc906b6b 100644 --- a/src/linear_algebra/projective_space/basic.lean +++ b/src/linear_algebra/projective_space/basic.lean @@ -43,7 +43,7 @@ def projectivization_setoid : setoid { v : V // v ≠ 0 } := /-- The projectivization of the `K`-vector space `V`. The notation `ℙ K V` is preferred. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def projectivization := quotient (projectivization_setoid K V) notation `ℙ` := projectivization diff --git a/src/linear_algebra/quadratic_form/isometry.lean b/src/linear_algebra/quadratic_form/isometry.lean index f51eef6dfcb0e..36af1c374f979 100644 --- a/src/linear_algebra/quadratic_form/isometry.lean +++ b/src/linear_algebra/quadratic_form/isometry.lean @@ -30,7 +30,7 @@ variables [module R M] [module R M₁] [module R M₂] [module R M₃] /-- An isometry between two quadratic spaces `M₁, Q₁` and `M₂, Q₂` over a ring `R`, is a linear equivalence between `M₁` and `M₂` that commutes with the quadratic forms. -/ -@[nolint has_inhabited_instance] structure isometry +@[nolint has_nonempty_instance] structure isometry (Q₁ : quadratic_form R M₁) (Q₂ : quadratic_form R M₂) extends M₁ ≃ₗ[R] M₂ := (map_app' : ∀ m, Q₂ (to_fun m) = Q₁ m) diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index e6637cf101f87..5692a3e5fc97f 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -159,7 +159,7 @@ end same_ray /-- Nonzero vectors, as used to define rays. This type depends on an unused argument `R` so that `ray_vector.setoid` can be an instance. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def ray_vector (R M : Type*) [has_zero M] := {v : M // v ≠ 0} instance ray_vector.has_coe {R M : Type*} [has_zero M] : @@ -177,7 +177,7 @@ instance : setoid (ray_vector R M) := λ x y z hxy hyz, hxy.trans hyz $ λ hy, (y.2 hy).elim⟩ } /-- A ray (equivalence class of nonzero vectors with common positive multiples) in a module. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def module.ray := quotient (ray_vector.setoid R M) variables {R M} diff --git a/src/logic/embedding.lean b/src/logic/embedding.lean index 2e6de95a72255..de2acbb334707 100644 --- a/src/logic/embedding.lean +++ b/src/logic/embedding.lean @@ -18,7 +18,7 @@ universes u v w x namespace function /-- `α ↪ β` is a bundled injective function. -/ -@[nolint has_inhabited_instance] -- depending on cardinalities, an injective function may not exist +@[nolint has_nonempty_instance] -- depending on cardinalities, an injective function may not exist structure embedding (α : Sort*) (β : Sort*) := (to_fun : α → β) (inj' : injective to_fun) diff --git a/src/logic/small.lean b/src/logic/small.lean index e2dc1c96593d1..ae110d8e29457 100644 --- a/src/logic/small.lean +++ b/src/logic/small.lean @@ -34,7 +34,7 @@ lemma small.mk' {α : Type v} {S : Type w} (e : α ≃ S) : small.{w} α := /-- An arbitrarily chosen model in `Type w` for a `w`-small type. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def shrink (α : Type v) [small.{w} α] : Type w := classical.some (@small.equiv_small α _) diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index 192072a166498..bd0d9afdbf04a 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -60,7 +60,7 @@ Vitali families are provided by covering theorems such as the Besicovitch coveri Vitali covering theorem. They make it possible to formulate general versions of theorems on differentiations of measure that apply in both contexts. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure vitali_family {m : measurable_space α} (μ : measure α) := (sets_at : Π (x : α), set (set α)) (measurable_set' : ∀ (x : α), ∀ (a : set α), a ∈ sets_at x → measurable_set a) diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 431bbcde4d5b0..90ecc0e58bdea 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2483,7 +2483,7 @@ lemma finite_at_bot {m0 : measurable_space α} (μ : measure α) : μ.finite_at_ about the sets, such as that they are monotone. `sigma_finite` is defined in terms of this: `μ` is σ-finite if there exists a sequence of finite spanning sets in the collection of all measurable sets. -/ -@[protect_proj, nolint has_inhabited_instance] +@[protect_proj, nolint has_nonempty_instance] structure finite_spanning_sets_in {m0 : measurable_space α} (μ : measure α) (C : set (set α)) := (set : ℕ → set α) (set_mem : ∀ i, set i ∈ C) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 4b4bcf682c69b..2c3df278de426 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -112,14 +112,14 @@ protected def sum (L : language.{u v}) (L' : language.{u' v'}) : language := variable (L : language.{u v}) /-- The type of constants in a given language. -/ -@[nolint has_inhabited_instance] protected def «constants» := L.functions 0 +@[nolint has_nonempty_instance] protected def «constants» := L.functions 0 @[simp] lemma constants_mk₂ (c f₁ f₂ : Type u) (r₁ r₂ : Type v) : (language.mk₂ c f₁ f₂ r₁ r₂).constants = c := rfl /-- The type of symbols in a given language. -/ -@[nolint has_inhabited_instance] def symbols := (Σl, L.functions l) ⊕ (Σl, L.relations l) +@[nolint has_nonempty_instance] def symbols := (Σl, L.functions l) ⊕ (Σl, L.relations l) /-- The cardinality of a language is the cardinality of its type of symbols. -/ def card : cardinal := # L.symbols diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index e94d147983708..9bdc44b5069be 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -151,7 +151,7 @@ end /-- Structure 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. -/ -@[nolint has_inhabited_instance] -- can't prove that they always exist +@[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 : pnat) (char : add_char R (cyclotomic_field n R')) diff --git a/src/order/galois_connection.lean b/src/order/galois_connection.lean index c52a6a1d7ae73..e38dd6af59b1f 100644 --- a/src/order/galois_connection.lean +++ b/src/order/galois_connection.lean @@ -389,7 +389,7 @@ end nat /-- A Galois insertion is a Galois connection where `l ∘ u = id`. It also contains a constructive choice function, to give better definitional equalities when lifting order structures. Dual to `galois_coinsertion` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure galois_insertion {α β : Type*} [preorder α] [preorder β] (l : α → β) (u : β → α) := (choice : Πx : α, u (l x) ≤ x → β) (gc : galois_connection l u) @@ -582,7 +582,7 @@ end galois_insertion /-- A Galois coinsertion is a Galois connection where `u ∘ l = id`. It also contains a constructive choice function, to give better definitional equalities when lifting order structures. Dual to `galois_insertion` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure galois_coinsertion [preorder α] [preorder β] (l : α → β) (u : β → α) := (choice : Πx : β, x ≤ l (u x) → α) (gc : galois_connection l u) diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 405b90051df03..d427a7c260d19 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -169,7 +169,7 @@ segments. /-- If `r` is a relation on `α` and `s` in a relation on `β`, then `f : r ≺i s` is an order embedding whose range is an open interval `(-∞, top)` for some element `top` of `β`. Such order embeddings are called principal segments -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure principal_seg {α β : Type*} (r : α → α → Prop) (s : β → β → Prop) extends r ↪r s := (top : β) (down : ∀ b, s b top ↔ ∃ a, to_rel_embedding a = b) diff --git a/src/order/prime_ideal.lean b/src/order/prime_ideal.lean index 09e5457b2ea80..d113dabce9f8e 100644 --- a/src/order/prime_ideal.lean +++ b/src/order/prime_ideal.lean @@ -40,7 +40,7 @@ namespace ideal /-- A pair of an `ideal` and a `pfilter` which form a partition of `P`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure prime_pair (P : Type*) [preorder P] := (I : ideal P) (F : pfilter P) diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index 60c811dfbeef1..1a93b746ec328 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -40,7 +40,7 @@ variables {α β γ : Type*} {r : α → α → Prop} {s : β → β → Prop} { /-- A relation homomorphism with respect to a given pair of relations `r` and `s` is a function `f : α → β` such that `r a b → s (f a) (f b)`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure rel_hom {α β : Type*} (r : α → α → Prop) (s : β → β → Prop) := (to_fun : α → β) (map_rel' : ∀ {a b}, r a b → s (to_fun a) (to_fun b)) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index f75b732dd923a..677414e077afd 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -831,7 +831,7 @@ variables [is_domain R] [is_dedekind_domain R] /-- The height one prime spectrum of a Dedekind domain `R` is the type of nonzero prime ideals of `R`. Note that this equals the maximal spectrum if `R` has Krull dimension 1. -/ -@[ext, nolint has_inhabited_instance unused_arguments] +@[ext, nolint has_nonempty_instance unused_arguments] structure height_one_spectrum := (as_ideal : ideal R) (is_prime : as_ideal.is_prime) diff --git a/src/ring_theory/graded_algebra/homogeneous_localization.lean b/src/ring_theory/graded_algebra/homogeneous_localization.lean index 4b8d5d8e9f1e6..ee1f4e61e9b0e 100644 --- a/src/ring_theory/graded_algebra/homogeneous_localization.lean +++ b/src/ring_theory/graded_algebra/homogeneous_localization.lean @@ -79,7 +79,7 @@ section Let `x` be a prime ideal, then `num_denom_same_deg 𝒜 x` is a structure with a numerator and a denominator with same grading such that the denominator is not contained in `x`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure num_denom_same_deg := (deg : ι) (num denom : 𝒜 deg) @@ -214,7 +214,7 @@ For `x : prime ideal of A`, `homogeneous_localization 𝒜 x` is `num_denom_same kernel of `embedding 𝒜 x`. This is essentially the subring of `Aₓ` where the numerator and denominator share the same grading. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def homogeneous_localization : Type* := quotient (setoid.ker $ homogeneous_localization.num_denom_same_deg.embedding 𝒜 x) diff --git a/src/ring_theory/perfection.lean b/src/ring_theory/perfection.lean index c4f8b0bb7bb69..f2ee9c480787b 100644 --- a/src/ring_theory/perfection.lean +++ b/src/ring_theory/perfection.lean @@ -201,7 +201,7 @@ end perfection /-- A perfection map to a ring of characteristic `p` is a map that is isomorphic to its perfection. -/ -@[nolint has_inhabited_instance] structure perfection_map (p : ℕ) [fact p.prime] +@[nolint has_nonempty_instance] structure perfection_map (p : ℕ) [fact p.prime] {R : Type u₁} [comm_semiring R] [char_p R p] {P : Type u₂} [comm_semiring P] [char_p P p] [perfect_ring P p] (π : P →+* R) : Prop := (injective : ∀ ⦃x y : P⦄, (∀ n, π (pth_root P p ^[n] x) = π (pth_root P p ^[n] y)) → x = y) @@ -321,7 +321,7 @@ variables (p : ℕ) include hv /-- `O/(p)` for `O`, ring of integers of `K`. -/ -@[nolint unused_arguments has_inhabited_instance] def mod_p := +@[nolint unused_arguments has_nonempty_instance] def mod_p := O ⧸ (ideal.span {p} : ideal O) variables [hp : fact p.prime] [hvp : fact (v p ≠ 1)] @@ -431,7 +431,7 @@ end classical end mod_p /-- Perfection of `O/(p)` where `O` is the ring of integers of `K`. -/ -@[nolint has_inhabited_instance] def pre_tilt := +@[nolint has_nonempty_instance] def pre_tilt := ring.perfection (mod_p K v O hv p) p include hp hvp @@ -559,7 +559,7 @@ end pre_tilt /-- The tilt of a field, as defined in Perfectoid Spaces by Peter Scholze, as in [scholze2011perfectoid]. Given a field `K` with valuation `K → ℝ≥0` and ring of integers `O`, this is implemented as the fraction field of the perfection of `O/(p)`. -/ -@[nolint has_inhabited_instance] def tilt := +@[nolint has_nonempty_instance] def tilt := fraction_ring (pre_tilt K v O hv p) namespace tilt diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index 65d72d8a8f874..349cf5064473d 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -55,7 +55,7 @@ This is a structure, not a class, since the same algebra can have many power bas For the common case where `S` is defined by adjoining an integral element to `R`, the canonical power basis is given by `{algebra,intermediate_field}.adjoin.power_basis`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure power_basis (R S : Type*) [comm_ring R] [ring S] [algebra R S] := (gen : S) (dim : ℕ) diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 0174b62776a02..172abda537bf4 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -75,7 +75,7 @@ variables (F R) (Γ₀ : Type*) [linear_ordered_comm_monoid_with_zero Γ₀] [ri /-- The type of `Γ₀`-valued valuations on `R`. When you extend this structure, make sure to extend `valuation_class`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure valuation extends R →*₀ Γ₀ := (map_add_le_max' : ∀ x y, to_fun (x + y) ≤ max (to_fun x) (to_fun y)) @@ -580,7 +580,7 @@ section add_monoid variables (R) [ring R] (Γ₀ : Type*) [linear_ordered_add_comm_monoid_with_top Γ₀] /-- The type of `Γ₀`-valued additive valuations on `R`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def add_valuation := valuation R (multiplicative Γ₀ᵒᵈ) end add_monoid diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index 6ad6d3fb1904e..7323a6d8f3312 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -110,12 +110,12 @@ variables (V) localized "notation `Φ(` p`,` k`)` := witt_vector.isocrystal.frobenius p k" in isocrystal /-- A homomorphism between isocrystals respects the Frobenius map. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure isocrystal_hom extends V →ₗ[K(p, k)] V₂ := ( frob_equivariant : ∀ x : V, Φ(p, k) (to_linear_map x) = to_linear_map (Φ(p, k) x) ) /-- An isomorphism between isocrystals respects the Frobenius map. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure isocrystal_equiv extends V ≃ₗ[K(p, k)] V₂ := ( frob_equivariant : ∀ x : V, Φ(p, k) (to_linear_equiv x) = to_linear_equiv (Φ(p, k) x) ) @@ -139,7 +139,7 @@ def fraction_ring.module : module K(p, k) K(p, k) := semiring.to_module Type synonym for `K(p, k)` to carry the standard 1-dimensional isocrystal structure of slope `m : ℤ`. -/ -@[nolint unused_arguments has_inhabited_instance, derive [add_comm_group, module K(p, k)]] +@[nolint unused_arguments has_nonempty_instance, derive [add_comm_group, module K(p, k)]] def standard_one_dim_isocrystal (m : ℤ) : Type* := K(p, k) diff --git a/src/tactic/derive_fintype.lean b/src/tactic/derive_fintype.lean index bffeae606ed1d..6aa82c1ba17b6 100644 --- a/src/tactic/derive_fintype.lean +++ b/src/tactic/derive_fintype.lean @@ -144,7 +144,7 @@ instance (α enum n) : inhabited (finset_above α enum n) := ⟨finset_above.nil /-- This is a finset covering a nontrivial variant (with one or more constructor arguments). The property `P` here is `λ a, enum a = n` where `n` is the discriminant for the current variant. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def finset_in {α} (P : α → Prop) := {s : finset α // ∀ x ∈ s, P x} /-- To construct the finset, we use an injective map from the type `Γ`, which will be the diff --git a/src/tactic/lint/default.lean b/src/tactic/lint/default.lean index 2db842856f084..2158d4eade770 100644 --- a/src/tactic/lint/default.lean +++ b/src/tactic/lint/default.lean @@ -40,7 +40,8 @@ The following linters are run by default: 4. `ge_or_gt` checks whether ≥/> is used in the declaration. 5. `instance_priority` checks that instances that always apply have priority below default. 6. `doc_blame` checks for missing doc strings on definitions and constants. -7. `has_inhabited_instance` checks whether every type has an associated `inhabited` instance. +7. `has_nonempty_instance` checks whether every type has an associated `inhabited`, `unique` + or `nonempty` instance. 8. `impossible_instance` checks for instances that can never fire. 9. `incorrect_type_class_argument` checks for arguments in [square brackets] that are not classes. 10. `dangerous_instance` checks for instances that generate type-class problems with metavariables. diff --git a/src/tactic/lint/type_classes.lean b/src/tactic/lint/type_classes.lean index e9b3ca7ecae7c..db71eccca7d85 100644 --- a/src/tactic/lint/type_classes.lean +++ b/src/tactic/lint/type_classes.lean @@ -14,7 +14,8 @@ This file defines several linters checking the correct usage of type classes and the appropriate definition of instances: * `instance_priority` ensures that blanket instances have low priority. - * `has_inhabited_instances` checks that every type has an `inhabited` instance. + * `has_nonempty_instances` checks that every type has a `nonempty` instance, an `inhabited` + instance, or a `unique` instance. * `impossible_instance` checks that there are no instances which can never apply. * `incorrect_type_class_argument` checks that only type classes are used in instance-implicit arguments. @@ -105,8 +106,10 @@ If you don't know what priority to choose, use priority 100. See note [lower instance priority] for instructions to change the priority.", auto_decls := tt } -/-- Reports declarations of types that do not have an associated `inhabited` instance. -/ -private meta def has_inhabited_instance (d : declaration) : tactic (option string) := do +/-- Reports declarations of types that do not have an nonemptiness instance. +A `nonempty`, `inhabited` or `unique` instance suffices, and we prefer a computable `inhabited` +or `unique` instance if possible. -/ +private meta def has_nonempty_instance (d : declaration) : tactic (option string) := do tt ← pure d.is_trusted | pure none, ff ← has_attribute' `reducible d.to_name | pure none, ff ← has_attribute' `class d.to_name | pure none, @@ -116,24 +119,26 @@ if ty = `(Prop) then pure none else do `(Sort _) ← whnf ty | pure none, insts ← attribute.get_instances `instance, insts_tys ← insts.mmap $ λ i, expr.pi_codomain <$> declaration.type <$> get_decl i, -let inhabited_insts := insts_tys.filter (λ i, - i.app_fn.const_name = ``inhabited ∨ i.app_fn.const_name = `unique), -let inhabited_tys := inhabited_insts.map (λ i, i.app_arg.get_app_fn.const_name), -if d.to_name ∈ inhabited_tys then +let nonempty_insts := insts_tys.filter + (λ i, i.app_fn.const_name ∈ [``nonempty, ``inhabited, `unique]), +let nonempty_tys := nonempty_insts.map (λ i, i.app_arg.get_app_fn.const_name), +if d.to_name ∈ nonempty_tys then pure none else - pure "inhabited instance missing" + pure "nonempty/inhabited/unique instance missing" -/-- A linter for missing `inhabited` instances. -/ +/-- A linter for missing `nonempty` instances. -/ @[linter] -meta def linter.has_inhabited_instance : linter := -{ test := has_inhabited_instance, +meta def linter.has_nonempty_instance : linter := +{ test := has_nonempty_instance, auto_decls := ff, - no_errors_found := "No types have missing inhabited instances.", - errors_found := "TYPES ARE MISSING INHABITED INSTANCES:", + no_errors_found := "No types have missing nonempty instances.", + errors_found := "TYPES ARE MISSING NONEMPTY INSTANCES. +The following types should have an associated instance of the class +`nonempty`, or if computably possible `inhabited` or `unique`:", is_fast := ff } -attribute [nolint has_inhabited_instance] pempty +attribute [nolint has_nonempty_instance] pempty /-- Checks whether an instance can never be applied. -/ private meta def impossible_instance (d : declaration) : tactic (option string) := do diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index acf06b2b4f516..031c5df8e24b4 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -283,7 +283,7 @@ continuous_semilinear_map_class F (ring_hom.id R) M M₂ /-- Continuous linear equivalences 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 semiring `R`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure continuous_linear_equiv {R : Type*} {S : Type*} [semiring R] [semiring S] (σ : R →+* S) {σ' : S →+* R} [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ] diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 81c938b9b92d8..e785f953d5fbb 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -68,7 +68,7 @@ section weak_topology /-- The space `E` equipped with the weak topology induced by the bilinear form `B`. -/ @[derive [add_comm_monoid, module 𝕜], -nolint has_inhabited_instance unused_arguments] +nolint has_nonempty_instance unused_arguments] def weak_bilin [comm_semiring 𝕜] [add_comm_monoid E] [module 𝕜 E] [add_comm_monoid F] [module 𝕜 F] (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) := E @@ -242,7 +242,7 @@ end weak_dual /-- The weak topology is the topology coarsest topology on `E` such that all functionals `λ x, top_dual_pairing 𝕜 E v x` are continuous. -/ @[derive [add_comm_monoid, module 𝕜, topological_space, has_continuous_add], -nolint has_inhabited_instance] +nolint has_nonempty_instance] def weak_space (𝕜 E) [comm_semiring 𝕜] [topological_space 𝕜] [has_continuous_add 𝕜] [has_continuous_const_smul 𝕜 𝕜] [add_comm_monoid E] [module 𝕜 E] [topological_space E] := weak_bilin (top_dual_pairing 𝕜 E).flip diff --git a/src/topology/fiber_bundle.lean b/src/topology/fiber_bundle.lean index baae1192a3ec6..b9d5ac358a962 100644 --- a/src/topology/fiber_bundle.lean +++ b/src/topology/fiber_bundle.lean @@ -164,7 +164,7 @@ below as `trivialization F proj`) if the total space has not been given a topolo have a topology on both the fiber and the base space. Through the construction `topological_fiber_prebundle F proj` it will be possible to promote a `pretrivialization F proj` to a `trivialization F proj`. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure topological_fiber_bundle.pretrivialization (proj : Z → B) extends local_equiv Z (B × F) := (open_target : is_open target) (base_set : set B) @@ -278,7 +278,7 @@ A structure extending local homeomorphisms, defining a local trivialization of a `proj : Z → B` with fiber `F`, as a local homeomorphism between `Z` and `B × F` defined between two sets of the form `proj ⁻¹' base_set` and `base_set × F`, acting trivially on the first coordinate. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure topological_fiber_bundle.trivialization (proj : Z → B) extends local_homeomorph Z (B × F) := (base_set : set B) @@ -833,7 +833,7 @@ Trivialization changes from `i` to `j` are given by continuous maps `coord_chang `base_set i ∩ base_set j` to the set of homeomorphisms of `F`, but we express them as maps `B → F → F` and require continuity on `(base_set i ∩ base_set j) × F` to avoid the topology on the space of continuous maps on `F`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure topological_fiber_bundle_core (ι : Type*) (B : Type*) [topological_space B] (F : Type*) [topological_space F] := (base_set : ι → set B) @@ -854,7 +854,7 @@ variables [topological_space B] [topological_space F] (Z : topological_fiber_bun include Z /-- The index set of a topological fiber bundle core, as a convenience function for dot notation -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def index := ι /-- The base space of a topological fiber bundle core, as a convenience function for dot notation -/ @@ -863,7 +863,7 @@ def base := B /-- The fiber of a topological fiber bundle core, as a convenience function for dot notation and typeclass inference -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def fiber (x : B) := F section fiber_instances @@ -1166,7 +1166,7 @@ open topological_fiber_bundle equivalences but there is not yet a topology on the total space. The total space is hence given a topology in such a way that there is a fiber bundle structure for which the local equivalences are also local homeomorphism and hence local trivializations. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure topological_fiber_prebundle (proj : Z → B) := (pretrivialization_atlas : set (pretrivialization F proj)) (pretrivialization_at : B → pretrivialization F proj) diff --git a/src/topology/gluing.lean b/src/topology/gluing.lean index 25543d9719b0b..934a506d6500b 100644 --- a/src/topology/gluing.lean +++ b/src/topology/gluing.lean @@ -77,7 +77,7 @@ that the `U i`'s are open subspaces of the glued space. Most of the times it would be easier to use the constructor `Top.glue_data.mk'` where the conditions are stated in a less categorical way. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure glue_data extends glue_data Top := (f_open : ∀ i j, open_embedding (f i j)) (f_mono := λ i j, (Top.mono_iff_injective _).mpr (f_open i j).to_embedding.inj) @@ -279,7 +279,7 @@ such that We can then glue the topological spaces `U i` together by identifying `V i j` with `V j i`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure mk_core := {J : Type u} (U : J → Top.{u}) diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 8dbee1460d6d0..33894f6bfa1a8 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -32,7 +32,7 @@ open_locale topological_space variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} /-- Homeomorphism between `α` and `β`, also called topological isomorphism -/ -@[nolint has_inhabited_instance] -- not all spaces are homeomorphic to each other +@[nolint has_nonempty_instance] -- not all spaces are homeomorphic to each other structure homeomorph (α : Type*) (β : Type*) [topological_space α] [topological_space β] extends α ≃ β := (continuous_to_fun : continuous to_fun . tactic.interactive.continuity') diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index 4c1bd0645075e..3593294b85b9c 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -49,7 +49,7 @@ variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} [topological_space α] [topological_space β] [topological_space γ] [topological_space δ] /-- local homeomorphisms, defined on open subsets of the space -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure local_homeomorph (α : Type*) (β : Type*) [topological_space α] [topological_space β] extends local_equiv α β := (open_source : is_open source) diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index fd9c2b517a402..61f22c5b4d64e 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -83,7 +83,7 @@ definition to_GH_space (X : Type u) [metric_space X] [compact_space X] [nonempty instance : inhabited GH_space := ⟨quot.mk _ ⟨⟨{0}, is_compact_singleton⟩, singleton_nonempty _⟩⟩ /-- A metric space representative of any abstract point in `GH_space` -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def GH_space.rep (p : GH_space) : Type := (quotient.out p : nonempty_compacts ℓ_infty_ℝ) lemma eq_to_GH_space_iff {X : Type u} [metric_space X] [compact_space X] [nonempty X] diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index f0623b8f2058b..c5f09a4aaed7d 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -451,7 +451,7 @@ def premetric_optimal_GH_dist : pseudo_metric_space (X ⊕ Y) := local attribute [instance] premetric_optimal_GH_dist pseudo_metric.dist_setoid /-- A metric space which realizes the optimal coupling between `X` and `Y` -/ -@[derive metric_space, nolint has_inhabited_instance] +@[derive metric_space, nolint has_nonempty_instance] definition optimal_GH_coupling : Type* := pseudo_metric_quot (X ⊕ Y) diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index d92a4530380cd..37e013925d5fe 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -232,7 +232,7 @@ begin end /-- `α` and `β` are isometric if there is an isometric bijection between them. -/ -@[nolint has_inhabited_instance] -- such a bijection need not exist +@[nolint has_nonempty_instance] -- such a bijection need not exist structure isometric (α : Type*) (β : Type*) [pseudo_emetric_space α] [pseudo_emetric_space β] extends α ≃ β := (isometry_to_fun : isometry to_fun) diff --git a/src/topology/metric_space/polish.lean b/src/topology/metric_space/polish.lean index b65a2b0d9ef48..d30db5f669224 100644 --- a/src/topology/metric_space/polish.lean +++ b/src/topology/metric_space/polish.lean @@ -180,7 +180,7 @@ lemma _root_.is_closed.polish_space {α : Type*} [topological_space α] [polish_ /-- A sequence of type synonyms of a given type `α`, useful in the proof of `exists_polish_space_forall_le` to endow each copy with a different topology. -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def aux_copy (α : Type*) {ι : Type*} (i : ι) : Type* := α /-- Given a Polish space, and countably many finer Polish topologies, there exists another Polish @@ -259,7 +259,7 @@ variables [metric_space α] {s : set α} /-- A type synonym for a subset `s` of a metric space, on which we will construct another metric for which it will be complete. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] def complete_copy {α : Type*} (s : set α) : Type* := s /-- A distance on a subset `s` of a metric space, designed to make it complete if `s` is open. diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index 52d02aceb005d..4397b7b123f04 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -67,7 +67,7 @@ variables {X Y : Type*} [topological_space X] [topological_space Y] {x y z : X} /-! ### Paths -/ /-- Continuous path connecting two points `x` and `y` in a topological space -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure path (x y : X) extends C(I, X) := (source' : to_fun 0 = x) (target' : to_fun 1 = y) diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index b27deb2fb58cb..06147fc2c2434 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -37,7 +37,7 @@ variables (C : Type u) [category.{v} C] namespace Top /-- The category of `C`-valued presheaves on a (bundled) topological space `X`. -/ -@[derive category, nolint has_inhabited_instance] +@[derive category, nolint has_nonempty_instance] def presheaf (X : Top.{w}) := (opens X)ᵒᵖ ⥤ C variables {C} diff --git a/src/topology/shrinking_lemma.lean b/src/topology/shrinking_lemma.lean index 7c346e2dd8443..cc359edf4b3ce 100644 --- a/src/topology/shrinking_lemma.lean +++ b/src/topology/shrinking_lemma.lean @@ -45,7 +45,7 @@ namespace shrinking_lemma This type is equipped with the folowing partial order: `v ≤ v'` if `v.carrier ⊆ v'.carrier` and `v i = v' i` for `i ∈ v.carrier`. We will use Zorn's lemma to prove that this type has a maximal element, then show that the maximal element must have `carrier = univ`. -/ -@[nolint has_inhabited_instance] -- the trivial refinement needs `u` to be a covering +@[nolint has_nonempty_instance] -- the trivial refinement needs `u` to be a covering structure partial_refinement (u : ι → set X) (s : set X) := (to_fun : ι → set X) (carrier : set ι) diff --git a/src/topology/uniform_space/equiv.lean b/src/topology/uniform_space/equiv.lean index 9ebf5a9481b64..037462f508c2a 100644 --- a/src/topology/uniform_space/equiv.lean +++ b/src/topology/uniform_space/equiv.lean @@ -27,7 +27,7 @@ open_locale variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} /-- Uniform isomorphism between `α` and `β` -/ -@[nolint has_inhabited_instance] -- not all spaces are homeomorphic to each other +@[nolint has_nonempty_instance] -- not all spaces are homeomorphic to each other structure uniform_equiv (α : Type*) (β : Type*) [uniform_space α] [uniform_space β] extends α ≃ β := (uniform_continuous_to_fun : uniform_continuous to_fun) diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 1f421ccfd5d5e..371393ce6255c 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -56,7 +56,7 @@ variables [semiring R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] /-- A pretrivialization for a (yet to be defined) topological vector bundle `total_space E` is a local equiv between sets of the form `proj ⁻¹' base_set` and `base_set × F` which respects the first coordinate, and is linear in each fiber. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure topological_vector_bundle.pretrivialization extends to_fiber_bundle_pretrivialization : topological_fiber_bundle.pretrivialization F (@total_space.proj B E) := (linear' : ∀ x ∈ base_set, is_linear_map R (λ y : E x, (to_fun (total_space_mk x y)).2)) @@ -234,7 +234,7 @@ A structure extending local homeomorphisms, defining a local trivialization of t and `B × F` defined between two sets of the form `proj ⁻¹' base_set` and `base_set × F`, acting trivially on the first coordinate and linear in the fibers. -/ -@[ext, nolint has_inhabited_instance] +@[ext, nolint has_nonempty_instance] structure topological_vector_bundle.trivialization extends to_fiber_bundle_trivialization : topological_fiber_bundle.trivialization F (@total_space.proj B E) := (linear' : ∀ x ∈ base_set, is_linear_map R (λ y : E x, (to_fun (total_space_mk x y)).2)) @@ -765,7 +765,7 @@ lemma coord_change_linear_comp (i j k : ι): ∀ x ∈ (Z.base_set i) ∩ (Z.bas λ x hx, by { ext v, exact Z.coord_change_comp i j k x hx v } /-- The index set of a topological vector bundle core, as a convenience function for dot notation -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def index := ι /-- The base space of a topological vector bundle core, as a convenience function for dot notation-/ @@ -774,7 +774,7 @@ def base := B /-- The fiber of a topological vector bundle core, as a convenience function for dot notation and typeclass inference -/ -@[nolint unused_arguments has_inhabited_instance] +@[nolint unused_arguments has_nonempty_instance] def fiber (x : B) := F instance topological_space_fiber (x : B) : topological_space (Z.fiber x) := @@ -934,7 +934,7 @@ The field `exists_coord_change` is stated as an existential statement (instead o fields), since it depends on propositional information (namely `e e' ∈ pretrivialization_atlas`). This makes it inconvenient to explicitly define a `coord_change` function when constructing a `topological_vector_prebundle`. -/ -@[nolint has_inhabited_instance] +@[nolint has_nonempty_instance] structure topological_vector_prebundle := (pretrivialization_atlas : set (pretrivialization R F E)) (pretrivialization_at : B → pretrivialization R F E) diff --git a/test/has_scalar_comp_loop.lean b/test/has_scalar_comp_loop.lean index 34d2504781e21..a853931895990 100644 --- a/test/has_scalar_comp_loop.lean +++ b/test/has_scalar_comp_loop.lean @@ -3,7 +3,7 @@ import group_theory.group_action.basic variables (R M S : Type*) /-- Some arbitrary type depending on `has_smul R M` -/ -@[irreducible, nolint has_inhabited_instance unused_arguments] +@[irreducible, nolint has_nonempty_instance unused_arguments] def foo [has_smul R M] : Type* := ℕ variables [has_smul R M] [has_smul S R] [has_smul S M] From 280d7059d40cf8af739b180fe1dae527bfc25dae Mon Sep 17 00:00:00 2001 From: tb65536 Date: Sat, 30 Jul 2022 19:08:11 +0000 Subject: [PATCH 148/173] feat(algebra/algebra/subalgebra/basic): Swap arguments of `map` and `comap` (#15744) As suggested by @eric-wieser, I swapped the arguments of `subalgebra.map` and `subalgebra.comap`. I also updated `subalgebra.map_injective` to use `function.injective` and to have `f` implicit. This is done to make the `subalgebra` API match the `subgroup`, `submodule`, etc... API. --- src/algebra/algebra/subalgebra/basic.lean | 22 +++++++++++----------- src/field_theory/intermediate_field.lean | 4 ++-- src/ring_theory/adjoin/fg.lean | 2 +- src/ring_theory/algebra_tower.lean | 3 +-- 4 files changed, 15 insertions(+), 16 deletions(-) diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index e37628792f246..7b5b072c843f6 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -305,7 +305,7 @@ def to_submodule_equiv (S : subalgebra R A) : S.to_submodule ≃ₗ[R] S := linear_equiv.of_eq _ _ rfl /-- Transport a subalgebra via an algebra homomorphism. -/ -def map (S : subalgebra R A) (f : A →ₐ[R] B) : subalgebra R B := +def map (f : A →ₐ[R] B) (S : subalgebra R A) : subalgebra R B := { algebra_map_mem' := λ r, f.commutes r ▸ set.mem_image_of_mem _ (S.algebra_map_mem r), .. S.to_subsemiring.map (f : A →+* B) } @@ -313,9 +313,9 @@ lemma map_mono {S₁ S₂ : subalgebra R A} {f : A →ₐ[R] B} : S₁ ≤ S₂ → S₁.map f ≤ S₂.map f := set.image_subset f -lemma map_injective {S₁ S₂ : subalgebra R A} (f : A →ₐ[R] B) - (hf : function.injective f) (ih : S₁.map f = S₂.map f) : S₁ = S₂ := -ext $ set.ext_iff.1 $ set.image_injective.2 hf $ set.ext $ set_like.ext_iff.mp ih +lemma map_injective {f : A →ₐ[R] B} (hf : function.injective f) : + function.injective (map f) := +λ S₁ S₂ ih, ext $ set.ext_iff.1 $ set.image_injective.2 hf $ set.ext $ set_like.ext_iff.mp ih @[simp] lemma map_id (S : subalgebra R A) : S.map (alg_hom.id R A) = S := set_like.coe_injective $ set.image_id _ @@ -325,7 +325,7 @@ lemma map_map (S : subalgebra R A) (g : B →ₐ[R] C) (f : A →ₐ[R] B) : set_like.coe_injective $ set.image_image _ _ _ lemma mem_map {S : subalgebra R A} {f : A →ₐ[R] B} {y : B} : - y ∈ map S f ↔ ∃ x ∈ S, f x = y := + y ∈ map f S ↔ ∃ x ∈ S, f x = y := subsemiring.mem_map lemma map_to_submodule {S : subalgebra R A} {f : A →ₐ[R] B} : @@ -341,16 +341,16 @@ set_like.coe_injective rfl rfl /-- Preimage of a subalgebra under an algebra homomorphism. -/ -def comap (S : subalgebra R B) (f : A →ₐ[R] B) : subalgebra R A := +def comap (f : A →ₐ[R] B) (S : subalgebra R B) : subalgebra R A := { algebra_map_mem' := λ r, show f (algebra_map R A r) ∈ S, from (f.commutes r).symm ▸ S.algebra_map_mem r, .. S.to_subsemiring.comap (f : A →+* B) } theorem map_le {S : subalgebra R A} {f : A →ₐ[R] B} {U : subalgebra R B} : - map S f ≤ U ↔ S ≤ comap U f := + map f S ≤ U ↔ S ≤ comap f U := set.image_subset_iff -lemma gc_map_comap (f : A →ₐ[R] B) : galois_connection (λ S, map S f) (λ S, comap S f) := +lemma gc_map_comap (f : A →ₐ[R] B) : galois_connection (map f) (comap f) := λ S U, map_le @[simp] lemma mem_comap (S : subalgebra R B) (f : A →ₐ[R] B) (x : A) : @@ -648,14 +648,14 @@ algebra.eq_top_iff @[simp] theorem range_id : (alg_hom.id R A).range = ⊤ := set_like.coe_injective set.range_id -@[simp] theorem map_top (f : A →ₐ[R] B) : subalgebra.map (⊤ : subalgebra R A) f = f.range := +@[simp] theorem map_top (f : A →ₐ[R] B) : (⊤ : subalgebra R A).map f = f.range := set_like.coe_injective set.image_univ -@[simp] theorem map_bot (f : A →ₐ[R] B) : subalgebra.map (⊥ : subalgebra R A) f = ⊥ := +@[simp] theorem map_bot (f : A →ₐ[R] B) : (⊥ : subalgebra R A).map f = ⊥ := set_like.coe_injective $ by simp only [← set.range_comp, (∘), algebra.coe_bot, subalgebra.coe_map, f.commutes] -@[simp] theorem comap_top (f : A →ₐ[R] B) : subalgebra.comap (⊤ : subalgebra R B) f = ⊤ := +@[simp] theorem comap_top (f : A →ₐ[R] B) : (⊤ : subalgebra R B).comap f = ⊤ := eq_top_iff.2 $ λ x, mem_top /-- `alg_hom` to `⊤ : subalgebra R A`. -/ diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index f2933d2f4e68f..36d0131dc3667 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -287,7 +287,7 @@ variables {L' : Type*} [field L'] [algebra K L'] /-- If `f : L →+* L'` fixes `K`, `S.map f` is the intermediate field between `L'` and `K` such that `x ∈ S ↔ f x ∈ S.map f`. -/ -def map (f : L →ₐ[K] L') : intermediate_field K L' := +def map (f : L →ₐ[K] L') (S : intermediate_field K L) : intermediate_field K L' := { inv_mem' := by { rintros _ ⟨x, hx, rfl⟩, exact ⟨x⁻¹, S.inv_mem hx, f.map_inv x⟩ }, neg_mem' := λ x hx, (S.to_subalgebra.map f).neg_mem hx, .. S.to_subalgebra.map f} @@ -397,7 +397,7 @@ section tower /-- Lift an intermediate_field of an intermediate_field -/ def lift {F : intermediate_field K L} (E : intermediate_field K F) : intermediate_field K L := -map E (val F) +E.map (val F) instance has_lift {F : intermediate_field K L} : has_lift_t (intermediate_field K F) (intermediate_field K L) := ⟨lift⟩ diff --git a/src/ring_theory/adjoin/fg.lean b/src/ring_theory/adjoin/fg.lean index 29c97f8ff97dd..a997fc1c51b34 100644 --- a/src/ring_theory/adjoin/fg.lean +++ b/src/ring_theory/adjoin/fg.lean @@ -126,7 +126,7 @@ end lemma fg_of_fg_map (S : subalgebra R A) (f : A →ₐ[R] B) (hf : function.injective f) (hs : (S.map f).fg) : S.fg := -let ⟨s, hs⟩ := hs in ⟨s.preimage f $ λ _ _ _ _ h, hf h, map_injective f hf $ +let ⟨s, hs⟩ := hs in ⟨s.preimage f $ λ _ _ _ _ h, hf h, map_injective hf $ by { rw [← algebra.adjoin_image, finset.coe_preimage, set.image_preimage_eq_of_subset, hs], rw [← alg_hom.coe_range, ← algebra.adjoin_le_iff, hs, ← algebra.map_top], exact map_mono le_top }⟩ diff --git a/src/ring_theory/algebra_tower.lean b/src/ring_theory/algebra_tower.lean index 8d019859161ba..9417f333be828 100644 --- a/src/ring_theory/algebra_tower.lean +++ b/src/ring_theory/algebra_tower.lean @@ -68,8 +68,7 @@ namespace algebra theorem adjoin_algebra_map (R : Type u) (S : Type v) (A : Type w) [comm_semiring R] [comm_semiring S] [semiring A] [algebra R S] [algebra S A] [algebra R A] [is_scalar_tower R S A] (s : set S) : - adjoin R (algebra_map S A '' s) = - subalgebra.map (adjoin R s) (is_scalar_tower.to_alg_hom R S A) := + adjoin R (algebra_map S A '' s) = (adjoin R s).map (is_scalar_tower.to_alg_hom R S A) := le_antisymm (adjoin_le $ set.image_subset_iff.2 $ λ y hy, ⟨y, subset_adjoin hy, rfl⟩) (subalgebra.map_le.2 $ adjoin_le $ λ y hy, subset_adjoin ⟨y, hy, rfl⟩) From b48983c9f6a442e3b5041c9144886e4123dc32b6 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sat, 30 Jul 2022 20:46:21 +0000 Subject: [PATCH 149/173] chore(logic/function/basic): make `function.injective.decidable_eq` protected (#15759) --- src/logic/function/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index 7af69f462df50..c80b3d209d686 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -86,7 +86,7 @@ h ▸ hf.ne_iff /-- If the co-domain `β` of an injective function `f : α → β` has decidable equality, then the domain `α` also has decidable equality. -/ -def injective.decidable_eq [decidable_eq β] (I : injective f) : decidable_eq α := +protected def injective.decidable_eq [decidable_eq β] (I : injective f) : decidable_eq α := λ a b, decidable_of_iff _ I.eq_iff lemma injective.of_comp {g : γ → α} (I : injective (f ∘ g)) : injective g := From 8d9599d6da1fa424d7d18093ce45e53ef5659480 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sun, 31 Jul 2022 03:48:52 +0000 Subject: [PATCH 150/173] chore(scripts): update nolints.txt (#15776) I am happy to remove some nolints for you! --- scripts/nolints.txt | 9 +-------- 1 file changed, 1 insertion(+), 8 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 3c203a8305e45..a72dba0ae7c0e 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -697,13 +697,6 @@ apply_nolint pow_gcd_card_eq_one_iff to_additive_doc apply_nolint powers_eq_zpowers fintype_finite apply_nolint subgroup.pow_index_mem fintype_finite --- group_theory/p_group.lean -apply_nolint is_p_group.bot_lt_center fintype_finite -apply_nolint is_p_group.center_nontrivial fintype_finite -apply_nolint is_p_group.exists_fixed_point_of_prime_dvd_card_of_fixed_point fintype_finite -apply_nolint is_p_group.index fintype_finite -apply_nolint is_p_group.nonempty_fixed_point_of_prime_not_dvd_card fintype_finite - -- group_theory/perm/cycle/basic.lean apply_nolint equiv.perm.closure_is_cycle fintype_finite apply_nolint equiv.perm.cycle_induction_on fintype_finite @@ -1629,4 +1622,4 @@ apply_nolint Cauchy.gen doc_blame apply_nolint uniform_space.completion.completion_separation_quotient_equiv doc_blame apply_nolint uniform_space.completion.cpkg doc_blame apply_nolint uniform_space.completion.extension₂ doc_blame -apply_nolint uniform_space.completion.map₂ doc_blame +apply_nolint uniform_space.completion.map₂ doc_blame \ No newline at end of file From 1474c566eab4d562b193869c2136b0a8b0e489ce Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 31 Jul 2022 05:16:10 +0000 Subject: [PATCH 151/173] feat(ring_theory/ideal/operations): A variant of `finsupp.total` for ideals. (#15460) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/ideal/operations.lean | 64 +++++++++++++++++++++++++++ 1 file changed, 64 insertions(+) diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 9e8d92e6e8df1..d41c4920d1210 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1397,6 +1397,70 @@ end⟩ end is_primary +section total + +variables (ι : Type*) +variables (M : Type*) [add_comm_group M] {R : Type*} [comm_ring R] [module R M] (I : ideal R) +variables (v : ι → M) (hv : submodule.span R (set.range v) = ⊤) + + +open_locale big_operators + +/-- A variant of `finsupp.total` that takes in vectors valued in `I`. -/ +noncomputable +def finsupp_total : (ι →₀ I) →ₗ[R] M := +(finsupp.total ι M R v).comp (finsupp.map_range.linear_map I.subtype) + +variables {ι M v} + +lemma finsupp_total_apply (f : ι →₀ I) : + finsupp_total ι M I v f = f.sum (λ i x, (x : R) • v i) := +begin + dsimp [finsupp_total], + rw [finsupp.total_apply, finsupp.sum_map_range_index], + exact λ _, zero_smul _ _ +end + +lemma finsupp_total_apply_eq_of_fintype [fintype ι] (f : ι →₀ I) : + finsupp_total ι M I v f = ∑ i, (f i : R) • v i := +begin + rw finsupp_total_apply, + apply finset.sum_subset (finset.subset_univ _), + intros x _ hx, + rw finsupp.not_mem_support_iff.mp hx, + exact zero_smul _ _ +end + +lemma range_finsupp_total : + (finsupp_total ι M I v).range = I • (submodule.span R (set.range v)) := +begin + apply le_antisymm, + { rintros x ⟨f, rfl⟩, + rw finsupp_total_apply, + apply submodule.sum_mem _ _, + intros c _, + apply submodule.smul_mem_smul (f c).2, + apply submodule.subset_span, + exact set.mem_range_self c }, + { rw submodule.smul_le, + rintros r hr m hm, + rw ← set.image_univ at hm, + obtain ⟨l, hl, rfl⟩ := (finsupp.mem_span_image_iff_total _).mp hm, + let l' : ι →₀ I := finsupp.map_range (λ x : R, (⟨x * r, I.mul_mem_left _ hr⟩ : I)) + (subtype.ext $ zero_mul _) l, + use l', + rw [finsupp_total_apply, finsupp.total_apply, finsupp.sum, finsupp.sum, finset.smul_sum], + dsimp, + simp only [← mul_smul, mul_comm r], + apply finset.sum_subset, + { exact finsupp.support_map_range }, + { intros x hx hx', + have : l x * r = 0 := by injection finsupp.not_mem_support_iff.mp hx', + rw [this, zero_smul] } } +end + +end total + end ideal lemma associates.mk_ne_zero' {R : Type*} [comm_semiring R] {r : R} : From 29fb82d9f18e7426e0022ff0243c26bdb26f7ed1 Mon Sep 17 00:00:00 2001 From: Patrick Stevens Date: Sun, 31 Jul 2022 11:53:14 +0000 Subject: [PATCH 152/173] feat(number_theory): Bertrand's postulate, slightly different approach (#8002) Co-authored-by: Smaug123 --- docs/100.yaml | 2 + src/data/nat/choose/central.lean | 2 +- src/data/nat/choose/factorization.lean | 29 +++ src/number_theory/bertrand.lean | 234 +++++++++++++++++++++++++ 4 files changed, 266 insertions(+), 1 deletion(-) create mode 100644 src/number_theory/bertrand.lean diff --git a/docs/100.yaml b/docs/100.yaml index c1affb6843c28..984c76cccc75f 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -390,6 +390,8 @@ author : Anne Baanen 98: title : Bertrand’s Postulate + decl : nat.bertrand + author : Bolton Bailey, Patrick Stevens 99: title : Buffon Needle Problem 100: diff --git a/src/data/nat/choose/central.lean b/src/data/nat/choose/central.lean index 220d0d421fab8..d5e9bebe5f13c 100644 --- a/src/data/nat/choose/central.lean +++ b/src/data/nat/choose/central.lean @@ -86,7 +86,7 @@ end /-- An exponential lower bound on the central binomial coefficient. -This bound is weaker than `four_pow_n_lt_n_mul_central_binom`, but it is of historical interest +This bound is weaker than `nat.four_pow_lt_mul_central_binom`, but it is of historical interest because it appears in Erdős's proof of Bertrand's postulate. -/ lemma four_pow_le_two_mul_self_mul_central_binom : ∀ (n : ℕ) (n_pos : 0 < n), diff --git a/src/data/nat/choose/factorization.lean b/src/data/nat/choose/factorization.lean index 8a13ee9a18910..0dc2dfedbecde 100644 --- a/src/data/nat/choose/factorization.lean +++ b/src/data/nat/choose/factorization.lean @@ -26,6 +26,8 @@ do not appear in the factorization of the `n`th central binomial coefficient. These results appear in the [Erdős proof of Bertrand's postulate](aigner1999proofs). -/ +open_locale big_operators + namespace nat variables {p n k : ℕ} @@ -137,4 +139,31 @@ lemma le_two_mul_of_factorization_central_binom_pos (h_pos : 0 < (central_binom n).factorization p) : p ≤ 2 * n := le_of_not_lt (pos_iff_ne_zero.mp h_pos ∘ factorization_central_binom_eq_zero_of_two_mul_lt) +/-- A binomial coefficient is the product of its prime factors, which are at most `n`. -/ +lemma prod_pow_factorization_choose (n k : ℕ) (hkn : k ≤ n) : + ∏ p in (finset.range (n + 1)), + p ^ ((nat.choose n k).factorization p) + = choose n k := +begin + nth_rewrite_rhs 0 ←factorization_prod_pow_eq_self (choose_pos hkn).ne', + rw eq_comm, + apply finset.prod_subset, + { intros p hp, + rw finset.mem_range, + contrapose! hp, + rw [finsupp.mem_support_iff, not_not, factorization_choose_eq_zero_of_lt hp] }, + { intros p _ h2, simp [not_not.1 (mt finsupp.mem_support_iff.2 h2)] }, +end + +/-- The `n`th central binomial coefficient is the product of its prime factors, which are +at most `2n`. -/ +lemma prod_pow_factorization_central_binom (n : ℕ) : + ∏ p in (finset.range (2 * n + 1)), + p ^ ((central_binom n).factorization p) + = central_binom n := +begin + apply prod_pow_factorization_choose, + linarith, +end + end nat diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean new file mode 100644 index 0000000000000..27480499ca9b3 --- /dev/null +++ b/src/number_theory/bertrand.lean @@ -0,0 +1,234 @@ +/- +Copyright (c) 2020 Patrick Stevens. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Stevens, Bolton Bailey +-/ +import data.nat.choose.factorization +import number_theory.primorial +import analysis.convex.specific_functions + +/-! +# Bertrand's Postulate + +This file contains a proof of Bertrand's postulate: That between any positive number and its +double there is a prime. + +The proof follows the outline of the Erdős proof presented in "Proofs from THE BOOK": One considers +the prime factorization of `(2 * n).choose n`, and splits the constituent primes up into various +groups, then upper bounds the contribution of each group. This upper bounds the central binomial +coefficient, and if the postulate does not hold, this upper bound conflicts with a simple lower +bound for large enough `n`. This proves the result holds for large enough `n`, and for smaller `n` +an explicit list of primes is provided which covers the remaining cases. + +As in the [Metamath implementation](carneiro2015arithmetic), we rely on some optimizations from +[Shigenori Tochiori](tochiori_bertrand). In particular we use the cleaner bound on the central +binomial coefficient given in `nat.four_pow_lt_mul_central_binom`. + +## References + +* [M. Aigner and G. M. Ziegler _Proofs from THE BOOK_][aigner1999proofs] +* [S. Tochiori, _Considering the Proof of “There is a Prime between n and 2n”_][tochiori_bertrand] +* [M. Carneiro, _Arithmetic in Metamath, Case Study: Bertrand's Postulate_][carneiro2015arithmetic] + +## Tags + +Bertrand, prime, binomial coefficients +-/ + +open_locale big_operators + +section real + +open real + +namespace bertrand + +/-- +A reified version of the `bertrand.main_inequality` below. +This is not best possible: it actually holds for 464 ≤ x. +-/ +lemma real_main_inequality {x : ℝ} (n_large : (512 : ℝ) ≤ x) : + x * (2 * x) ^ (sqrt (2 * x)) * 4 ^ (2 * x / 3) ≤ 4 ^ x := +begin + let f : ℝ → ℝ := λ x, log x + sqrt (2 * x) * log (2 * x) - log 4 / 3 * x, + have hf' : ∀ x, 0 < x → 0 < x * (2 * x) ^ sqrt (2 * x) / 4 ^ (x / 3) := + λ x h, div_pos (mul_pos h (rpow_pos_of_pos (mul_pos two_pos h) _)) (rpow_pos_of_pos four_pos _), + have hf : ∀ x, 0 < x → f x = log (x * (2 * x) ^ sqrt (2 * x) / 4 ^ (x / 3)), + { intros x h5, + have h6 := mul_pos two_pos h5, + have h7 := rpow_pos_of_pos h6 (sqrt (2 * x)), + rw [log_div (mul_pos h5 h7).ne' (rpow_pos_of_pos four_pos _).ne', log_mul h5.ne' h7.ne', + log_rpow h6, log_rpow zero_lt_four, ← mul_div_right_comm, ← mul_div, mul_comm x] }, + have h5 : 0 < x := lt_of_lt_of_le (by norm_num1) n_large, + rw [← div_le_one (rpow_pos_of_pos four_pos x), ← div_div_eq_mul_div, ← rpow_sub four_pos, + ← mul_div 2 x, mul_div_left_comm, ← mul_one_sub, (by norm_num1 : (1 : ℝ) - 2 / 3 = 1 / 3), + mul_one_div, ← log_nonpos_iff (hf' x h5), ← hf x h5], + have h : concave_on ℝ (set.Ioi 0.5) f, + { refine ((strict_concave_on_log_Ioi.concave_on.subset (set.Ioi_subset_Ioi _) + (convex_Ioi 0.5)).add ((strict_concave_on_sqrt_mul_log_Ioi.concave_on.comp_linear_map + ((2 : ℝ) • linear_map.id)).subset + (λ a ha, lt_of_eq_of_lt _ ((mul_lt_mul_left two_pos).mpr ha)) (convex_Ioi 0.5))).sub + ((convex_on_id (convex_Ioi 0.5)).smul (div_nonneg (log_nonneg _) _)); norm_num1 }, + suffices : ∃ x1 x2, 0.5 < x1 ∧ x1 < x2 ∧ x2 ≤ x ∧ 0 ≤ f x1 ∧ f x2 ≤ 0, + { obtain ⟨x1, x2, h1, h2, h0, h3, h4⟩ := this, + exact (h.right_le_of_le_left'' h1 ((h1.trans h2).trans_le h0) h2 h0 (h4.trans h3)).trans h4 }, + refine ⟨18, 512, by norm_num1, by norm_num1, le_trans (by norm_num1) n_large, _, _⟩, + { have : sqrt (2 * 18) = 6 := + (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1), + rw [hf, log_nonneg_iff (hf' 18 _), this]; norm_num1 }, + { have : sqrt (2 * 512) = 32, + { exact (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1) }, + rw [hf, log_nonpos_iff (hf' _ _), this, div_le_one (rpow_pos_of_pos four_pos _), + ← rpow_le_rpow_iff _ (rpow_pos_of_pos four_pos _).le three_pos, ← rpow_mul]; norm_num1 }, +end + +end bertrand + +end real + +section nat + +open nat + +/-- +The inequality which contradicts Bertrand's postulate, for large enough `n`. +-/ +lemma bertrand_main_inequality {n : ℕ} (n_large : 512 ≤ n) : + n * (2 * n) ^ sqrt (2 * n) * 4 ^ (2 * n / 3) ≤ 4 ^ n := +begin + rw ← @cast_le ℝ, + simp only [cast_bit0, cast_add, cast_one, cast_mul, cast_pow, ← real.rpow_nat_cast], + have n_pos : 0 < n := (dec_trivial : 0 < 512).trans_le n_large, + have n2_pos : 1 ≤ 2 * n := mul_pos dec_trivial n_pos, + refine trans (mul_le_mul _ _ _ _) (bertrand.real_main_inequality (by exact_mod_cast n_large)), + { refine mul_le_mul_of_nonneg_left _ (nat.cast_nonneg _), + refine real.rpow_le_rpow_of_exponent_le (by exact_mod_cast n2_pos) _, + exact_mod_cast real.nat_sqrt_le_real_sqrt }, + { exact real.rpow_le_rpow_of_exponent_le (by norm_num1) (cast_div_le.trans (by norm_cast)) }, + { exact real.rpow_nonneg_of_nonneg (by norm_num1) _ }, + { refine mul_nonneg (nat.cast_nonneg _) _, + exact real.rpow_nonneg_of_nonneg (mul_nonneg zero_le_two (nat.cast_nonneg _)) _, }, +end + +/-- +A lemma that tells us that, in the case where Bertrand's postulate does not hold, the prime +factorization of the central binomial coefficent only has factors at most `2 * n / 3 + 1`. +-/ +lemma central_binom_factorization_small (n : ℕ) (n_large : 2 < n) + (no_prime: ¬∃ (p : ℕ), p.prime ∧ n < p ∧ p ≤ 2 * n) : + central_binom n = ∏ p in finset.range (2 * n / 3 + 1), p ^ ((central_binom n).factorization p) := +begin + refine (eq.trans _ n.prod_pow_factorization_central_binom).symm, + apply finset.prod_subset, + { exact finset.range_subset.2 (add_le_add_right (nat.div_le_self _ _) _) }, + intros x hx h2x, + rw [finset.mem_range, lt_succ_iff] at hx h2x, + rw [not_le, div_lt_iff_lt_mul' three_pos, mul_comm x] at h2x, + replace no_prime := not_exists.mp no_prime x, + rw [←and_assoc, not_and', not_and_distrib, not_lt] at no_prime, + cases no_prime hx with h h, + { rw [factorization_eq_zero_of_non_prime n.central_binom h, pow_zero] }, + { rw [factorization_central_binom_of_two_mul_self_lt_three_mul n_large h h2x, pow_zero] }, +end + +/-- +An upper bound on the central binomial coefficient used in the proof of Bertrand's postulate. +The bound splits the prime factors of `central_binom n` into those +1. At most `sqrt (2 * n)`, which contribute at most `2 * n` for each such prime. +2. Between `sqrt (2 * n)` and `2 * n / 3`, which contribute at most `4^(2 * n / 3)` in total. +3. Between `2 * n / 3` and `n`, which do not exist. +4. Between `n` and `2 * n`, which would not exist in the case where Bertrand's postulate is false. +5. Above `2 * n`, which do not exist. +-/ +lemma central_binom_le_of_no_bertrand_prime (n : ℕ) (n_big : 2 < n) + (no_prime : ¬∃ (p : ℕ), nat.prime p ∧ n < p ∧ p ≤ 2 * n) : + central_binom n ≤ (2 * n) ^ sqrt (2 * n) * 4 ^ (2 * n / 3) := +begin + have n_pos : 0 < n := (nat.zero_le _).trans_lt n_big, + have n2_pos : 1 ≤ 2 * n := mul_pos two_pos n_pos, + let S := (finset.range (2 * n / 3 + 1)).filter nat.prime, + let f := λ x, x ^ n.central_binom.factorization x, + have : ∏ (x : ℕ) in S, f x = ∏ (x : ℕ) in finset.range (2 * n / 3 + 1), f x, + { refine finset.prod_filter_of_ne (λ p hp h, _), + contrapose! h, dsimp only [f], + rw [factorization_eq_zero_of_non_prime n.central_binom h, pow_zero] }, + 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 _, + { 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] }, + rw finset.prod_const, + refine pow_le_pow n2_pos ((finset.card_le_of_subset (λ x hx, _)).trans this.le), + obtain ⟨h1, h2⟩ := finset.mem_filter.1 hx, + exact finset.mem_Icc.mpr ⟨(finset.mem_filter.1 h1).2.one_lt.le, h2⟩ }, + { refine le_trans _ (primorial_le_4_pow (2 * n / 3)), + refine (finset.prod_le_prod' (λ p hp, (_ : f p ≤ p))).trans _, + { obtain ⟨h1, h2⟩ := finset.mem_filter.1 hp, + refine (pow_le_pow (finset.mem_filter.1 h1).2.one_lt.le _).trans (pow_one p).le, + exact nat.factorization_choose_le_one (sqrt_lt'.mp $ not_le.1 h2) }, + refine finset.prod_le_prod_of_subset_of_one_le' (finset.filter_subset _ _) _, + exact λ p hp _, (finset.mem_filter.1 hp).2.one_lt.le } +end + +namespace nat + +/-- +Proves that Bertrand's postulate holds for all sufficiently large `n`. +-/ +lemma exists_prime_lt_and_le_two_mul_eventually (n : ℕ) (n_big : 512 ≤ n) : + ∃ (p : ℕ), p.prime ∧ n < p ∧ p ≤ 2 * n := +begin + -- Assume there is no prime in the range. + by_contradiction no_prime, + -- Then we have the above sub-exponential bound on the size of this central binomial coefficient. + -- We now couple this bound with an exponential lower bound on the central binomial coefficient, + -- yielding an inequality which we have seen is false for large enough n. + have H1 : n * (2 * n) ^ sqrt (2 * n) * 4 ^ (2 * n / 3) ≤ 4 ^ n := bertrand_main_inequality n_big, + have H2 : 4 ^ n < n * n.central_binom := + nat.four_pow_lt_mul_central_binom n (le_trans (by norm_num1) n_big), + have H3 : n.central_binom ≤ (2 * n) ^ sqrt (2 * n) * 4 ^ (2 * n / 3) := + central_binom_le_of_no_bertrand_prime n (lt_of_lt_of_le (by norm_num1) n_big) no_prime, + rw mul_assoc at H1, exact not_le.2 H2 ((mul_le_mul_left' H3 n).trans H1), +end + +/-- +Proves that Bertrand's postulate holds over all positive naturals less than n by identifying a +descending list of primes, each no more than twice the next, such that the list contains a witness +for each number ≤ n. +-/ +lemma exists_prime_lt_and_le_two_mul_succ {n} (q) + {p : ℕ} (prime_p : nat.prime p) (covering : p ≤ 2 * q) + (H : n < q → ∃ (p : ℕ), p.prime ∧ n < p ∧ p ≤ 2 * n) + (hn : n < p) : ∃ (p : ℕ), p.prime ∧ n < p ∧ p ≤ 2 * n := +begin + by_cases p ≤ 2 * n, { exact ⟨p, prime_p, hn, h⟩ }, + exact H (lt_of_mul_lt_mul_left' (lt_of_lt_of_le (not_le.1 h) covering)) +end + +/-- +**Bertrand's Postulate**: For any positive natural number, there is a prime which is greater than +it, but no more than twice as large. +-/ +theorem exists_prime_lt_and_le_two_mul (n : ℕ) (hn0 : n ≠ 0) : + ∃ p, nat.prime p ∧ n < p ∧ p ≤ 2 * n := +begin + -- Split into cases whether `n` is large or small + cases lt_or_le 511 n, + -- If `n` is large, apply the lemma derived from the inequalities on the central binomial + -- coefficient. + { exact exists_prime_lt_and_le_two_mul_eventually n h, }, + replace h : n < 521 := h.trans_lt (by norm_num1), + revert h, + -- For small `n`, supply a list of primes to cover the initial cases. + ([317, 163, 83, 43, 23, 13, 7, 5, 3, 2].mmap' $ λ n, + `[refine exists_prime_lt_and_le_two_mul_succ %%(reflect n) (by norm_num1) (by norm_num1) _]), + exact λ h2, ⟨2, prime_two, h2, nat.mul_le_mul_left 2 (nat.pos_of_ne_zero hn0)⟩, +end + +alias nat.exists_prime_lt_and_le_two_mul ← bertrand + +end nat + +end nat From 5c1739f54a5be6bcb46580ece49b894699adf47b Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Sun, 31 Jul 2022 14:14:04 +0000 Subject: [PATCH 153/173] feat(data/nat/factorization/basic): add lemma `prime.factorization_self` (#15767) Adds `lemma prime.factorization_self (hp : prime p) : p.factorization p = 1`. We already have `prime.factorization (hp : prime p) : p.factorization = single p 1`, but this sometimes needs a little more help to simp down to `1`. --- src/data/nat/factorization/basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 8846a22839aef..fa807937a1fec 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -187,6 +187,10 @@ begin refl, end +/-- The only prime factor of prime `p` is `p` itself, with multiplicity `1` -/ +@[simp] lemma prime.factorization_self {p : ℕ} (hp : prime p) : p.factorization p = 1 := +by simp [hp] + /-- For prime `p` the only prime factor of `p^k` is `p` with multiplicity `k` -/ lemma prime.factorization_pow {p k : ℕ} (hp : prime p) : factorization (p ^ k) = single p k := From 0fca9dc3c45b54918541f7e5e3c5778db1bc624b Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sun, 31 Jul 2022 15:17:20 +0000 Subject: [PATCH 154/173] feat(probability/density): probability of event in uniform distribution (#15678) Co-authored-by: Bhavik Mehta --- src/probability/density.lean | 39 ++++++++++++++++++++++++++++-------- 1 file changed, 31 insertions(+), 8 deletions(-) diff --git a/src/probability/density.lean b/src/probability/density.lean index 1ffefadfb1bad..2a360c8166877 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -325,14 +325,14 @@ pdf X ℙ μ =ᵐ[μ] support.indicator ((μ support)⁻¹ • 1) namespace is_uniform lemma has_pdf {m : measurable_space α} {X : α → E} {ℙ : measure α} {μ : measure E} - {support : set E} (hns : μ support ≠ 0) (hnt : μ support ≠ ⊤) (hu : is_uniform X support ℙ μ) : + {s : set E} (hns : μ s ≠ 0) (hnt : μ s ≠ ∞) (hu : is_uniform X s ℙ μ) : has_pdf X ℙ μ := has_pdf_of_pdf_ne_zero begin intro hpdf, rw [is_uniform, hpdf] at hu, - suffices : μ (support ∩ function.support ((μ support)⁻¹ • 1)) = 0, - { have heq : function.support ((μ support)⁻¹ • (1 : E → ℝ≥0∞)) = set.univ, + suffices : μ (s ∩ function.support ((μ s)⁻¹ • 1)) = 0, + { have heq : function.support ((μ s)⁻¹ • (1 : E → ℝ≥0∞)) = set.univ, { ext x, rw [function.mem_support], simp [hnt] }, @@ -347,12 +347,35 @@ lemma pdf_to_real_ae_eq {m : measurable_space α} (λ x, (s.indicator ((μ s)⁻¹ • (1 : E → ℝ≥0∞)) x).to_real) := filter.eventually_eq.fun_comp hX ennreal.to_real -variables [is_finite_measure ℙ] {X : α → ℝ} -variables {s : set ℝ} (hms : measurable_set s) (hns : volume s ≠ 0) +lemma measure_preimage {m : measurable_space α} {X : α → E} {ℙ : measure α} {μ : measure E} + {s : set E} (hns : μ s ≠ 0) (hnt : μ s ≠ ∞) (hms : measurable_set s) + (hu : is_uniform X s ℙ μ) + {A : set E} (hA : measurable_set A) : + ℙ (X ⁻¹' A) = μ (s ∩ A) / μ s := +begin + haveI := hu.has_pdf hns hnt, + rw [←measure.map_apply (has_pdf.measurable X ℙ μ) hA, map_eq_set_lintegral_pdf X ℙ μ hA, + lintegral_congr_ae hu.restrict], + simp only [hms, hA, lintegral_indicator, pi.smul_apply, pi.one_apply, algebra.id.smul_eq_mul, + mul_one, lintegral_const, restrict_apply', set.univ_inter], + rw ennreal.div_eq_inv_mul, +end + +lemma is_probability_measure {m : measurable_space α} {X : α → E} {ℙ : measure α} {μ : measure E} + {s : set E} (hns : μ s ≠ 0) (hnt : μ s ≠ ∞) (hms : measurable_set s) + (hu : is_uniform X s ℙ μ) : + is_probability_measure ℙ := +⟨begin + have : X ⁻¹' set.univ = set.univ, { simp only [set.preimage_univ] }, + rw [←this, hu.measure_preimage hns hnt hms measurable_set.univ, set.inter_univ, + ennreal.div_self hns hnt], +end⟩ + +variables {X : α → ℝ} {s : set ℝ} (hms : measurable_set s) (hns : volume s ≠ 0) include hms hns -lemma mul_pdf_integrable (hcs : is_compact s) (huX : is_uniform X s ℙ) : +lemma mul_pdf_integrable [is_finite_measure ℙ] (hcs : is_compact s) (huX : is_uniform X s ℙ) : integrable (λ x : ℝ, x * (pdf X ℙ volume x).to_real) := begin by_cases hsupp : volume s = ∞, @@ -379,12 +402,12 @@ end /-- A real uniform random variable `X` with support `s` has expectation `(λ s)⁻¹ * ∫ x in s, x ∂λ` where `λ` is the Lebesgue measure. -/ -lemma integral_eq (hnt : volume s ≠ ⊤) (huX : is_uniform X s ℙ) : +lemma integral_eq (hnt : volume s ≠ ∞) (huX : is_uniform X s ℙ) : ∫ x, X x ∂ℙ = (volume s)⁻¹.to_real * ∫ x in s, x := begin haveI := has_pdf hns hnt huX, + haveI := huX.is_probability_measure hns hnt hms, rw ← integral_mul_eq_integral, - all_goals { try { apply_instance } }, rw integral_congr_ae (filter.eventually_eq.mul (ae_eq_refl _) (pdf_to_real_ae_eq huX)), have : ∀ x, x * (s.indicator ((volume s)⁻¹ • (1 : ℝ → ℝ≥0∞)) x).to_real = x * (s.indicator ((volume s)⁻¹.to_real • (1 : ℝ → ℝ)) x), From c7457041d6352081220b851bf335f1e30825236e Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Sun, 31 Jul 2022 15:17:21 +0000 Subject: [PATCH 155/173] chore(ring_theory/adjoin_root): remove duplicate namespace (#15775) --- 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 ef8f030aa82c5..69f6fbe3aafb0 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -107,7 +107,7 @@ def root : adjoin_root f := mk f X variables {f} -instance adjoin_root.has_coe_t : has_coe_t R (adjoin_root f) := ⟨of f⟩ +instance has_coe_t : has_coe_t R (adjoin_root f) := ⟨of f⟩ @[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 From 4209ebe1a296cdfd3cfc0efb05940b0b42a347d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 31 Jul 2022 20:14:27 +0000 Subject: [PATCH 156/173] =?UTF-8?q?feat(order/succ=5Fpred/basic):=20`a=20?= =?UTF-8?q?=E2=89=A4=20succ=20=E2=8A=A5=20=E2=86=94=20a=20=3D=20=E2=8A=A5?= =?UTF-8?q?=20=E2=88=A8=20a=20=3D=20succ=20=E2=8A=A5`=20and=20related=20(#?= =?UTF-8?q?15567)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We use these new results to golf down the ordinal API. --- src/order/succ_pred/basic.lean | 20 ++++++++++++++++++-- src/set_theory/ordinal/arithmetic.lean | 11 ++--------- src/set_theory/ordinal/basic.lean | 6 ++---- 3 files changed, 22 insertions(+), 15 deletions(-) diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index c3a1904915278..5f731b9006596 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -362,7 +362,15 @@ lt_succ_iff_not_is_max.trans not_is_max_iff_ne_top end order_top section order_bot -variables [order_bot α] [nontrivial α] +variable [order_bot α] + +@[simp] lemma lt_succ_bot_iff [no_max_order α] : a < succ ⊥ ↔ a = ⊥ := +by rw [lt_succ_iff, le_bot_iff] + +lemma le_succ_bot_iff : a ≤ succ ⊥ ↔ a = ⊥ ∨ a = succ ⊥ := +by rw [le_succ_iff_eq_or_le, le_bot_iff, or_comm] + +variable [nontrivial α] lemma bot_lt_succ (a : α) : ⊥ < succ a := (lt_succ_of_not_is_max not_is_max_bot).trans_le $ succ_mono bot_le @@ -575,7 +583,15 @@ variables [order_bot α] end order_bot section order_top -variables [order_top α] [nontrivial α] + +variable [order_top α] + +@[simp] lemma pred_top_lt_iff [no_min_order α] : pred ⊤ < a ↔ a = ⊤ := +@lt_succ_bot_iff αᵒᵈ _ _ _ _ _ + +lemma pred_top_le_iff : pred ⊤ ≤ a ↔ a = ⊤ ∨ a = pred ⊤ := @le_succ_bot_iff αᵒᵈ _ _ _ _ + +variable [nontrivial α] lemma pred_lt_top (a : α) : pred a < ⊤ := (pred_mono le_top).trans_lt $ pred_lt_of_not_is_min not_is_min_top diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 5524d109f77eb..eccff6128b1a2 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -120,8 +120,7 @@ theorem nat_cast_succ (n : ℕ) : ↑n.succ = succ (n : ordinal) := rfl theorem add_left_cancel (a) {b c : ordinal} : a + b = a + c ↔ b = c := by simp only [le_antisymm_iff, add_le_add_iff_left] -theorem lt_one_iff_zero {a : ordinal} : a < 1 ↔ a = 0 := -by rw [←succ_zero, lt_succ_iff, ordinal.le_zero] +theorem lt_one_iff_zero {a : ordinal} : a < 1 ↔ a = 0 := by simpa using @lt_succ_bot_iff _ _ _ a _ _ private theorem add_lt_add_iff_left' (a) {b c : ordinal} : a + b < a + c ↔ b < c := by rw [← not_le, ← not_le, add_le_add_iff_left] @@ -185,13 +184,7 @@ unique.eq_default x by rw [one_out_eq x, typein_enum] theorem le_one_iff {a : ordinal} : a ≤ 1 ↔ a = 0 ∨ a = 1 := -begin - refine ⟨λ ha, _, _⟩, - { rcases eq_or_lt_of_le ha with rfl | ha, - exacts [or.inr rfl, or.inl (lt_one_iff_zero.1 ha)], }, - { rintro (rfl | rfl), - exacts [zero_le_one, le_rfl], } -end +by simpa using @le_succ_bot_iff _ _ _ a _ theorem add_eq_zero_iff {a b : ordinal} : a + b = 0 ↔ (a = 0 ∧ b = 0) := induction_on a $ λ α r _, induction_on b $ λ β s _, begin diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 0a4f52e70e0c9..551a395dc6917 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -408,11 +408,9 @@ induction_on o $ λ α r _, ⟨⟨⟨embedding.of_is_empty, is_empty_elim⟩, is instance : order_bot ordinal := ⟨0, ordinal.zero_le⟩ -@[simp] protected theorem le_zero {o : ordinal} : o ≤ 0 ↔ o = 0 := -by simp only [le_antisymm_iff, ordinal.zero_le, and_true] +@[simp] protected theorem le_zero {o : ordinal} : o ≤ 0 ↔ o = 0 := le_bot_iff -protected theorem pos_iff_ne_zero {o : ordinal} : 0 < o ↔ o ≠ 0 := -by simp only [lt_iff_le_and_ne, ordinal.zero_le, true_and, ne.def, eq_comm] +protected theorem pos_iff_ne_zero {o : ordinal} : 0 < o ↔ o ≠ 0 := bot_lt_iff_ne_bot @[simp] theorem out_empty_iff_eq_zero {o : ordinal} : is_empty o.out.α ↔ o = 0 := by rw [←@type_eq_zero_iff_is_empty o.out.α (<), type_lt] From 4a529fa5779251f349ce990c713267c7360878ba Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 1 Aug 2022 08:42:17 +0000 Subject: [PATCH 157/173] fix(data/polynomial/basic): remove `monomial_fun` to remove diamonds (#15778) This was previously irreducible to prevent any `finsupp/add_monoid_algebra` leakage, but it causes a non-defeq diamond in instances. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/data/polynomial/basic.lean | 38 +++++++++------------------------- test/instance_diamonds.lean | 3 +++ 2 files changed, 13 insertions(+), 28 deletions(-) diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 3841580de0b48..f5af0efd8be91 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -83,8 +83,6 @@ they unfold around `polynomial.of_finsupp` and `polynomial.to_finsupp`. -/ section add_monoid_algebra -/-- The function version of `monomial`. Use `monomial` instead of this one. -/ -@[irreducible] def monomial_fun (n : ℕ) (a : R) : R[X] := ⟨finsupp.single n a⟩ @[irreducible] private def add : R[X] → R[X] → R[X] | ⟨a⟩ ⟨b⟩ := ⟨a + b⟩ @[irreducible] private def neg {R : Type u} [ring R] : R[X] → R[X] @@ -93,7 +91,7 @@ section add_monoid_algebra | ⟨a⟩ ⟨b⟩ := ⟨a * b⟩ instance : has_zero R[X] := ⟨⟨0⟩⟩ -instance : has_one R[X] := ⟨monomial_fun 0 (1 : R)⟩ +instance : has_one R[X] := ⟨⟨1⟩⟩ instance : has_add R[X] := ⟨add⟩ instance {R : Type u} [ring R] : has_neg R[X] := ⟨neg⟩ instance {R : Type u} [ring R] : has_sub R[X] := ⟨λ a b, a + -b⟩ @@ -103,15 +101,8 @@ instance {S : Type*} [monoid S] [distrib_mul_action S R] : has_smul S R[X] := @[priority 1] -- to avoid a bug in the `ring` tactic instance has_pow : has_pow R[X] ℕ := { pow := λ p n, npow_rec n p } -@[simp] lemma of_finsupp_zero : (⟨0⟩ : R[X]) = 0 := -rfl - -@[simp] lemma of_finsupp_one : (⟨1⟩ : R[X]) = 1 := -begin - change (⟨1⟩ : R[X]) = monomial_fun 0 (1 : R), - rw [monomial_fun], - refl -end +@[simp] lemma of_finsupp_zero : (⟨0⟩ : R[X]) = 0 := rfl +@[simp] lemma of_finsupp_one : (⟨1⟩ : R[X]) = 1 := rfl @[simp] lemma of_finsupp_add {a b} : (⟨a + b⟩ : R[X]) = ⟨a⟩ + ⟨b⟩ := show _ = add _ _, by rw add @[simp] lemma of_finsupp_neg {R : Type u} [ring R] {a} : (⟨-a⟩ : R[X]) = -⟨a⟩ := @@ -133,12 +124,7 @@ end @[simp] lemma to_finsupp_zero : (0 : R[X]).to_finsupp = 0 := rfl -@[simp] lemma to_finsupp_one : (1 : R[X]).to_finsupp = 1 := -begin - change to_finsupp (monomial_fun _ _) = _, - rw monomial_fun, - refl, -end +@[simp] lemma to_finsupp_one : (1 : R[X]).to_finsupp = 1 := rfl @[simp] lemma to_finsupp_add (a b : R[X]) : (a + b).to_finsupp = a.to_finsupp + b.to_finsupp := by { cases a, cases b, rw ←of_finsupp_add } @@ -263,17 +249,17 @@ by simp /-- `monomial s a` is the monomial `a * X^s` -/ def monomial (n : ℕ) : R →ₗ[R] R[X] := -{ to_fun := monomial_fun n, - map_add' := by simp [monomial_fun], - map_smul' := by simp [monomial_fun, ←of_finsupp_smul] } +{ to_fun := λ t, ⟨finsupp.single n t⟩, + map_add' := by simp, + map_smul' := by simp [←of_finsupp_smul] } @[simp] lemma to_finsupp_monomial (n : ℕ) (r : R) : (monomial n r).to_finsupp = finsupp.single n r := -by simp [monomial, monomial_fun] +by simp [monomial] @[simp] lemma of_finsupp_single (n : ℕ) (r : R) : (⟨finsupp.single n r⟩ : R[X]) = monomial n r := -by simp [monomial, monomial_fun] +by simp [monomial] @[simp] lemma monomial_zero_right (n : ℕ) : @@ -308,11 +294,7 @@ to_finsupp_injective $ by simp lemma monomial_injective (n : ℕ) : function.injective (monomial n : R → R[X]) := -begin - convert (to_finsupp_iso R).symm.injective.comp (single_injective n), - ext, - simp -end +(to_finsupp_iso R).symm.injective.comp (single_injective n) @[simp] lemma monomial_eq_zero_iff (t : R) (n : ℕ) : monomial n t = 0 ↔ t = 0 := diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index 75001714fb61c..9755f42425164 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -170,6 +170,9 @@ example [comm_semiring R] [nontrivial R] : polynomial.has_smul_pi' _ _ _ = (polynomial.has_smul_pi _ _ : has_smul R[X] (R → R[X])) := rfl +/-- `polynomial.algebra_of_algebra` is consistent with `algebra_nat`. -/ +example [comm_semiring R] : (polynomial.algebra_of_algebra : algebra ℕ R[X]) = algebra_nat := rfl + end polynomial /-! ## `subtype` instances -/ From 7fb60f679f81e5fd543c44a5da129c339c1905ad Mon Sep 17 00:00:00 2001 From: tb65536 Date: Mon, 1 Aug 2022 10:42:39 +0000 Subject: [PATCH 158/173] feat(topology/homeomorph): Product of sets is isomorphic to product of sets (#15664) This PR adds `homeomorph.set.prod`, analogous to `homeomorph.set.univ`. --- src/topology/homeomorph.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 33894f6bfa1a8..16e15b877d069 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -442,6 +442,15 @@ def set.univ (α : Type*) [topological_space α] : (univ : set α) ≃ₜ α := continuous_to_fun := continuous_subtype_coe, continuous_inv_fun := continuous_subtype_mk _ continuous_id } +/-- `s ×ˢ t` is homeomorphic to `s × t`. -/ +@[simps] def set.prod (s : set α) (t : set β) : ↥(s ×ˢ t) ≃ₜ s × t := +{ to_equiv := equiv.set.prod s t, + continuous_to_fun := continuous.prod_mk + (continuous_subtype_mk _ (continuous_fst.comp continuous_induced_dom)) + (continuous_subtype_mk _ (continuous_snd.comp continuous_induced_dom)), + continuous_inv_fun := continuous_subtype_mk _ (continuous.prod_mk + (continuous_induced_dom.comp continuous_fst) (continuous_induced_dom.comp continuous_snd)) } + end homeomorph /-- An inducing equiv between topological spaces is a homeomorphism. -/ From a03afc03f0e459f56fb9f5f4ac050cbd0cd8c2b9 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 1 Aug 2022 10:42:40 +0000 Subject: [PATCH 159/173] chore(data/zmod): make `zmod.int_cast_zmod_cast` `@[norm_cast]` (#15753) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I was suprised that `norm_cast` couldn't do anything with an expression including `(((r : zmod 4) : ℤ) : zmod 4)`, turns out this lemma was missing a `norm_cast` attribute. --- 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 9a49c33fa81a5..52752b7e99160 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -138,7 +138,7 @@ nat_cast_right_inverse.surjective /-- So-named because the outer coercion is `int.cast` into `zmod`. For `int.cast` into an arbitrary ring, see `zmod.int_cast_cast`. -/ -lemma int_cast_zmod_cast (a : zmod n) : ((a : ℤ) : zmod n) = a := +@[norm_cast] lemma int_cast_zmod_cast (a : zmod n) : ((a : ℤ) : zmod n) = a := begin cases n, { rw [int.cast_id a, int.cast_id a], }, From d5b8b40ca43732d5cf83b52d3b040f297177f002 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 1 Aug 2022 10:42:41 +0000 Subject: [PATCH 160/173] feat(algebra/order/floor): `fract_one` (#15785) This seems to be an appropriate `simp` lemma, but is currently missing. --- src/algebra/order/floor.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 25af371180a3d..71ab9a85379ee 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -521,6 +521,9 @@ lemma fract_lt_one (a : α) : fract a < 1 := sub_lt.1 $ sub_one_lt_floor _ @[simp] lemma fract_zero : fract (0 : α) = 0 := by rw [fract, floor_zero, cast_zero, sub_self] +@[simp] lemma fract_one : fract (1 : α) = 0 := +by simp [fract] + @[simp] lemma fract_coe (z : ℤ) : fract (z : α) = 0 := by { unfold fract, rw floor_coe, exact sub_self _ } From e5c2fab8152a435ddf7a2e4910f7f8da1503ad36 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 1 Aug 2022 12:20:50 +0000 Subject: [PATCH 161/173] feat(geometry/euclidean/oriented_angle): relation to unoriented angles (#15722) Add some versions of the result that an oriented angle is plus or minus the unoriented angle between the same two vectors. There will be a lot more lemmas to add subsequently in this area (including, in particular, lemmas about when two angles have the same or different signs, that are needed to go from equality of unoriented angles to equality of oriented angles). This just provides the starting point for adding such results. --- src/geometry/euclidean/oriented_angle.lean | 70 ++++++++++++++++++++++ 1 file changed, 70 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index ca13248f0fdf6..472c924018e6e 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -6,6 +6,7 @@ Authors: Joseph Myers import analysis.inner_product_space.orientation import analysis.inner_product_space.pi_L2 import analysis.special_functions.complex.circle +import geometry.euclidean.basic /-! # Oriented angles. @@ -39,6 +40,7 @@ generally use the definitions and results in the `orientation` namespace instead noncomputable theory open_locale real +open_locale real_inner_product_space namespace orthonormal @@ -788,6 +790,50 @@ begin rw [←submonoid.coe_mul, mul_left_inv, submonoid.coe_one, mul_one] end +/-- The inner product of two vectors is the product of the norms and the cosine of the oriented +angle between the vectors. -/ +lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : + ⟪x, y⟫ = ∥x∥ * ∥y∥ * real.angle.cos (hb.oangle x y) := +begin + by_cases hx : x = 0, { simp [hx] }, + by_cases hy : y = 0, { simp [hy] }, + rw [oangle, real.angle.cos_coe, complex.cos_arg], swap, { simp [hx, hy] }, + simp_rw [complex.abs_div, ←complex.norm_eq_abs, linear_isometry_equiv.norm_map, complex.div_re, + ←complex.sq_abs, ←complex.norm_eq_abs, linear_isometry_equiv.norm_map, + complex.isometry_of_orthonormal_symm_apply, complex.add_re, complex.add_im, + is_R_or_C.I, complex.mul_I_re, complex.mul_I_im, complex.of_real_re, + complex.of_real_im, basis.coord_apply, neg_zero, zero_add, add_zero], + conv_lhs { rw [←b.sum_repr x, ←b.sum_repr y] }, + simp_rw [hb.inner_sum, (dec_trivial : (finset.univ : finset (fin 2)) = {0, 1}), + star_ring_end_apply, star_trivial], + rw [finset.sum_insert (dec_trivial : (0 : fin 2) ∉ ({1} : finset (fin 2))), + finset.sum_singleton], + field_simp [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy], + ring +end + +/-- The cosine of the oriented angle between two nonzero vectors is the inner product divided by +the product of the norms. -/ +lemma cos_oangle_eq_inner_div_norm_mul_norm {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + real.angle.cos (hb.oangle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := +begin + rw hb.inner_eq_norm_mul_norm_mul_cos_oangle, + field_simp [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy], + ring +end + +/-- The cosine of the oriented angle between two nonzero vectors equals that of the unoriented +angle. -/ +lemma cos_oangle_eq_cos_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + real.angle.cos (hb.oangle x y) = real.cos (inner_product_geometry.angle x y) := +by rw [hb.cos_oangle_eq_inner_div_norm_mul_norm hx hy, inner_product_geometry.cos_angle] + +/-- The oriented angle between two nonzero vectors is plus or minus the unoriented angle. -/ +lemma oangle_eq_angle_or_eq_neg_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + hb.oangle x y = inner_product_geometry.angle x y ∨ + hb.oangle x y = -inner_product_geometry.angle x y := +real.angle.cos_eq_real_cos_iff_eq_or_eq_neg.1 $ hb.cos_oangle_eq_cos_angle hx hy + end orthonormal namespace orientation @@ -1215,4 +1261,28 @@ begin simp_rw orientation.fin_orthonormal_basis_orientation end +/-- The inner product of two vectors is the product of the norms and the cosine of the oriented +angle between the vectors. -/ +lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : + ⟪x, y⟫ = ∥x∥ * ∥y∥ * real.angle.cos (o.oangle x y) := +(ob).inner_eq_norm_mul_norm_mul_cos_oangle x y + +/-- The cosine of the oriented angle between two nonzero vectors is the inner product divided by +the product of the norms. -/ +lemma cos_oangle_eq_inner_div_norm_mul_norm {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + real.angle.cos (o.oangle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := +(ob).cos_oangle_eq_inner_div_norm_mul_norm hx hy + +/-- The cosine of the oriented angle between two nonzero vectors equals that of the unoriented +angle. -/ +lemma cos_oangle_eq_cos_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + real.angle.cos (o.oangle x y) = real.cos (inner_product_geometry.angle x y) := +(ob).cos_oangle_eq_cos_angle hx hy + +/-- The oriented angle between two nonzero vectors is plus or minus the unoriented angle. -/ +lemma oangle_eq_angle_or_eq_neg_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + o.oangle x y = inner_product_geometry.angle x y ∨ + o.oangle x y = -inner_product_geometry.angle x y := +(ob).oangle_eq_angle_or_eq_neg_angle hx hy + end orientation From 58b8f8089a3a2cc0718f9de8a0a3005f5e140bb8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 1 Aug 2022 13:04:56 +0000 Subject: [PATCH 162/173] feat(order/rel_iso): relation embeddings for `sum.lex` and `prod.lex` (#15355) --- src/order/rel_iso.lean | 59 +++++++++++++++++++++++++++++++++++++++++- 1 file changed, 58 insertions(+), 1 deletion(-) diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index 1a93b746ec328..88527c486f303 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -36,7 +36,8 @@ isomorphisms. open function universes u v w -variables {α β γ : Type*} {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} +variables {α β γ δ : Type*} + {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} {u : δ → δ → Prop} /-- A relation homomorphism with respect to a given pair of relations `r` and `s` is a function `f : α → β` such that `r a b → s (f a) (f b)`. -/ @@ -212,6 +213,8 @@ initialize_simps_projections rel_embedding (to_embedding_to_fun → apply, -to_e theorem injective (f : r ↪r s) : injective f := f.inj' +@[simp] theorem inj (f : r ↪r s) {a b} : f a = f b ↔ a = b := f.injective.eq_iff + theorem map_rel_iff (f : r ↪r s) : ∀ {a b}, s (f a) (f b) ↔ r a b := f.map_rel_iff' @[simp] theorem coe_fn_mk (f : α ↪ β) (o) : @@ -341,6 +344,60 @@ end def of_is_empty (r : α → α → Prop) (s : β → β → Prop) [is_empty α] : r ↪r s := ⟨embedding.of_is_empty, is_empty_elim⟩ +/-- `sum.inl` as a relation embedding into `sum.lift_rel r s`. -/ +@[simps] def sum_lift_rel_inl (r : α → α → Prop) (s : β → β → Prop) : r ↪r sum.lift_rel r s := +{ to_fun := sum.inl, + inj' := sum.inl_injective, + map_rel_iff' := λ a b, sum.lift_rel_inl_inl } + +/-- `sum.inr` as a relation embedding into `sum.lift_rel r s`. -/ +@[simps] def sum_lift_rel_inr (r : α → α → Prop) (s : β → β → Prop) : s ↪r sum.lift_rel r s := +{ to_fun := sum.inr, + inj' := sum.inr_injective, + map_rel_iff' := λ a b, sum.lift_rel_inr_inr } + +/-- `sum.map` as a relation embedding between `sum.lift_rel` relations. -/ +@[simps] def sum_lift_rel_map (f : r ↪r s) (g : t ↪r u) : sum.lift_rel r t ↪r sum.lift_rel s u := +{ to_fun := sum.map f g, + inj' := f.injective.sum_map g.injective, + map_rel_iff' := by { rintro (a | b) (c | d); simp [f.map_rel_iff, g.map_rel_iff] } } + +/-- `sum.inl` as a relation embedding into `sum.lex r s`. -/ +@[simps] def sum_lex_inl (r : α → α → Prop) (s : β → β → Prop) : r ↪r sum.lex r s := +{ to_fun := sum.inl, + inj' := sum.inl_injective, + map_rel_iff' := λ a b, sum.lex_inl_inl } + +/-- `sum.inr` as a relation embedding into `sum.lex r s`. -/ +@[simps] def sum_lex_inr (r : α → α → Prop) (s : β → β → Prop) : s ↪r sum.lex r s := +{ to_fun := sum.inr, + inj' := sum.inr_injective, + map_rel_iff' := λ a b, sum.lex_inr_inr } + +/-- `sum.map` as a relation embedding between `sum.lex` relations. -/ +@[simps] def sum_lex_map (f : r ↪r s) (g : t ↪r u) : sum.lex r t ↪r sum.lex s u := +{ to_fun := sum.map f g, + inj' := f.injective.sum_map g.injective, + map_rel_iff' := by { rintro (a | b) (c | d); simp [f.map_rel_iff, g.map_rel_iff] } } + +/-- `λ b, prod.mk a b` as a relation embedding. -/ +@[simps] def prod_lex_mk_left (s : β → β → Prop) {a : α} (h : ¬ r a a) : s ↪r prod.lex r s := +{ to_fun := prod.mk a, + inj' := prod.mk.inj_left a, + map_rel_iff' := λ b₁ b₂, by simp [prod.lex_def, h] } + +/-- `λ a, prod.mk a b` as a relation embedding. -/ +@[simps] def prod_lex_mk_right (r : α → α → Prop) {b : β} (h : ¬ s b b) : r ↪r prod.lex r s := +{ to_fun := λ a, (a, b), + inj' := prod.mk.inj_right b, + map_rel_iff' := λ a₁ a₂, by simp [prod.lex_def, h] } + +/-- `prod.map` as a relation embedding. -/ +@[simps] def prod_lex_map (f : r ↪r s) (g : t ↪r u) : prod.lex r t ↪r prod.lex s u := +{ to_fun := prod.map f g, + inj' := f.injective.prod_map g.injective, + map_rel_iff' := λ a b, by simp [prod.lex_def, f.map_rel_iff, g.map_rel_iff] } + end rel_embedding /-- A relation isomorphism is an equivalence that is also a relation embedding. -/ From dab828766a1c01afb7a51de89ebe511d8a111f73 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 1 Aug 2022 15:32:47 +0000 Subject: [PATCH 163/173] refactor(category_theory): custom structure for full_subcategory (#14767) Full subcategories are now a custom structure rather than the usual subtype. The advantage of this is that we don't have to fight the `simp`-normal form of subtypes, as the coercion does more harm than good for full subcategories. We saw a similar refactor for discrete categories, and in both cases, erring on the side of explicitness seems to pay off. --- src/algebra/category/FinVect.lean | 37 +++++++------------ src/algebra/category/FinVect/limits.lean | 2 +- src/algebraic_geometry/AffineScheme.lean | 18 ++++++--- src/algebraic_geometry/properties.lean | 2 +- src/algebraic_topology/simplex_category.lean | 2 +- .../adjunction/reflective.lean | 14 ++++--- .../concrete_category/basic.lean | 4 +- src/category_theory/connected_components.lean | 4 +- src/category_theory/essential_image.lean | 13 ++++--- src/category_theory/full_subcategory.lean | 26 ++++++++----- src/category_theory/monoidal/subcategory.lean | 22 +++++------ .../preadditive/additive_functor.lean | 2 +- src/category_theory/sites/sieves.lean | 2 +- src/category_theory/subobject/basic.lean | 2 +- src/category_theory/subobject/mono_over.lean | 14 +++---- src/category_theory/subterminal.lean | 4 +- src/representation_theory/fdRep.lean | 6 +-- src/topology/category/Top/open_nhds.lean | 6 +-- .../sheaf_condition/opens_le_cover.lean | 9 ++--- src/topology/sheaves/sheafify.lean | 4 +- src/topology/sheaves/stalks.lean | 5 +-- 21 files changed, 101 insertions(+), 97 deletions(-) diff --git a/src/algebra/category/FinVect.lean b/src/algebra/category/FinVect.lean index f78a110c0bf5d..39dd0267240df 100644 --- a/src/algebra/category/FinVect.lean +++ b/src/algebra/category/FinVect.lean @@ -32,21 +32,15 @@ instance monoidal_predicate_finite_dimensional : prop_tensor' := λ X Y hX hY, by exactI module.finite.tensor_product K X Y } /-- Define `FinVect` as the subtype of `Module.{u} K` of finite dimensional vector spaces. -/ -@[derive [large_category, λ α, has_coe_to_sort α (Sort*), concrete_category, monoidal_category, - symmetric_category]] -def FinVect := { V : Module.{u} K // finite_dimensional K V } +@[derive [large_category, concrete_category, monoidal_category, symmetric_category]] +def FinVect := full_subcategory (λ (V : Module.{u} K), finite_dimensional K V) namespace FinVect -instance finite_dimensional (V : FinVect K) : finite_dimensional K V := V.prop +instance finite_dimensional (V : FinVect K) : finite_dimensional K V.obj := V.property instance : inhabited (FinVect K) := ⟨⟨Module.of K K, finite_dimensional.finite_dimensional_self K⟩⟩ -instance : has_coe (FinVect.{u} K) (Module.{u} K) := { coe := λ V, V.1, } - -protected lemma coe_comp {U V W : FinVect K} (f : U ⟶ V) (g : V ⟶ W) : - ((f ≫ g) : U → W) = (g : V → W) ∘ (f : U → V) := rfl - /-- Lift an unbundled vector space to `FinVect K`. -/ def of (V : Type u) [add_comm_group V] [module K V] [finite_dimensional K V] : FinVect K := ⟨Module.of K V, by { change finite_dimensional K V, apply_instance }⟩ @@ -61,42 +55,39 @@ variables (V : FinVect K) /-- The dual module is the dual in the rigid monoidal category `FinVect K`. -/ def FinVect_dual : FinVect K := -⟨Module.of K (module.dual K V), subspace.module.dual.finite_dimensional⟩ - -instance : has_coe_to_fun (FinVect_dual K V) (λ _, V → K) := -{ coe := λ v, by { change V →ₗ[K] K at v, exact v, } } +⟨Module.of K (module.dual K V.obj), subspace.module.dual.finite_dimensional⟩ open category_theory.monoidal_category /-- The coevaluation map is defined in `linear_algebra.coevaluation`. -/ def FinVect_coevaluation : 𝟙_ (FinVect K) ⟶ V ⊗ (FinVect_dual K V) := -by apply coevaluation K V +by apply coevaluation K V.obj lemma FinVect_coevaluation_apply_one : FinVect_coevaluation K V (1 : K) = - ∑ (i : basis.of_vector_space_index K V), - (basis.of_vector_space K V) i ⊗ₜ[K] (basis.of_vector_space K V).coord i := -by apply coevaluation_apply_one K V + ∑ (i : basis.of_vector_space_index K V.obj), + (basis.of_vector_space K V.obj) i ⊗ₜ[K] (basis.of_vector_space K V.obj).coord i := +by apply coevaluation_apply_one K V.obj /-- The evaluation morphism is given by the contraction map. -/ def FinVect_evaluation : (FinVect_dual K V) ⊗ V ⟶ 𝟙_ (FinVect K) := -by apply contract_left K V +by apply contract_left K V.obj @[simp] -lemma FinVect_evaluation_apply (f : (FinVect_dual K V)) (x : V) : - (FinVect_evaluation K V) (f ⊗ₜ x) = f x := +lemma FinVect_evaluation_apply (f : (FinVect_dual K V).obj) (x : V.obj) : + (FinVect_evaluation K V) (f ⊗ₜ x) = f.to_fun x := by apply contract_left_apply f x private theorem coevaluation_evaluation : let V' : FinVect K := FinVect_dual K V in (𝟙 V' ⊗ (FinVect_coevaluation K V)) ≫ (α_ V' V V').inv ≫ (FinVect_evaluation K V ⊗ 𝟙 V') = (ρ_ V').hom ≫ (λ_ V').inv := -by apply contract_left_assoc_coevaluation K V +by apply contract_left_assoc_coevaluation K V.obj private theorem evaluation_coevaluation : (FinVect_coevaluation K V ⊗ 𝟙 V) ≫ (α_ V (FinVect_dual K V) V).hom ≫ (𝟙 V ⊗ FinVect_evaluation K V) = (λ_ V).hom ≫ (ρ_ V).inv := -by apply contract_left_assoc_coevaluation' K V +by apply contract_left_assoc_coevaluation' K V.obj instance exact_pairing : exact_pairing V (FinVect_dual K V) := { coevaluation := FinVect_coevaluation K V, @@ -112,7 +103,7 @@ variables {K V} (W : FinVect K) /-- Converts and isomorphism in the category `FinVect` to a `linear_equiv` between the underlying vector spaces. -/ -def iso_to_linear_equiv {V W : FinVect K} (i : V ≅ W) : V ≃ₗ[K] W := +def iso_to_linear_equiv {V W : FinVect K} (i : V ≅ W) : V.obj ≃ₗ[K] W.obj := ((forget₂ (FinVect.{u} K) (Module.{u} K)).map_iso i).to_linear_equiv lemma iso.conj_eq_conj {V W : FinVect K} (i : V ≅ W) (f : End V) : diff --git a/src/algebra/category/FinVect/limits.lean b/src/algebra/category/FinVect/limits.lean index e27d7c05aa3c2..70188668e71db 100644 --- a/src/algebra/category/FinVect/limits.lean +++ b/src/algebra/category/FinVect/limits.lean @@ -48,7 +48,7 @@ instance (F : J ⥤ FinVect k) : finite_dimensional k (limit (F ⋙ forget₂ (FinVect k) (Module.{v} k)) : Module.{v} k) := begin haveI : ∀ j, finite_dimensional k ((F ⋙ forget₂ (FinVect k) (Module.{v} k)).obj j), - { intro j, change finite_dimensional k (F.obj j), apply_instance, }, + { intro j, change finite_dimensional k (F.obj j).obj, apply_instance, }, exact finite_dimensional.of_injective (limit_subobject_product (F ⋙ forget₂ (FinVect k) (Module.{v} k))) ((Module.mono_iff_injective _).1 (by apply_instance)), diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 4444625b845fe..fb26ada924bac 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -41,7 +41,8 @@ namespace algebraic_geometry open Spec (structure_sheaf) /-- The category of affine schemes -/ -def AffineScheme := Scheme.Spec.ess_image +@[derive category, nolint has_nonempty_instance] +def AffineScheme := Scheme.Spec.ess_image_subcategory /-- A Scheme is affine if the canonical map `X ⟶ Spec Γ(X)` is an isomorphism. -/ class is_affine (X : Scheme) : Prop := @@ -54,18 +55,23 @@ def Scheme.iso_Spec (X : Scheme) [is_affine X] : X ≅ Scheme.Spec.obj (op $ Scheme.Γ.obj $ op X) := as_iso (Γ_Spec.adjunction.unit.app X) -lemma mem_AffineScheme (X : Scheme) : X ∈ AffineScheme ↔ is_affine X := +/-- Construct an affine scheme from a scheme and the information that it is affine. -/ +@[simps] +def AffineScheme.mk (X : Scheme) (h : is_affine X) : AffineScheme := +⟨X, @@mem_ess_image_of_unit_is_iso _ _ _ _ h.1⟩ + +lemma mem_Spec_ess_image (X : Scheme) : X ∈ Scheme.Spec.ess_image ↔ is_affine X := ⟨λ h, ⟨functor.ess_image.unit_is_iso h⟩, λ h, @@mem_ess_image_of_unit_is_iso _ _ _ X h.1⟩ -instance is_affine_AffineScheme (X : AffineScheme.{u}) : is_affine (X : Scheme.{u}) := -(mem_AffineScheme _).mp X.prop +instance is_affine_AffineScheme (X : AffineScheme.{u}) : is_affine X.obj := +⟨functor.ess_image.unit_is_iso X.property⟩ instance Spec_is_affine (R : CommRingᵒᵖ) : is_affine (Scheme.Spec.obj R) := -(mem_AffineScheme _).mp (Scheme.Spec.obj_mem_ess_image R) +algebraic_geometry.is_affine_AffineScheme ⟨_, Scheme.Spec.obj_mem_ess_image R⟩ lemma is_affine_of_iso {X Y : Scheme} (f : X ⟶ Y) [is_iso f] [h : is_affine Y] : is_affine X := -by { rw [← mem_AffineScheme] at h ⊢, exact functor.ess_image.of_iso (as_iso f).symm h } +by { rw [← mem_Spec_ess_image] at h ⊢, exact functor.ess_image.of_iso (as_iso f).symm h } namespace AffineScheme diff --git a/src/algebraic_geometry/properties.lean b/src/algebraic_geometry/properties.lean index 4043ecffb4a34..a942db176b528 100644 --- a/src/algebraic_geometry/properties.lean +++ b/src/algebraic_geometry/properties.lean @@ -34,7 +34,7 @@ begin rw t0_space_iff_inseparable, intros x y h, obtain ⟨U, R, ⟨e⟩⟩ := X.local_affine x, - have hy : y ∈ U.val := (h.mem_open_iff U.1.2).1 U.2, + have hy : y ∈ U.obj := (h.mem_open_iff U.1.2).1 U.2, erw ← subtype_inseparable_iff (⟨x, U.2⟩ : U.1.1) (⟨y, hy⟩ : U.1.1) at h, let e' : U.1 ≃ₜ prime_spectrum R := homeo_of_iso ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _).map_iso e), diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index 0dc093321ccff..335a502fd6a9d 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -382,7 +382,7 @@ def is_skeleton_of : is_skeleton_of NonemptyFinLinOrd simplex_category skeletal_ /-- The truncated simplex category. -/ @[derive small_category] -def truncated (n : ℕ) := {a : simplex_category // a.len ≤ n} +def truncated (n : ℕ) := full_subcategory (λ a : simplex_category, a.len ≤ n) namespace truncated diff --git a/src/category_theory/adjunction/reflective.lean b/src/category_theory/adjunction/reflective.lean index 67ce6368ea73d..39f172553e620 100644 --- a/src/category_theory/adjunction/reflective.lean +++ b/src/category_theory/adjunction/reflective.lean @@ -156,17 +156,19 @@ by rw [←equiv.eq_symm_apply, unit_comp_partial_bijective_symm_natural A h, equ /-- If `i : D ⥤ C` is reflective, the inverse functor of `i ≌ F.ess_image` can be explicitly defined by the reflector. -/ @[simps] -def equiv_ess_image_of_reflective [reflective i] : D ≌ i.ess_image := +def equiv_ess_image_of_reflective [reflective i] : D ≌ i.ess_image_subcategory := { functor := i.to_ess_image, inverse := i.ess_image_inclusion ⋙ (left_adjoint i : _), unit_iso := nat_iso.of_components (λ X, (as_iso $ (of_right_adjoint i).counit.app X).symm) (by { intros X Y f, dsimp, simp only [is_iso.eq_inv_comp, is_iso.comp_inv_eq, category.assoc], exact ((of_right_adjoint i).counit.naturality _).symm }), - counit_iso := nat_iso.of_components - (λ X, by { refine (iso.symm $ as_iso _), exact (of_right_adjoint i).unit.app X, + counit_iso := + nat_iso.of_components + (λ X, by { refine (iso.symm $ as_iso _), exact (of_right_adjoint i).unit.app X.obj, apply_with (is_iso_of_reflects_iso _ i.ess_image_inclusion) { instances := ff }, - exact functor.ess_image.unit_is_iso X.prop }) - (by { intros X Y f, dsimp, simp only [is_iso.eq_inv_comp, is_iso.comp_inv_eq, category.assoc], - exact ((of_right_adjoint i).unit.naturality f).symm }) } + exact functor.ess_image.unit_is_iso X.property }) + (by { intros X Y f, dsimp, rw [is_iso.comp_inv_eq, assoc], + have h := ((of_right_adjoint i).unit.naturality f).symm, + rw [functor.id_map] at h, erw [← h, is_iso.inv_hom_id_assoc, functor.comp_map] }) } end category_theory diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index bf12a8c81f664..29e9eb2079c0b 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -171,11 +171,11 @@ instance induced_category.has_forget₂ {C : Type v} {D : Type v'} [category D] forget_comp := rfl } instance full_subcategory.concrete_category {C : Type v} [category C] [concrete_category C] - (Z : C → Prop) : concrete_category {X : C // Z X} := + (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] - (Z : C → Prop) : has_forget₂ {X : C // Z X} C := + (Z : C → Prop) : has_forget₂ (full_subcategory Z) C := { forget₂ := full_subcategory_inclusion Z, forget_comp := rfl } diff --git a/src/category_theory/connected_components.lean b/src/category_theory/connected_components.lean index e4f1e2bc3d2f1..ba3c0b93c162f 100644 --- a/src/category_theory/connected_components.lean +++ b/src/category_theory/connected_components.lean @@ -41,7 +41,7 @@ instance [inhabited J] : inhabited (connected_components J) := ⟨quotient.mk' d /-- Given an index for a connected component, produce the actual component as a full subcategory. -/ @[derive category] -def component (j : connected_components J) : Type u₁ := {k : J // quotient.mk' k = j} +def component (j : connected_components J) : Type u₁ := full_subcategory (λ k, quotient.mk' k = j) /-- The inclusion functor from a connected component to the whole category. -/ @[derive [full, faithful], simps {rhs_md := semireducible}] @@ -82,7 +82,7 @@ begin { refine @@list.chain_pmap_of_chain _ _ _ f (λ x y _ _ h, _) hl₁ h₁₂ _, exact zag_of_zag_obj (component.ι _) h }, { erw list.last_pmap _ f (j₁ :: l) (by simpa [h₁₂] using hf) (list.cons_ne_nil _ _), - exact subtype.ext hl₂ }, + exact full_subcategory.ext _ _ hl₂ }, end /-- diff --git a/src/category_theory/essential_image.lean b/src/category_theory/essential_image.lean index 7a832885783c3..e5cbab18144e6 100644 --- a/src/category_theory/essential_image.lean +++ b/src/category_theory/essential_image.lean @@ -66,11 +66,13 @@ set.ext $ λ A, ⟨ess_image.of_nat_iso h, ess_image.of_nat_iso h.symm⟩ /-- An object in the image is in the essential image. -/ lemma obj_mem_ess_image (F : D ⥤ C) (Y : D) : F.obj Y ∈ ess_image F := ⟨Y, ⟨iso.refl _⟩⟩ -instance : category F.ess_image := category_theory.full_subcategory _ +/-- The essential image of a functor, interpreted of a full subcategory of the target category. -/ +@[derive category, nolint has_nonempty_instance] +def ess_image_subcategory (F : C ⥤ D) := full_subcategory F.ess_image /-- The essential image as a subcategory has a fully faithful inclusion into the target category. -/ @[derive [full, faithful], simps] -def ess_image_inclusion (F : C ⥤ D) : F.ess_image ⥤ D := +def ess_image_inclusion (F : C ⥤ D) : F.ess_image_subcategory ⥤ D := full_subcategory_inclusion _ /-- @@ -78,9 +80,8 @@ Given a functor `F : C ⥤ D`, we have an (essentially surjective) functor from image of `F`. -/ @[simps] -def to_ess_image (F : C ⥤ D) : C ⥤ F.ess_image := -{ obj := λ X, ⟨_, obj_mem_ess_image _ X⟩, - map := λ X Y f, (ess_image_inclusion F).preimage (F.map f) } +def to_ess_image (F : C ⥤ D) : C ⥤ F.ess_image_subcategory := +full_subcategory.lift _ F (obj_mem_ess_image _) /-- The functor `F` factorises through its essential image, where the first functor is essentially @@ -89,7 +90,7 @@ surjective and the second is fully faithful. @[simps] def to_ess_image_comp_essential_image_inclusion (F : C ⥤ D) : F.to_ess_image ⋙ F.ess_image_inclusion ≅ F := -nat_iso.of_components (λ X, iso.refl _) (by tidy) +full_subcategory.lift_comp_inclusion _ _ _ end functor diff --git a/src/category_theory/full_subcategory.lean b/src/category_theory/full_subcategory.lean index e34c71dcd3e33..52397e39da4f6 100644 --- a/src/category_theory/full_subcategory.lean +++ b/src/category_theory/full_subcategory.lean @@ -81,35 +81,41 @@ variables {C : Type u₁} [category.{v} C] variables (Z : C → Prop) /-- -The category structure on a subtype; morphisms just ignore the property. +A subtype-like structure for full subcategories. Morphisms just ignore the property. We don't use +actual subtypes since the simp-normal form `↑X` of `X.val` does not work well for full +subcategories. See . We do not define 'strictly full' subcategories. -/ -instance full_subcategory : category.{v} {X : C // Z X} := -induced_category.category subtype.val +@[ext, nolint has_nonempty_instance] structure full_subcategory := +(obj : C) +(property : Z obj) + +instance full_subcategory.category : category.{v} (full_subcategory Z) := +induced_category.category full_subcategory.obj /-- The forgetful functor from a full subcategory into the original category ("forgetting" the condition). -/ -def full_subcategory_inclusion : {X : C // Z X} ⥤ C := -induced_functor subtype.val +def full_subcategory_inclusion : full_subcategory Z ⥤ C := +induced_functor full_subcategory.obj @[simp] lemma full_subcategory_inclusion.obj {X} : - (full_subcategory_inclusion Z).obj X = X.val := rfl + (full_subcategory_inclusion Z).obj X = X.obj := rfl @[simp] lemma full_subcategory_inclusion.map {X Y} {f : X ⟶ Y} : (full_subcategory_inclusion Z).map f = f := rfl instance full_subcategory.full : full (full_subcategory_inclusion Z) := -induced_category.full subtype.val +induced_category.full _ instance full_subcategory.faithful : faithful (full_subcategory_inclusion Z) := -induced_category.faithful subtype.val +induced_category.faithful _ variables {Z} {Z' : C → Prop} /-- An implication of predicates `Z → Z'` induces a functor between full subcategories. -/ @[simps] -def full_subcategory.map (h : ∀ ⦃X⦄, Z X → Z' X) : {X // Z X} ⥤ {X // Z' X} := +def full_subcategory.map (h : ∀ ⦃X⦄, Z X → Z' X) : full_subcategory Z ⥤ full_subcategory Z' := { obj := λ X, ⟨X.1, h X.2⟩, map := λ X Y f, f } @@ -128,7 +134,7 @@ variables {D : Type u₂} [category.{v₂} D] (P Q : D → Prop) /-- A functor which maps objects to objects satisfying a certain property induces a lift through the full subcategory of objects satisfying that property. -/ @[simps] -def full_subcategory.lift (F : C ⥤ D) (hF : ∀ X, P (F.obj X)) : C ⥤ {X // P X} := +def full_subcategory.lift (F : C ⥤ D) (hF : ∀ X, P (F.obj X)) : C ⥤ full_subcategory P := { obj := λ X, ⟨F.obj X, hF X⟩, map := λ X Y f, F.map f } diff --git a/src/category_theory/monoidal/subcategory.lean b/src/category_theory/monoidal/subcategory.lean index 8fef005cad002..2944ad4c599d5 100644 --- a/src/category_theory/monoidal/subcategory.lean +++ b/src/category_theory/monoidal/subcategory.lean @@ -45,11 +45,11 @@ open monoidal_predicate variables [monoidal_predicate P] /-- -When `P` is a monoidal predicate, the full subcategory `{X : C // P X}` inherits the monoidal -structure of `C` +When `P` is a monoidal predicate, the full subcategory for `P` inherits the monoidal structure of + `C`. -/ -instance full_monoidal_subcategory : monoidal_category {X : C // P X} := -{ tensor_obj := λ X Y, ⟨X ⊗ Y, prop_tensor X.2 Y.2⟩, +instance full_monoidal_subcategory : monoidal_category (full_subcategory P) := +{ tensor_obj := λ X Y, ⟨X.1 ⊗ Y.1, prop_tensor X.2 Y.2⟩, tensor_hom := λ X₁ Y₁ X₂ Y₂ f g, by { change X₁.1 ⊗ X₂.1 ⟶ Y₁.1 ⊗ Y₂.1, change X₁.1 ⟶ Y₁.1 at f, change X₂.1 ⟶ Y₂.1 at g, exact f ⊗ g }, tensor_unit := ⟨𝟙_ C, prop_id⟩, @@ -71,7 +71,7 @@ The forgetful monoidal functor from a full monoidal subcategory into the origina ("forgetting" the condition). -/ @[simps] -def full_monoidal_subcategory_inclusion : monoidal_functor {X : C // P X} C := +def full_monoidal_subcategory_inclusion : monoidal_functor (full_subcategory P) C := { to_functor := full_subcategory_inclusion P, ε := 𝟙 _, μ := λ X Y, 𝟙 _ } @@ -87,7 +87,7 @@ variables {P} {P' : C → Prop} [monoidal_predicate P'] subcategories. -/ @[simps] def full_monoidal_subcategory.map (h : ∀ ⦃X⦄, P X → P' X) : - monoidal_functor {X : C // P X} {X : C // P' X} := + monoidal_functor (full_subcategory P) (full_subcategory P') := { to_functor := full_subcategory.map h, ε := 𝟙 _, μ := λ X Y, 𝟙 _ } @@ -102,9 +102,9 @@ section braided variables (P) [braided_category C] /-- -The braided structure on `{X : C // P X}` inherited by the braided structure on `C`. +The braided structure on a full subcategory inherited by the braided structure on `C`. -/ -instance full_braided_subcategory : braided_category {X : C // P X} := +instance full_braided_subcategory : braided_category (full_subcategory P) := braided_category_of_faithful (full_monoidal_subcategory_inclusion P) (λ X Y, ⟨(β_ X.1 Y.1).hom, (β_ X.1 Y.1).inv, (β_ X.1 Y.1).hom_inv_id, (β_ X.1 Y.1).inv_hom_id⟩) (λ X Y, by tidy) @@ -114,7 +114,7 @@ The forgetful braided functor from a full braided subcategory into the original ("forgetting" the condition). -/ @[simps] -def full_braided_subcategory_inclusion : braided_functor {X : C // P X} C := +def full_braided_subcategory_inclusion : braided_functor (full_subcategory P) C := { to_monoidal_functor := full_monoidal_subcategory_inclusion P, braided' := λ X Y, by { rw [is_iso.eq_inv_comp], tidy } } @@ -129,7 +129,7 @@ variables {P} subcategories. -/ @[simps] def full_braided_subcategory.map (h : ∀ ⦃X⦄, P X → P' X) : - braided_functor {X : C // P X} {X : C // P' X} := + braided_functor (full_subcategory P) (full_subcategory P') := { to_monoidal_functor := full_monoidal_subcategory.map h, braided' := λ X Y, by { rw [is_iso.eq_inv_comp], tidy } } @@ -144,7 +144,7 @@ section symmetric variables (P) [symmetric_category C] -instance full_symmetric_subcategory : symmetric_category {X : C // P X} := +instance full_symmetric_subcategory : symmetric_category (full_subcategory P) := symmetric_category_of_faithful (full_braided_subcategory_inclusion P) end symmetric diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index bf80733103044..d5119cad4094f 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -147,7 +147,7 @@ variables (C D : Type*) [category C] [category D] [preadditive C] [preadditive D /-- Bundled additive functors. -/ @[derive category, nolint has_nonempty_instance] def AdditiveFunctor := -{ F : C ⥤ D // functor.additive F } +full_subcategory (λ (F : C ⥤ D), F.additive) infixr ` ⥤+ `:26 := AdditiveFunctor diff --git a/src/category_theory/sites/sieves.lean b/src/category_theory/sites/sieves.lean index 0eaba6acc1718..f0453fb0c5c4d 100644 --- a/src/category_theory/sites/sieves.lean +++ b/src/category_theory/sites/sieves.lean @@ -44,7 +44,7 @@ instance : inhabited (presieve X) := ⟨⊤⟩ /-- Given a sieve `S` on `X : C`, its associated diagram `S.diagram` is defined to be the natural functor from the full subcategory of the over category `C/X` consisting of arrows in `S` to `C`. -/ -abbreviation diagram (S : presieve X) : {f : over X // S f.hom} ⥤ C := +abbreviation diagram (S : presieve X) : full_subcategory (λ (f : over X), S f.hom) ⥤ C := full_subcategory_inclusion _ ⋙ over.forget X /-- Given a sieve `S` on `X : C`, its associated cocone `S.cocone` is defined to be diff --git a/src/category_theory/subobject/basic.lean b/src/category_theory/subobject/basic.lean index 5a8913a0d8811..46565079d87dd 100644 --- a/src/category_theory/subobject/basic.lean +++ b/src/category_theory/subobject/basic.lean @@ -149,7 +149,7 @@ The morphism in `C` from the arbitrarily chosen underlying object to the ambient -/ noncomputable def arrow {X : C} (Y : subobject X) : (Y : C) ⟶ X := -(representative.obj Y).val.hom +(representative.obj Y).obj.hom instance arrow_mono {X : C} (Y : subobject X) : mono (Y.arrow) := (representative.obj Y).property diff --git a/src/category_theory/subobject/mono_over.lean b/src/category_theory/subobject/mono_over.lean index 0325ade408c04..98ccbe1a4b4e1 100644 --- a/src/category_theory/subobject/mono_over.lean +++ b/src/category_theory/subobject/mono_over.lean @@ -49,19 +49,19 @@ This isn't skeletal, so it's not a partial order. Later we define `subobject X` as the quotient of this by isomorphisms. -/ @[derive [category]] -def mono_over (X : C) := {f : over X // mono f.hom} +def mono_over (X : C) := full_subcategory (λ (f : over X), mono f.hom) namespace mono_over /-- Construct a `mono_over X`. -/ @[simps] -def mk' {X A : C} (f : A ⟶ X) [hf : mono f] : mono_over X := { val := over.mk f, property := hf } +def mk' {X A : C} (f : A ⟶ X) [hf : mono f] : mono_over X := { obj := over.mk f, property := hf } /-- The inclusion from monomorphisms over X to morphisms over X. -/ def forget (X : C) : mono_over X ⥤ over X := full_subcategory_inclusion _ instance : has_coe (mono_over X) C := -{ coe := λ Y, Y.val.left, } +{ coe := λ Y, Y.obj.left, } @[simp] lemma forget_obj_left {f} : ((forget X).obj f).left = (f : C) := rfl @@ -93,13 +93,13 @@ end⟩ @[reassoc] lemma w {f g : mono_over X} (k : f ⟶ g) : k.left ≫ g.arrow = f.arrow := over.w _ /-- Convenience constructor for a morphism in monomorphisms over `X`. -/ -abbreviation hom_mk {f g : mono_over X} (h : f.val.left ⟶ g.val.left) (w : h ≫ g.arrow = f.arrow) : +abbreviation hom_mk {f g : mono_over X} (h : f.obj.left ⟶ g.obj.left) (w : h ≫ g.arrow = f.arrow) : f ⟶ g := over.hom_mk h w /-- Convenience constructor for an isomorphism in monomorphisms over `X`. -/ @[simps] -def iso_mk {f g : mono_over X} (h : f.val.left ≅ g.val.left) (w : h.hom ≫ g.arrow = f.arrow) : +def iso_mk {f g : mono_over X} (h : f.obj.left ≅ g.obj.left) (w : h.hom ≫ g.arrow = f.arrow) : f ≅ g := { hom := hom_mk h.hom w, inv := hom_mk h.inv (by rw [h.inv_comp_eq, w]) } @@ -217,7 +217,7 @@ def map_id : map (𝟙 X) ≅ 𝟭 _ := lift_iso _ _ over.map_id ≪≫ lift_id @[simp] lemma map_obj_left (f : X ⟶ Y) [mono f] (g : mono_over X) : - (((map f).obj g) : C) = g.val.left := + (((map f).obj g) : C) = g.obj.left := rfl @[simp] @@ -324,7 +324,7 @@ adjunction.mk_of_hom_equiv inv_fun := λ k, begin refine over.hom_mk _ _, - refine image.lift {I := g.val.left, m := g.arrow, e := k.left, fac' := over.w k}, + refine image.lift {I := g.obj.left, m := g.arrow, e := k.left, fac' := over.w k}, apply image.lift_fac, end, left_inv := λ k, subsingleton.elim _ _, diff --git a/src/category_theory/subterminal.lean b/src/category_theory/subterminal.lean index ddf8015a321e3..5adb0eedb5aae 100644 --- a/src/category_theory/subterminal.lean +++ b/src/category_theory/subterminal.lean @@ -117,7 +117,7 @@ to the lattice of open subsets of `X`. More generally, if `C` is a topos, this i -/ @[derive category] def subterminals (C : Type u₁) [category.{v₁} C] := -{A : C // is_subterminal A} +full_subcategory (λ (A : C), is_subterminal A) instance [has_terminal C] : inhabited (subterminals C) := ⟨⟨⊤_ C, is_subterminal_of_terminal⟩⟩ @@ -140,7 +140,7 @@ def subterminals_equiv_mono_over_terminal [has_terminal C] : { obj := λ X, ⟨over.mk (terminal.from X.1), X.2.mono_terminal_from⟩, map := λ X Y f, mono_over.hom_mk f (by ext1 ⟨⟨⟩⟩) }, inverse := - { obj := λ X, ⟨X.val.left, λ Z f g, by { rw ← cancel_mono X.arrow, apply subsingleton.elim }⟩, + { obj := λ X, ⟨X.obj.left, λ Z f g, by { rw ← cancel_mono X.arrow, apply subsingleton.elim }⟩, map := λ X Y f, f.1 }, unit_iso := { hom := { app := λ X, 𝟙 _ }, diff --git a/src/representation_theory/fdRep.lean b/src/representation_theory/fdRep.lean index bdc8703f7707c..054983741e6a1 100644 --- a/src/representation_theory/fdRep.lean +++ b/src/representation_theory/fdRep.lean @@ -43,13 +43,13 @@ variables {k G : Type u} [field k] [monoid G] instance : has_coe_to_sort (fdRep k G) (Type u) := concrete_category.has_coe_to_sort _ instance (V : fdRep k G) : add_comm_group V := -by { change add_comm_group ((forget₂ (fdRep k G) (FinVect k)).obj V), apply_instance, } +by { change add_comm_group ((forget₂ (fdRep k G) (FinVect k)).obj V).obj, apply_instance, } instance (V : fdRep k G) : module k V := -by { change module k ((forget₂ (fdRep k G) (FinVect k)).obj V), apply_instance, } +by { change module k ((forget₂ (fdRep k G) (FinVect k)).obj V).obj, apply_instance, } instance (V : fdRep k G) : finite_dimensional k V := -by { change finite_dimensional k ((forget₂ (fdRep k G) (FinVect k)).obj V), apply_instance, } +by { change finite_dimensional k ((forget₂ (fdRep k G) (FinVect k)).obj V).obj, apply_instance, } /-- The monoid homomorphism corresponding to the action of `G` onto `V : fdRep k G`. -/ def ρ (V : fdRep k G) : G →* (V →ₗ[k] V) := V.ρ diff --git a/src/topology/category/Top/open_nhds.lean b/src/topology/category/Top/open_nhds.lean index 5fad2e3daddcc..19d3a759f5b6c 100644 --- a/src/topology/category/Top/open_nhds.lean +++ b/src/topology/category/Top/open_nhds.lean @@ -35,7 +35,7 @@ variables {X Y : Top.{u}} (f : X ⟶ Y) namespace topological_space /-- The type of open neighbourhoods of a point `x` in a (bundled) topological space. -/ -def open_nhds (x : X) := { U : opens X // x ∈ U } +def open_nhds (x : X) := full_subcategory (λ (U : opens X), x ∈ U) namespace open_nhds @@ -43,7 +43,7 @@ instance (x : X) : partial_order (open_nhds x) := { le := λ U V, U.1 ≤ V.1, le_refl := λ _, le_rfl, le_trans := λ _ _ _, le_trans, - le_antisymm := λ _ _ i j, subtype.eq $ le_antisymm i j } + le_antisymm := λ _ _ i j, full_subcategory.ext _ _ $ le_antisymm i j } instance (x : X) : lattice (open_nhds x) := { inf := λ U V, ⟨U.1 ⊓ V.1, ⟨U.2, V.2⟩⟩, @@ -92,7 +92,7 @@ lemma open_embedding {x : X} (U : open_nhds x) : open_embedding (U.1.inclusion) U.1.open_embedding def map (x : X) : open_nhds (f x) ⥤ open_nhds x := -{ obj := λ U, ⟨(opens.map f).obj U.1, by tidy⟩, +{ obj := λ U, ⟨(opens.map f).obj U.1, U.2⟩, map := λ U V i, (opens.map f).map i } @[simp] lemma map_obj (x : X) (U) (q) : (map f x).obj ⟨U, q⟩ = ⟨(opens.map f).obj U, by tidy⟩ := diff --git a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean index 7af25b4afdc98..081f85a57a0bd 100644 --- a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean +++ b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean @@ -47,13 +47,12 @@ namespace sheaf_condition /-- The category of open sets contained in some element of the cover. -/ -def opens_le_cover : Type v := { V : opens X // ∃ i, V ≤ U i } +@[derive category] +def opens_le_cover : Type v := full_subcategory (λ (V : opens X), ∃ i, V ≤ U i) instance [inhabited ι] : inhabited (opens_le_cover U) := ⟨⟨⊥, default, bot_le⟩⟩ -instance : category (opens_le_cover U) := category_theory.full_subcategory _ - namespace opens_le_cover variables {U} @@ -66,7 +65,7 @@ def index (V : opens_le_cover U) : ι := V.property.some /-- The morphism from `V` to `U i` for some `i`. -/ -def hom_to_index (V : opens_le_cover U) : V.val ⟶ U (index V) := +def hom_to_index (V : opens_le_cover U) : V.obj ⟶ U (index V) := (V.property.some_spec).hom end opens_le_cover @@ -244,7 +243,7 @@ variables {Y : opens X} (hY : Y = supr U) in the sieve. This full subcategory is equivalent to `opens_le_cover U`, the (poset) category of opens contained in some `U i`. -/ @[simps] def generate_equivalence_opens_le : - {f : over Y // (sieve.generate (presieve_of_covering_aux U Y)).arrows f.hom} ≌ + full_subcategory (λ (f : over Y), (sieve.generate (presieve_of_covering_aux U Y)).arrows f.hom) ≌ opens_le_cover U := { functor := { obj := λ f, ⟨f.1.left, let ⟨_,h,_,⟨i,hY⟩,_⟩ := f.2 in ⟨i, hY ▸ h.le⟩⟩, diff --git a/src/topology/sheaves/sheafify.lean b/src/topology/sheaves/sheafify.lean index 46368756c0bce..3c707fbc5aa5e 100644 --- a/src/topology/sheaves/sheafify.lean +++ b/src/topology/sheaves/sheafify.lean @@ -104,9 +104,9 @@ begin dsimp at e', use ⟨W ⊓ (U' ⊓ V'), ⟨mW, mU, mV⟩⟩, refine ⟨_, _, _⟩, - { change W ⊓ (U' ⊓ V') ⟶ U.val, + { change W ⊓ (U' ⊓ V') ⟶ U.obj, exact (opens.inf_le_right _ _) ≫ (opens.inf_le_left _ _) ≫ iU, }, - { change W ⊓ (U' ⊓ V') ⟶ V.val, + { change W ⊓ (U' ⊓ V') ⟶ V.obj, exact (opens.inf_le_right _ _) ≫ (opens.inf_le_right _ _) ≫ iV, }, { intro w, dsimp, diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 5ff3e104a9df6..90f4f1ce8cb0c 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -155,7 +155,7 @@ begin dsimp [stalk_pushforward, stalk_functor], ext1, tactic.op_induction', - cases j, cases j_val, + rcases j with ⟨⟨_, _⟩, _⟩, rw [colimit.ι_map_assoc, colimit.ι_map, colimit.ι_pre, whisker_left_app, whisker_right_app, pushforward.id_hom_app, eq_to_hom_map, eq_to_hom_refl], dsimp, @@ -172,8 +172,7 @@ begin dsimp [stalk_pushforward, stalk_functor], ext U, induction U using opposite.rec, - cases U, - cases U_val, + rcases U with ⟨⟨_, _⟩, _⟩, simp only [colimit.ι_map_assoc, colimit.ι_pre_assoc, whisker_right_app, category.assoc], dsimp, From 63971d1fa708efaa102e15fa327d6e919c892992 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 1 Aug 2022 15:32:48 +0000 Subject: [PATCH 164/173] fix(tactic/rcases): support `rcases x with y` renames (#15773) fixes #15741 --- src/tactic/rcases.lean | 18 ++++++++++++++++-- test/rcases.lean | 13 +++++++++++++ 2 files changed, 29 insertions(+), 2 deletions(-) diff --git a/src/tactic/rcases.lean b/src/tactic/rcases.lean index e21a0367df6b7..19b94534a857f 100644 --- a/src/tactic/rcases.lean +++ b/src/tactic/rcases.lean @@ -385,7 +385,14 @@ meta def rcases (h : option name) (p : pexpr) (pat : rcases_patt) : tactic unit | none := i_to_expr p end, if e.is_local_constant then - focus1 (rcases_core pat e >>= clear_goals) + match pat.name with + | some x := do + n ← revert e, + e ← intro x, + intron (n - 1), + focus1 (rcases_core pat e >>= clear_goals) + | none := focus1 (rcases_core pat e >>= clear_goals) + end else do x ← pat.name.elim mk_fresh_name pure, n ← revert_kdependencies e semireducible, @@ -409,7 +416,14 @@ meta def rcases_many (ps : listΠ pexpr) (pat : rcases_patt) : tactic unit := do end, e ← i_to_expr p, if e.is_local_constant then - pure (pat, e) + match pat.name with + | some x := do + n ← revert e, + e ← intro x, + intron (n - 1), + pure (pat, e) + | none := pure (pat, e) + end else do x ← pat.name.elim mk_fresh_name pure, n ← revert_kdependencies e semireducible, diff --git a/test/rcases.lean b/test/rcases.lean index 23ca4c6ee27e3..e295c99c60943 100644 --- a/test/rcases.lean +++ b/test/rcases.lean @@ -189,6 +189,19 @@ example : bool → false → true | ff := by rintro ⟨⟩ | tt := by rintro ⟨⟩ +example : true := +begin + obtain h : true, + { trivial }, + exact h +end + +example {a b} (h : a ∧ b) : a ∧ b := +begin + rcases h with t, + exact t +end + open tactic meta def test_rcases_hint (s : string) (num_goals : ℕ) (depth := 5) : tactic unit := do change `(true), From 1f1a93ed8b31cb10e0b1aaf9c8f47419a0ab48f7 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 1 Aug 2022 15:32:49 +0000 Subject: [PATCH 165/173] fix(algebra/monoid_algebra/basic): add `int_cast` to `monoid_algebra` instances (#15779) Note that this reshuffles how some instances are created to ensure that they get the `int_cast` field. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/monoid_algebra/basic.lean | 23 +++++++++++++++++++---- src/data/polynomial/expand.lean | 2 +- test/instance_diamonds.lean | 5 ++++- 3 files changed, 24 insertions(+), 6 deletions(-) diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index 0ad604262a1a7..09d7c6b790759 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -183,6 +183,8 @@ instance : non_assoc_semiring (monoid_algebra k G) := single_zero, sum_zero, add_zero, mul_one, sum_single], ..monoid_algebra.non_unital_non_assoc_semiring } +lemma nat_cast_def (n : ℕ) : (n : monoid_algebra k G) = single 1 n := rfl + end mul_one_class /-! #### Semiring structure -/ @@ -244,11 +246,16 @@ instance [ring k] [semigroup G] : non_unital_ring (monoid_algebra k G) := .. monoid_algebra.non_unital_semiring } instance [ring k] [mul_one_class G] : non_assoc_ring (monoid_algebra k G) := -{ .. monoid_algebra.add_comm_group, +{ int_cast := λ z, single 1 (z : k), + int_cast_of_nat := λ n, by simpa, + int_cast_neg_succ_of_nat := λ n, by simpa, + .. monoid_algebra.add_comm_group, .. monoid_algebra.non_assoc_semiring } +lemma int_cast_def [ring k] [mul_one_class G] (z : ℤ) : (z : monoid_algebra k G) = single 1 z := rfl + instance [ring k] [monoid G] : ring (monoid_algebra k G) := -{ .. monoid_algebra.non_unital_non_assoc_ring, +{ .. monoid_algebra.non_assoc_ring, .. monoid_algebra.semiring } instance [comm_ring k] [comm_semigroup G] : non_unital_comm_ring (monoid_algebra k G) := @@ -1029,6 +1036,8 @@ instance : non_assoc_semiring (add_monoid_algebra k G) := single_zero, sum_zero, add_zero, mul_one, sum_single], .. add_monoid_algebra.non_unital_non_assoc_semiring } +lemma nat_cast_def (n : ℕ) : (n : add_monoid_algebra k G) = single 0 n := rfl + end mul_one_class /-! #### Semiring structure -/ @@ -1091,11 +1100,17 @@ instance [ring k] [add_semigroup G] : non_unital_ring (add_monoid_algebra k G) : .. add_monoid_algebra.non_unital_semiring } instance [ring k] [add_zero_class G] : non_assoc_ring (add_monoid_algebra k G) := -{ .. add_monoid_algebra.add_comm_group, +{ int_cast := λ z, single 0 (z : k), + int_cast_of_nat := λ n, by simpa, + int_cast_neg_succ_of_nat := λ n, by simpa, + .. add_monoid_algebra.add_comm_group, .. add_monoid_algebra.non_assoc_semiring } +lemma int_cast_def [ring k] [add_zero_class G] (z : ℤ) : + (z : add_monoid_algebra k G) = single 0 z := rfl + instance [ring k] [add_monoid G] : ring (add_monoid_algebra k G) := -{ .. add_monoid_algebra.non_unital_non_assoc_ring, +{ .. add_monoid_algebra.non_assoc_ring, .. add_monoid_algebra.semiring } instance [comm_ring k] [add_comm_semigroup G] : non_unital_comm_ring (add_monoid_algebra k G) := diff --git a/src/data/polynomial/expand.lean b/src/data/polynomial/expand.lean index c5d4f12bd692d..22d0a2234866d 100644 --- a/src/data/polynomial/expand.lean +++ b/src/data/polynomial/expand.lean @@ -240,7 +240,7 @@ variable {R} theorem of_irreducible_expand {p : ℕ} (hp : p ≠ 0) {f : R[X]} (hf : irreducible (expand R p f)) : irreducible f := -@@of_irreducible_map _ _ _ (is_local_ring_hom_expand R hp.bot_lt) hf +let _ := is_local_ring_hom_expand R hp.bot_lt in by exactI of_irreducible_map ↑(expand R p) hf theorem of_irreducible_expand_pow {p : ℕ} (hp : p ≠ 0) {f : R[X]} {n : ℕ} : irreducible (expand R (p ^ n) f) → irreducible f := diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index 9755f42425164..1ba113934c6d8 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -171,7 +171,10 @@ example [comm_semiring R] [nontrivial R] : rfl /-- `polynomial.algebra_of_algebra` is consistent with `algebra_nat`. -/ -example [comm_semiring R] : (polynomial.algebra_of_algebra : algebra ℕ R[X]) = algebra_nat := rfl +example [semiring R] : (polynomial.algebra_of_algebra : algebra ℕ R[X]) = algebra_nat := rfl + +/-- `polynomial.algebra_of_algebra` is consistent with `algebra_int`. -/ +example [ring R] : (polynomial.algebra_of_algebra : algebra ℤ R[X]) = algebra_int _ := rfl end polynomial From cd8620de95f485de61db28e37fb0aab02f7247d7 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Mon, 1 Aug 2022 20:44:13 -0400 Subject: [PATCH 166/173] implement proper power computation --- scripts/polyrith_sage.py | 22 +++++++++++++++++++++- src/tactic/polyrith.lean | 21 ++++++++++++++------- test/polyrith.lean | 38 +++++++++++++++++++++++++++++++++++--- 3 files changed, 70 insertions(+), 11 deletions(-) diff --git a/scripts/polyrith_sage.py b/scripts/polyrith_sage.py index 23f5d0d4a09d5..993a41b65f347 100644 --- a/scripts/polyrith_sage.py +++ b/scripts/polyrith_sage.py @@ -32,6 +32,26 @@ def create_query(type: str, n_vars: int, eq_list, goal_type): ''' return query +def create_query_radical(type: str, n_vars: int, eq_list, goal_type): + """ Create a query to invoke Sage's `MPolynomial_libsingular.lift`. See + https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518 + for a description of this method. """ + var_list = [f"var{i}" for i in range(n_vars)] + ['aux'] + query = f''' +import json +P = {type_str(type)}{var_list} +[{", ".join(var_list)}] = P.gens() +p = P({goal_type}) +gens = {eq_list} + [1 - p*aux] +I = ideal(gens) +coeffs = P(1).lift(I) +power = max(cf.degree(aux) for cf in coeffs) +coeffs = [P(cf.subs(aux = 1/p)*p^power) for cf in coeffs[:int(-1)]] +js = {{'power': power, 'coeffs': [polynomial_to_string(c) for c in coeffs]}} +print(json.dumps(js)) +''' + return query + class EvaluationError(Exception): def __init__(self, ename, evalue, message='Error in Sage communication'): self.ename = ename @@ -68,7 +88,7 @@ def main(): error_value: Optional[str] } ``` ''' - command = create_query(sys.argv[2], int(sys.argv[3]), sys.argv[4], sys.argv[5]) + command = create_query_radical(sys.argv[2], int(sys.argv[3]), sys.argv[4], sys.argv[5]) final_query = polynomial_formatting_functions + "\n" + command if sys.argv[1] == 'tt': # trace dry run enabled output = dict(success=True, trace=command) diff --git a/src/tactic/polyrith.lean b/src/tactic/polyrith.lean index 17c6ad9872e9c..7f6354e9fa682 100644 --- a/src/tactic/polyrith.lean +++ b/src/tactic/polyrith.lean @@ -300,12 +300,18 @@ meta instance : non_null_json_serializable poly := | sum.inr p := pure p end} +/-- A schema for the data reported by the Sage calculation -/ +@[derive [non_null_json_serializable, inhabited]] +structure sage_json_coeffs_and_power := +(coeffs : list poly) +(power : ℕ) + /-- A schema for success messages from the python script -/ @[derive [non_null_json_serializable, inhabited]] structure sage_json_success := (success : {b : bool // b = tt}) (trace : option string := none) -(data : option (list poly) := none) +(data : option sage_json_coeffs_and_power := none) /-- A schema for failure messages from the python script -/ @[derive [non_null_json_serializable, inhabited]] @@ -316,7 +322,7 @@ structure sage_json_failure := /-- Parse the json output from `scripts/polyrith.py` into either an error message, a list of `poly` objects, or `none` if only trace output was requested. -/ -meta def convert_sage_output (j : json) : tactic (option (list poly)) := +meta def convert_sage_output (j : json) : tactic (option (ℕ × list poly)) := do r : sage_json_success ⊕ sage_json_failure ← decorate_ex "internal json error: " -- try the error format first, so that if both fail we get the message from the success parser @@ -326,7 +332,7 @@ do fail!"polyrith failed to retrieve a solution from Sage! {f.error_name}: {f.error_value}" | sum.inl s := do s.trace.mmap trace, - pure s.data + pure $ s.data.map $ λ sd, (sd.power, sd.coeffs) end /-! @@ -484,18 +490,19 @@ a call to `linear_combination`. -/ meta def process_output (eq_names : list expr) (m : list expr) (R : expr) (sage_out : json) : tactic format := focus1 $ do - some coeffs_as_poly ← convert_sage_output sage_out | fail!"internal error: No output available", + some (power, coeffs_as_poly) ← convert_sage_output sage_out | fail!"internal error: No output available", coeffs_as_pexpr ← coeffs_as_poly.mmap (poly.to_pexpr m), let eq_names_pexpr := eq_names.map to_pexpr, coeffs_as_expr ← coeffs_as_pexpr.mmap $ λ e, to_expr ``(%%e : %%R), - linear_combo.linear_combination eq_names_pexpr coeffs_as_pexpr, + linear_combo.linear_combination eq_names_pexpr coeffs_as_pexpr (some power), let components := (eq_names.zip coeffs_as_expr).filter $ λ pr, bnot $ pr.2.is_app_of `has_zero.zero, expr_string ← components_to_lc_format components, - let lc_fmt : format := "linear_combination " ++ format.nest 2 (format.group expr_string), + let lc_exp : format := if power = 1 then "" else format!" with exponent {power}", + let lc_fmt : format := "linear_combination " ++ format.nest 2 (format.group expr_string ++ lc_exp), done <|> fail!"polyrith found the following certificate, but it failed to close the goal:\n{lc_fmt}", - return $ "linear_combination " ++ format.nest 2 (format.group expr_string) + return lc_fmt /-- Tactic for the special case when no hypotheses are available. -/ meta def no_hypotheses_case : tactic (option format) := diff --git a/test/polyrith.lean b/test/polyrith.lean index c3b71514e1e42..2ba08073c5130 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -98,10 +98,13 @@ end tactic ## SageCell communcation tests -/ +-- set_option trace.polyrith true + example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : x*y = -2*y + 1 := begin - test_sage_output "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", + polyrith, + -- test_sage_output "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", linear_combination h1 - 2 * h2 end @@ -436,10 +439,39 @@ by test_polyrith "linear_combination h1 - h2" --- We comment the following tests so that we don't overwhelm the SageCell API. - +/- ### Cases with powers -/ +example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := +by test_polyrith + "{\"data\":{\"coeffs\":[\"(poly.add (poly.neg (poly.mul (poly.var 1) (poly.pow (poly.var 2) 2))) (poly.var 0))\",\"(poly.add (poly.add (poly.pow (poly.var 2) 2) (poly.mul (poly.const 2/1) (poly.var 2))) (poly.const 1/1))\"],\"power\":2},\"success\":true}" + ["ff", + "rat", + "3", + "[(var0 - var1), ((var0 * var1) - 0)]", + "((var0 + (var1 * var2)) - 0)"] + "linear_combination (-(y * z ^ 2) + x) * h + (z ^ 2 + 2 * z + 1) * h2 with exponent 2" + +. +example (K : Type) (h_this : 3 - 1 = 2) + [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 test_polyrith + "{\"data\":{\"coeffs\":[\"(poly.mul (poly.mul (poly.const 2/7) (poly.var 1)) (poly.pow (poly.var 2) 2))\",\"(poly.sub (poly.pow (poly.var 0) 3) (poly.mul (poly.mul (poly.const 1/7) (poly.pow (poly.var 1) 2)) (poly.var 2)))\",\"(poly.neg (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.var 2)))\",\"(poly.mul (poly.mul (poly.const 1/7) (poly.var 1)) (poly.var 2))\"],\"power\":6},\"success\":true}" + ["ff", + "K", + "3", + "[(((var1 ^ 3) + (var0 * (3 * (var2 ^ 2)))) - 0), (((var0 ^ 3) + (var2 * (3 * (var1 ^ 2)))) - 0), (((var1 * (3 * (var0 ^ 2))) + (var2 ^ 3)) - 0), (((((var0 ^ 3) * var1) + ((var1 ^ 3) * var2)) + ((var2 ^ 3) * var0)) - 0)]", + "(var0 - 0)"] + "linear_combination 2 / 7 * y * z ^ 2 * h₂ + (x ^ 3 - 1 / 7 * y ^ 2 * z) * h₁ - x * y * z * h₀ + 1 / 7 * y * + z * h with exponent 6" +-- We comment the following tests so that we don't overwhelm the SageCell API. /- From f5546b01d878de5e6856d7215f27ae81d47df2d0 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Mon, 1 Aug 2022 20:48:11 -0400 Subject: [PATCH 167/173] dead code --- scripts/polyrith_sage.py | 19 +------------------ 1 file changed, 1 insertion(+), 18 deletions(-) diff --git a/scripts/polyrith_sage.py b/scripts/polyrith_sage.py index 993a41b65f347..96eddcbaa6aad 100644 --- a/scripts/polyrith_sage.py +++ b/scripts/polyrith_sage.py @@ -16,23 +16,6 @@ def type_str(type): return "QQ" def create_query(type: str, n_vars: int, eq_list, goal_type): - """ Create a query to invoke Sage's `MPolynomial_libsingular.lift`. See - https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518 - for a description of this method. """ - var_list = ", ".join([f"var{i}" for i in range(n_vars)]) - query = f''' -import json -P = PolynomialRing({type_str(type)}, 'var', {n_vars!r}) -[{var_list}] = P.gens() -gens = {eq_list} -p = P({goal_type}) -I = ideal(gens) -coeffs = p.lift(I) -print(json.dumps([polynomial_to_string(c) for c in coeffs])) -''' - return query - -def create_query_radical(type: str, n_vars: int, eq_list, goal_type): """ Create a query to invoke Sage's `MPolynomial_libsingular.lift`. See https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518 for a description of this method. """ @@ -88,7 +71,7 @@ def main(): error_value: Optional[str] } ``` ''' - command = create_query_radical(sys.argv[2], int(sys.argv[3]), sys.argv[4], sys.argv[5]) + command = create_query(sys.argv[2], int(sys.argv[3]), sys.argv[4], sys.argv[5]) final_query = polynomial_formatting_functions + "\n" + command if sys.argv[1] == 'tt': # trace dry run enabled output = dict(success=True, trace=command) From 41da1687a460a7be25a082711a2717142eedee72 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Mon, 1 Aug 2022 20:50:11 -0400 Subject: [PATCH 168/173] clean up tests --- test/polyrith.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/test/polyrith.lean b/test/polyrith.lean index 2ba08073c5130..abac683afc414 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -98,13 +98,10 @@ end tactic ## SageCell communcation tests -/ --- set_option trace.polyrith true - example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : x*y = -2*y + 1 := begin - polyrith, - -- test_sage_output "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", + test_sage_output "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", linear_combination h1 - 2 * h2 end From 6f18ae9ff10c0c321a0530ed54d2632125b6bdd5 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 22 Sep 2022 19:59:01 -0400 Subject: [PATCH 169/173] update syntax to use config object --- src/tactic/linear_combination.lean | 13 ++++--------- test/linear_combination.lean | 6 +++--- 2 files changed, 7 insertions(+), 12 deletions(-) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index 41435923f9f52..17bdd9338eb11 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 -/ @@ -320,14 +321,14 @@ Note: The left and right sides of all the equalities should have the same * Output: N/A -/ meta def linear_combination (h_eqs_names : list pexpr) (coeffs : list pexpr) - (exponent : option ℕ := none) (config : linear_combination_config := {}) : tactic unit := + (config : linear_combination_config := {}) : tactic unit := do `(@eq %%ext _ _) ← target | fail "linear_combination can only be used to prove equality goals", h_eqs ← h_eqs_names.mmap to_expr, 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 (exponent.get_or_else 1), + raise_goal_to_power config.exponent, set_goal_to_hleft_sub_tleft hsum_on_left, normalize_if_desired config @@ -435,16 +436,10 @@ by linear_combination 3 * h a b + hqc meta def _root_.tactic.interactive.linear_combination (input : parse (as_linear_combo ff [] <$> texpr)?) (_ : parse (tk "with")?) - (exponent : parse (prod.mk <$> ident <*> small_nat)?) (config : linear_combination_config := {}) : tactic unit := do - exponent ← match exponent with - | none := return none - | some (`exponent, n) := return $ some n - | some (id, _) := fail!"linear_combination does not support the modifier {id}" - end, let (h_eqs_names, coeffs) := list.unzip (input.get_or_else []) in -linear_combination h_eqs_names coeffs exponent config +linear_combination h_eqs_names coeffs config add_tactic_doc { name := "linear_combination", diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 8f5355aea530f..c1f0921b46dee 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -202,12 +202,12 @@ 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 +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}, + with {exponent := 2, normalize := ff}, ring end @@ -221,7 +221,7 @@ example (K : Type) (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 + x * y * z * h₀ + y * z * h / 7 with {exponent := 6} /-! ### Cases where the goal is not closed -/ From 101bcbdc3a84e9fed0aec2e0fa861398348e6d22 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 22 Sep 2022 20:01:07 -0400 Subject: [PATCH 170/173] golf --- src/tactic/linear_combination.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index 17bdd9338eb11..9675d584fc630 100644 --- a/src/tactic/linear_combination.lean +++ b/src/tactic/linear_combination.lean @@ -274,11 +274,9 @@ If an exponent `n` is provided, changes the goal from `t = 0` to `t^n = 0`. * Output: N/A -/ -meta def raise_goal_to_power (exponent : ℕ) : tactic unit := -match exponent with +meta def raise_goal_to_power : ℕ → tactic unit | 1 := skip | n := refine ``(@pow_eq_zero _ _ _ _ %%`(n) _) -end /-- This tactic attempts to prove the goal by normalizing the target if the From 72ef223abb7451b9fc2729cbde58685e8f759c54 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 22 Sep 2022 20:12:37 -0400 Subject: [PATCH 171/173] update syntax to provide exponent in config object --- src/tactic/polyrith.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/tactic/polyrith.lean b/src/tactic/polyrith.lean index 7f6354e9fa682..875f970c49d49 100644 --- a/src/tactic/polyrith.lean +++ b/src/tactic/polyrith.lean @@ -494,11 +494,11 @@ meta def process_output (eq_names : list expr) (m : list expr) (R : expr) (sage_ coeffs_as_pexpr ← coeffs_as_poly.mmap (poly.to_pexpr m), let eq_names_pexpr := eq_names.map to_pexpr, coeffs_as_expr ← coeffs_as_pexpr.mmap $ λ e, to_expr ``(%%e : %%R), - linear_combo.linear_combination eq_names_pexpr coeffs_as_pexpr (some power), + linear_combo.linear_combination eq_names_pexpr coeffs_as_pexpr {exponent := power}, let components := (eq_names.zip coeffs_as_expr).filter $ λ pr, bnot $ pr.2.is_app_of `has_zero.zero, expr_string ← components_to_lc_format components, - let lc_exp : format := if power = 1 then "" else format!" with exponent {power}", + let lc_exp : format := if power = 1 then "" else format!" with {{exponent := {power}}}", let lc_fmt : format := "linear_combination " ++ format.nest 2 (format.group expr_string ++ lc_exp), done <|> fail!"polyrith found the following certificate, but it failed to close the goal:\n{lc_fmt}", From d00de7412ca66f352350e7b067649460a471259f Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 22 Sep 2022 20:20:13 -0400 Subject: [PATCH 172/173] update tests --- test/polyrith.lean | 98 ++++++++++++++++++++++++++++------------------ 1 file changed, 59 insertions(+), 39 deletions(-) diff --git a/test/polyrith.lean b/test/polyrith.lean index abac683afc414..572e7b47e167f 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -101,7 +101,7 @@ end tactic example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : x*y = -2*y + 1 := begin - test_sage_output "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", + test_sage_output "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"power\":1},\"success\":true}", linear_combination h1 - 2 * h2 end @@ -109,7 +109,7 @@ example (w x y z : ℝ) (h1 : x + 2.1*y + 2*z = 2) (h2 : x + 8*z + 5*w = -6.5) (h3 : x + y + 5*z + 5*w = 3) : x + 2.2*y + 2*z - 5*w = -8.5 := begin - test_sage_output "{\"data\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}", + test_sage_output "{\"data\":{\"coeffs\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"power\":1},\"success\":true}", linear_combination 2 * h1 + h2 - 2 * h3 end @@ -120,7 +120,7 @@ end example (x y : ℤ) (h1 : 3*x + 2*y = 10): 3*x + 2*y = 10 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "int", "2", @@ -131,7 +131,7 @@ by test_polyrith example (x y : ℚ) (h1 : x*y + 2*x = 1) (h2 : x = y) : x*y = -2*y + 1 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "2", @@ -142,7 +142,7 @@ by test_polyrith example (x y : ℝ) (h1 : x + 2 = -3) (h2 : y = 10) : -y + 2*x + 4 = -16 := by test_polyrith - "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\"],\"power\":1},\"success\":true}" ["ff", "real", "2", @@ -154,7 +154,7 @@ example (x y z : ℝ) (ha : x + 2*y - z = 4) (hb : 2*x + y + z = -2) (hc : x + 2*y + z = 2) : -3*x - 3*y - 4*z = 2 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\",\"(poly.const -2/1)\"],\"power\":1},\"success\":true}" ["ff", "real", "3", @@ -166,7 +166,7 @@ example (w x y z : ℝ) (h1 : x + 2.1*y + 2*z = 2) (h2 : x + 8*z + 5*w = -6.5) (h3 : x + y + 5*z + 5*w = 3) : x + 2.2*y + 2*z - 5*w = -8.5 := by test_polyrith - "{\"data\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 2/1)\",\"(poly.const 1/1)\",\"(poly.const -2/1)\"],\"power\":1},\"success\":true}" ["ff", "real", "4", @@ -177,7 +177,7 @@ by test_polyrith example (a b c d : ℚ) (h1 : a = 4) (h2 : 3 = b) (h3 : c*3 = d) (h4 : -d = a) : 2*a - 3 + 9*c + 3*d = 8 - b + 3*d - 3*a := by test_polyrith - "{\"data\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\",\"(poly.const 3/1)\",\"(poly.const -3/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 2/1)\",\"(poly.const -1/1)\",\"(poly.const 3/1)\",\"(poly.const -3/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "4", @@ -190,7 +190,7 @@ by test_polyrith example («def evil» y : ℤ) (h1 : 3*«def evil» + 2*y = 10): 3*«def evil» + 2*y = 10 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "int", "2", @@ -201,7 +201,7 @@ by test_polyrith example («¥» y : ℤ) (h1 : 3*«¥» + 2*y = 10): «¥» * (3*«¥» + 2*y) = 10 * «¥» := by test_polyrith - "{\"data\":[\"(poly.var 0)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.var 0)\"],\"power\":1},\"success\":true}" ["ff", "int", "2", @@ -214,7 +214,7 @@ by test_polyrith example (a b : ℤ) (h : a = b) : a * a = a * b := by test_polyrith - "{\"data\":[\"(poly.var 0)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.var 0)\"],\"power\":1},\"success\":true}" ["ff", "int", "2", @@ -225,7 +225,7 @@ by test_polyrith example (a b c : ℤ) (h : a = b) : a * c = b * c := by test_polyrith - "{\"data\":[\"(poly.var 1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.var 1)\"],\"power\":1},\"success\":true}" ["ff", "int", "3", @@ -236,7 +236,7 @@ by test_polyrith example (a b c : ℤ) (h1 : a = b) (h2 : b = 1) : c * a + b = c * b + 1 := by test_polyrith - "{\"data\":[\"(poly.var 0)\",\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.var 0)\",\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "int", "3", @@ -247,7 +247,7 @@ by test_polyrith example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : x*x*y + y*x*y + 6*x = 3*x*y + 14 := by test_polyrith - "{\"data\":[\"(poly.mul (poly.var 0) (poly.var 1))\",\"(poly.const 2/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.mul (poly.var 0) (poly.var 1))\",\"(poly.const 2/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "2", @@ -257,7 +257,7 @@ by test_polyrith example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := by test_polyrith - "{\"data\":[\"(poly.add (poly.var 0) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.add (poly.var 0) (poly.mul (poly.const 2/1) (poly.var 2)))\"],\"power\":1},\"success\":true}" ["ff", "rat", "4", @@ -270,7 +270,7 @@ by test_polyrith example (a b : ℝ) (ha : 2*a = 4) (hab : 2*b = a - b) (hignore : 3 = a + b) : b = 2 / 3 := by test_polyrith only [ha, hab] - "{\"data\":[\"(poly.const 1/6)\",\"(poly.const 1/3)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/6)\",\"(poly.const 1/3)\"],\"power\":1},\"success\":true}" ["ff", "real", "2", @@ -282,7 +282,7 @@ constant term : ∀ a b : ℚ, a + b = 0 example (a b c d : ℚ) (h : a + b = 0) (h2: b + c = 0): a + b + c + d = 0 := by test_polyrith only [term c d, h] - "{\"data\":[\"(poly.const 1/1)\",\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\",\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "4", @@ -294,7 +294,7 @@ constants (qc : ℚ) (hqc : qc = 2*qc) example (a b : ℚ) (h : ∀ p q : ℚ, p = q) : 3*a + qc = 3*b + 2*qc := by test_polyrith [h a b, hqc] - "{\"data\":[\"(poly.const 3/1)\",\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 3/1)\",\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "3", @@ -306,7 +306,7 @@ constant bad (q : ℚ) : q = 0 example (a b : ℚ) : a + b^3 = 0 := by test_polyrith [bad a, bad (b^2)] - "{\"data\":[\"(poly.const 1/1)\",\"(poly.var 1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\",\"(poly.var 1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "2", @@ -319,7 +319,7 @@ by test_polyrith [bad a, bad (b^2)] example {α} [h : comm_ring α] {a b c d e f : α} (h1 : a*d = b*c) (h2 : c*f = e*d) : c * (a*f - b*e) = 0 := by test_polyrith - "{\"data\":[\"(poly.var 4)\",\"(poly.var 1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.var 4)\",\"(poly.var 1)\"],\"power\":1},\"success\":true}" ["ff", "α", "6", @@ -336,7 +336,7 @@ begin have hs_nonzero : s ≠ 0, { contrapose! hp_nonzero with hs_nonzero, test_polyrith - "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.var 4)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.const -1/1)\",\"(poly.const 0/1)\",\"(poly.var 4)\"],\"power\":1},\"success\":true}" ["ff", "K", "6", @@ -346,7 +346,7 @@ begin have H' : 2 * q = s ^ 3 - t ^ 3, { rw ← mul_left_inj' (pow_ne_zero 3 hs_nonzero), test_polyrith - "{\"data\":[\"(poly.const -1/1)\",\"(poly.sub (poly.add (poly.neg (poly.pow (poly.var 1) 3)) (poly.var 0)) (poly.var 3))\",\"(poly.add (poly.add (poly.mul (poly.pow (poly.var 1) 2) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.var 1) (poly.var 2)) (poly.var 4))) (poly.pow (poly.var 4) 2))\",\"(poly.const 0/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const -1/1)\",\"(poly.sub (poly.add (poly.neg (poly.pow (poly.var 1) 3)) (poly.var 0)) (poly.var 3))\",\"(poly.add (poly.add (poly.mul (poly.pow (poly.var 1) 2) (poly.pow (poly.var 2) 2)) (poly.mul (poly.mul (poly.var 1) (poly.var 2)) (poly.var 4))) (poly.pow (poly.var 4) 2))\",\"(poly.const 0/1)\"],\"power\":1},\"success\":true}" ["ff", "K", "6", @@ -354,7 +354,7 @@ begin "(((2 * var0) * (var1 ^ 3)) - (((var1 ^ 3) - (var2 ^ 3)) * (var1 ^ 3)))"] "linear_combination -hr + (-s ^ 3 + q - r) * hs3 + (s ^ 2 * t ^ 2 + s * t * p + p ^ 2) * ht"}, test_polyrith - "{\"data\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.mul (poly.var 0) (poly.pow (poly.var 5) 4)) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 3))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.add (poly.sub (poly.add (poly.sub (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.sub (poly.add (poly.neg (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 3) 2)) (poly.var 5))) (poly.mul (poly.pow (poly.var 3) 3) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.pow (poly.var 4) 3) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 3))) (poly.pow (poly.var 3) 3)) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 4))) (poly.pow (poly.var 4) 3)) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1)))\",\"(poly.const -1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 0/1)\",\"(poly.const 0/1)\",\"(poly.add (poly.add (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.mul (poly.var 0) (poly.pow (poly.var 5) 4)) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 4))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 3))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 3))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 3) (poly.pow (poly.var 5) 2))) (poly.mul (poly.var 4) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 0)) (poly.var 5)))\",\"(poly.sub (poly.add (poly.add (poly.sub (poly.sub (poly.add (poly.add (poly.sub (poly.add (poly.sub (poly.neg (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 3) 2)) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.pow (poly.var 4) 2)) (poly.var 5))) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 3)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.mul (poly.var 1) (poly.var 4)) (poly.pow (poly.var 5) 2))) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 3))) (poly.mul (poly.pow (poly.var 0) 2) (poly.var 4))) (poly.mul (poly.mul (poly.var 0) (poly.var 1)) (poly.var 5))) (poly.mul (poly.mul (poly.const 3/1) (poly.var 0)) (poly.var 1))) (poly.mul (poly.mul (poly.const 2/1) (poly.var 2)) (poly.var 5))) (poly.mul (poly.const 2/1) (poly.var 2)))\",\"(poly.neg (poly.pow (poly.var 5) 3))\"],\"power\":1},\"success\":true}" ["ff", "K", "7", @@ -362,14 +362,13 @@ begin "((((var0 ^ 3) + ((3 * var1) * var0)) - (2 * var2)) - (((var0 - (var3 - var4)) * (var0 - ((var3 * var5) - (var4 * (var5 ^ 2))))) * (var0 - ((var3 * (var5 ^ 2)) - (var4 * var5)))))"] "linear_combination (x * ω ^ 4 - s * ω ^ 4 + t * ω ^ 4 - s * ω ^ 3 + t * ω ^ 3 + 3 * x * ω ^ 2 - s * ω ^ 2 + t * ω ^ 2 + - 2 * x * ω) * ht + (-(x * s ^ 2 * ω) + s ^ 3 * ω - x * t ^ 2 * ω - t ^ 3 * ω + x * p * ω ^ 2 - p * s * ω ^ 2 + - p * t * ω ^ 2 + + 2 * x * ω) * ht + (-(x * s ^ 2 * ω) - x * t ^ 2 * ω + x * p * ω ^ 2 - p * s * ω ^ 2 + p * t * ω ^ 2 + x ^ 2 * s - - s ^ 3 - - x ^ 2 * t + - t ^ 3 - - x * p * ω + - 3 * x * p) * H - H'" + x ^ 2 * t - + x * p * ω + + 3 * x * p + + 2 * q * ω - + 2 * q) * H - ω ^ 3 * H'" end @@ -377,7 +376,7 @@ end example {K : Type*} [field K] [char_zero K] {s : K} (hs : 3 * s + 1 = 4) : s = 1 := by test_polyrith - "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/3)\"],\"power\":1},\"success\":true}" ["ff", "K", "1", @@ -387,7 +386,7 @@ by test_polyrith example {x : ℤ} (h1 : x + 4 = 2) : x = -2 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "int", "1", @@ -397,7 +396,7 @@ by test_polyrith example {w : ℚ} (h1 : 3 * w + 1 = 4) : w = 1 := by test_polyrith - "{\"data\":[\"(poly.const 1/3)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/3)\"],\"power\":1},\"success\":true}" ["ff", "rat", "1", @@ -407,7 +406,7 @@ by test_polyrith example {x : ℤ} (h1 : 2 * x + 3 = x) : x = -3 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "int", "1", @@ -417,7 +416,7 @@ by test_polyrith example {c : ℚ} (h1 : 4 * c + 1 = 3 * c - 2) : c = -3 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const 1/1)\"],\"power\":1},\"success\":true}" ["ff", "rat", "1", @@ -427,13 +426,13 @@ by test_polyrith example (z : ℤ) (h1 : z + 1 = 2) (h2 : z + 2 = 2) : (1 : ℤ) = 2 := by test_polyrith - "{\"data\":[\"(poly.const 1/1)\",\"(poly.const -1/1)\"],\"success\":true}" + "{\"data\":{\"coeffs\":[\"(poly.const -1/1)\",\"(poly.const 1/1)\"],\"power\":0},\"success\":true}" ["ff", "int", "1", "[((var0 + 1) - 2), ((var0 + 2) - 2)]", "(1 - 2)"] - "linear_combination h1 - h2" + "linear_combination -h1 + h2 with {exponent := 0}" /- ### Cases with powers -/ @@ -446,7 +445,7 @@ by test_polyrith "3", "[(var0 - var1), ((var0 * var1) - 0)]", "((var0 + (var1 * var2)) - 0)"] - "linear_combination (-(y * z ^ 2) + x) * h + (z ^ 2 + 2 * z + 1) * h2 with exponent 2" + "linear_combination (-(y * z ^ 2) + x) * h + (z ^ 2 + 2 * z + 1) * h2 with {exponent := 2}" . example (K : Type) (h_this : 3 - 1 = 2) @@ -466,7 +465,10 @@ by test_polyrith "[(((var1 ^ 3) + (var0 * (3 * (var2 ^ 2)))) - 0), (((var0 ^ 3) + (var2 * (3 * (var1 ^ 2)))) - 0), (((var1 * (3 * (var0 ^ 2))) + (var2 ^ 3)) - 0), (((((var0 ^ 3) * var1) + ((var1 ^ 3) * var2)) + ((var2 ^ 3) * var0)) - 0)]", "(var0 - 0)"] "linear_combination 2 / 7 * y * z ^ 2 * h₂ + (x ^ 3 - 1 / 7 * y ^ 2 * z) * h₁ - x * y * z * h₀ + 1 / 7 * y * - z * h with exponent 6" + z * h with {exponent := 6}" + + + -- We comment the following tests so that we don't overwhelm the SageCell API. @@ -741,4 +743,22 @@ example (z : ℤ) (h1 : z + 1 = 2) (h2 : z + 2 = 2) : (1 : ℤ) = 2 := by create_polyrith_test +/- ### Cases with powers -/ + +example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := +by create_polyrith_test + +. +example (K : Type) (h_this : 3 - 1 = 2) + [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 create_polyrith_test + + -/ From 44cd468cb50a11a296cc4491509895b8584b05fd Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Thu, 22 Sep 2022 20:21:44 -0400 Subject: [PATCH 173/173] lint --- src/tactic/polyrith.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/tactic/polyrith.lean b/src/tactic/polyrith.lean index 875f970c49d49..8f4a80f250f81 100644 --- a/src/tactic/polyrith.lean +++ b/src/tactic/polyrith.lean @@ -490,7 +490,8 @@ a call to `linear_combination`. -/ meta def process_output (eq_names : list expr) (m : list expr) (R : expr) (sage_out : json) : tactic format := focus1 $ do - some (power, coeffs_as_poly) ← convert_sage_output sage_out | fail!"internal error: No output available", + some (power, coeffs_as_poly) ← convert_sage_output sage_out + | fail!"internal error: No output available", coeffs_as_pexpr ← coeffs_as_poly.mmap (poly.to_pexpr m), let eq_names_pexpr := eq_names.map to_pexpr, coeffs_as_expr ← coeffs_as_pexpr.mmap $ λ e, to_expr ``(%%e : %%R), @@ -499,7 +500,8 @@ meta def process_output (eq_names : list expr) (m : list expr) (R : expr) (sage_ $ λ pr, bnot $ pr.2.is_app_of `has_zero.zero, expr_string ← components_to_lc_format components, let lc_exp : format := if power = 1 then "" else format!" with {{exponent := {power}}}", - let lc_fmt : format := "linear_combination " ++ format.nest 2 (format.group expr_string ++ lc_exp), + let lc_fmt : format := "linear_combination " ++ + format.nest 2 (format.group expr_string ++ lc_exp), done <|> fail!"polyrith found the following certificate, but it failed to close the goal:\n{lc_fmt}", return lc_fmt