diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 24d5639dc4898..104230a778405 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -70,8 +70,11 @@ open_locale pointwise topological_space variables {π•œ E : Type*} -section normed_field -variables (π•œ) [normed_field π•œ] [add_comm_group E] [module π•œ E] +section semi_normed_ring +variables [semi_normed_ring π•œ] + +section has_scalar +variables (π•œ) [has_scalar π•œ E] /-- A set `A` absorbs another set `B` if `B` is contained in all scalings of `A` by elements of sufficiently large norms. -/ @@ -84,77 +87,44 @@ def absorbent (A : set E) := βˆ€ x, βˆƒ r, 0 < r ∧ βˆ€ a : π•œ, r ≀ βˆ₯aβˆ₯ has norm less than or equal to one. -/ def balanced (A : set E) := βˆ€ a : π•œ, βˆ₯aβˆ₯ ≀ 1 β†’ a β€’ A βŠ† A -variables {π•œ} (a : π•œ) {A B : set E} +variables {π•œ} {A B : set E} -/-- A balanced set absorbs itself. -/ -lemma balanced.absorbs_self (hA : balanced π•œ A) : absorbs π•œ A A := -begin - use [1, zero_lt_one], - intros a ha x hx, - rw mem_smul_set_iff_inv_smul_memβ‚€, - { apply hA a⁻¹, - { rw norm_inv, exact inv_le_one ha }, - { rw mem_smul_set, use [x, hx] }}, - { rw ←norm_pos_iff, calc 0 < 1 : zero_lt_one ... ≀ βˆ₯aβˆ₯ : ha, } -end +lemma balanced.univ : balanced π•œ (univ : set E) := Ξ» a ha, subset_univ _ -lemma balanced.univ : balanced π•œ (univ : set E) := -Ξ» a ha, subset_univ _ - -lemma balanced.union {A₁ Aβ‚‚ : set E} (hA₁ : balanced π•œ A₁) (hAβ‚‚ : balanced π•œ Aβ‚‚) : - balanced π•œ (A₁ βˆͺ Aβ‚‚) := +lemma balanced.union (hA : balanced π•œ A) (hB : balanced π•œ B) : balanced π•œ (A βˆͺ B) := begin intros a ha t ht, rw [smul_set_union] at ht, - exact ht.imp (Ξ» x, hA₁ _ ha x) (Ξ» x, hAβ‚‚ _ ha x), + exact ht.imp (Ξ» x, hA _ ha x) (Ξ» x, hB _ ha x), end -lemma balanced.inter {A₁ Aβ‚‚ : set E} (hA₁ : balanced π•œ A₁) (hAβ‚‚ : balanced π•œ Aβ‚‚) : - balanced π•œ (A₁ ∩ Aβ‚‚) := +end has_scalar + +section add_comm_group +variables [add_comm_group E] [module π•œ E] {A B : set E} + +lemma balanced.inter (hA : balanced π•œ A) (hB : balanced π•œ B) : balanced π•œ (A ∩ B) := begin rintro a ha _ ⟨x, ⟨hx₁, hxβ‚‚βŸ©, rfl⟩, - exact ⟨hA₁ _ ha ⟨_, hx₁, rfl⟩, hAβ‚‚ _ ha ⟨_, hxβ‚‚, rfl⟩⟩, + exact ⟨hA _ ha ⟨_, hx₁, rfl⟩, hB _ ha ⟨_, hxβ‚‚, rfl⟩⟩, end -lemma balanced.add {A₁ Aβ‚‚ : set E} (hA₁ : balanced π•œ A₁) (hAβ‚‚ : balanced π•œ Aβ‚‚) : - balanced π•œ (A₁ + Aβ‚‚) := +lemma balanced.add (hA₁ : balanced π•œ A) (hAβ‚‚ : balanced π•œ B) : balanced π•œ (A + B) := begin rintro a ha _ ⟨_, ⟨x, y, hx, hy, rfl⟩, rfl⟩, rw smul_add, exact ⟨_, _, hA₁ _ ha ⟨_, hx, rfl⟩, hAβ‚‚ _ ha ⟨_, hy, rfl⟩, rfl⟩, end -lemma balanced.smul (hA : balanced π•œ A) : balanced π•œ (a β€’ A) := -begin - rintro b hb _ ⟨_, ⟨x, hx, rfl⟩, rfl⟩, - exact ⟨b β€’ x, hA _ hb ⟨_, hx, rfl⟩, smul_comm _ _ _⟩, -end - -lemma balanced.subset_smul (hA : balanced π•œ A) {a : π•œ} (ha : 1 ≀ βˆ₯aβˆ₯) : A βŠ† a β€’ A := -begin - refine (subset_set_smul_iffβ‚€ _).2 (hA (a⁻¹) _), - { rintro rfl, - rw norm_zero at ha, - exact zero_lt_one.not_le ha }, - { rw norm_inv, - exact inv_le_one ha } -end - -lemma balanced.smul_eq (hA : balanced π•œ A) {a : π•œ} (ha : βˆ₯aβˆ₯ = 1) : a β€’ A = A := -(hA _ ha.le).antisymm $ hA.subset_smul ha.ge - lemma absorbent.subset (hA : absorbent π•œ A) (hAB : A βŠ† B) : absorbent π•œ B := begin rintro x, obtain ⟨r, hr, hx⟩ := hA x, - refine ⟨r, hr, Ξ» a ha, (set_smul_subset_set_smul_iffβ‚€ _).2 hAB $ hx a ha⟩, - rintro rfl, - rw norm_zero at ha, - exact hr.not_le ha, + exact ⟨r, hr, Ξ» a ha, set.smul_set_mono hAB $ hx a ha⟩, end lemma absorbent_iff_forall_absorbs_singleton : absorbent π•œ A ↔ βˆ€ x, absorbs π•œ A {x} := -by simp [absorbs, absorbent] +by simp_rw [absorbs, absorbent, singleton_subset_iff] lemma absorbent_iff_nonneg_lt : absorbent π•œ A ↔ βˆ€ x, βˆƒ r, 0 ≀ r ∧ βˆ€ a : π•œ, r < βˆ₯aβˆ₯ β†’ x ∈ a β€’ A := begin @@ -168,9 +138,50 @@ begin Ξ» a ha, hx a ((lt_add_of_pos_right r zero_lt_one).trans_le ha)⟩ } end -/-! -Properties of balanced and absorbent sets in a topological vector space: --/ +end add_comm_group +end semi_normed_ring + +section normed_comm_ring +variables [normed_comm_ring π•œ] [add_comm_monoid E] [module π•œ E] {A B : set E} (a : π•œ) + +lemma balanced.smul (hA : balanced π•œ A) : balanced π•œ (a β€’ A) := +begin + rintro b hb _ ⟨_, ⟨x, hx, rfl⟩, rfl⟩, + exact ⟨b β€’ x, hA _ hb ⟨_, hx, rfl⟩, smul_comm _ _ _⟩, +end + +end normed_comm_ring + +section normed_field +variables [normed_field π•œ] [add_comm_group E] [module π•œ E] {A B : set E} {a : π•œ} + +/-- A balanced set absorbs itself. -/ +lemma balanced.absorbs_self (hA : balanced π•œ A) : absorbs π•œ A A := +begin + use [1, zero_lt_one], + intros a ha x hx, + rw mem_smul_set_iff_inv_smul_memβ‚€, + { apply hA a⁻¹, + { rw norm_inv, exact inv_le_one ha }, + { rw mem_smul_set, use [x, hx] }}, + { rw ←norm_pos_iff, calc 0 < 1 : zero_lt_one ... ≀ βˆ₯aβˆ₯ : ha, } +end + +lemma balanced.subset_smul (hA : balanced π•œ A) (ha : 1 ≀ βˆ₯aβˆ₯) : A βŠ† a β€’ A := +begin + refine (subset_set_smul_iffβ‚€ _).2 (hA (a⁻¹) _), + { rintro rfl, + rw norm_zero at ha, + exact zero_lt_one.not_le ha }, + { rw norm_inv, + exact inv_le_one ha } +end + +lemma balanced.smul_eq (hA : balanced π•œ A) (ha : βˆ₯aβˆ₯ = 1) : a β€’ A = A := +(hA _ ha.le).antisymm $ hA.subset_smul ha.ge + +/-! #### Topological vector space -/ + variables [topological_space E] [has_continuous_smul π•œ E] /-- Every neighbourhood of the origin is absorbent. -/ @@ -185,19 +196,15 @@ begin use [(r/2)⁻¹, hr₃], intros a ha₁, have haβ‚‚ : 0 < βˆ₯aβˆ₯ := hr₃.trans_le ha₁, - have ha₃ : a ⁻¹ β€’ x ∈ w, - { apply hrβ‚‚, - rw [metric.mem_ball, dist_zero_right, norm_inv], - calc βˆ₯aβˆ₯⁻¹ ≀ r/2 : (inv_le (half_pos hr₁) haβ‚‚).mp ha₁ - ... < r : half_lt_self hr₁ }, rw [mem_smul_set_iff_inv_smul_memβ‚€ (norm_pos_iff.mp haβ‚‚)], - exact hw₁ ha₃, + refine hw₁ (hrβ‚‚ _), + rw [metric.mem_ball, dist_zero_right, norm_inv], + calc βˆ₯aβˆ₯⁻¹ ≀ r/2 : (inv_le (half_pos hr₁) haβ‚‚).mp ha₁ + ... < r : half_lt_self hr₁, end -/-- The union of `{0}` with the interior of a balanced set - is balanced. -/ -lemma balanced_zero_union_interior (hA : balanced π•œ A) : - balanced π•œ ({(0 : E)} βˆͺ interior A) := +/-- The union of `{0}` with the interior of a balanced set is balanced. -/ +lemma balanced_zero_union_interior (hA : balanced π•œ A) : balanced π•œ ({(0 : E)} βˆͺ interior A) := begin intros a ha, by_cases a = 0, { rw [h, zero_smul_set], @@ -233,18 +240,23 @@ end normed_field /-- A seminorm on a vector space over a normed field is a function to the reals that is positive semidefinite, positive homogeneous, and subadditive. -/ -structure seminorm (π•œ : Type*) (E : Type*) - [normed_field π•œ] [add_comm_group E] [module π•œ E] := +structure seminorm (π•œ : Type*) (E : Type*) [semi_normed_ring π•œ] [add_monoid E] [has_scalar π•œ E] := (to_fun : E β†’ ℝ) (smul' : βˆ€ (a : π•œ) (x : E), to_fun (a β€’ x) = βˆ₯aβˆ₯ * to_fun x) (triangle' : βˆ€ x y : E, to_fun (x + y) ≀ to_fun x + to_fun y) namespace seminorm -section normed_field -variables [normed_field π•œ] [add_comm_group E] [module π•œ E] +section semi_normed_ring +variables [semi_normed_ring π•œ] + +section add_monoid +variables [add_monoid E] + +section has_scalar +variables [has_scalar π•œ E] instance : inhabited (seminorm π•œ E) := -⟨{ to_fun := Ξ» _, 0, +⟨{ to_fun := Ξ» _, 0, smul' := Ξ» _ _, (mul_zero _).symm, triangle' := Ξ» x y, by rw add_zero }⟩ @@ -262,23 +274,35 @@ variables (p : seminorm π•œ E) (c : π•œ) (x y : E) (r : ℝ) protected lemma smul : p (c β€’ x) = βˆ₯cβˆ₯ * p x := p.smul' _ _ protected lemma triangle : p (x + y) ≀ p x + p y := p.triangle' _ _ -protected lemma sub_le : p (x - y) ≀ p x + p y := -calc - p (x - y) - = p (x + -y) : by rw sub_eq_add_neg - ... ≀ p x + p (-y) : p.triangle x (-y) - ... = p x + p y : by rw [←neg_one_smul π•œ y, p.smul, norm_neg, norm_one, one_mul] + +end has_scalar + +section smul_with_zero +variables [smul_with_zero π•œ E] (p : seminorm π•œ E) @[simp] protected lemma zero : p 0 = 0 := calc p 0 = p ((0 : π•œ) β€’ 0) : by rw zero_smul ... = 0 : by rw [p.smul, norm_zero, zero_mul] +end smul_with_zero +end add_monoid + +section norm_one_class +variables [norm_one_class π•œ] [add_comm_group E] [module π•œ E] (p : seminorm π•œ E) (x y : E) (r : ℝ) + @[simp] protected lemma neg : p (-x) = p x := calc p (-x) = p ((-1 : π•œ) β€’ x) : by rw neg_one_smul ... = p x : by rw [p.smul, norm_neg, norm_one, one_mul] +protected lemma sub_le : p (x - y) ≀ p x + p y := +calc + p (x - y) + = p (x + -y) : by rw sub_eq_add_neg + ... ≀ p x + p (-y) : p.triangle x (-y) + ... = p x + p y : by rw p.neg + lemma nonneg : 0 ≀ p x := have h: 0 ≀ 2 * p x, from calc 0 = p (x + (- x)) : by rw [add_neg_self, p.zero] @@ -286,24 +310,37 @@ calc 0 = p (x + (- x)) : by rw [add_neg_self, p.zero] ... = 2 * p x : by rw [p.neg, two_mul], nonneg_of_mul_nonneg_left h zero_lt_two -lemma sub_rev : p (x - y) = p (y - x) := -by rw [←neg_sub, p.neg] +lemma sub_rev : p (x - y) = p (y - x) := by rw [←neg_sub, p.neg] + +end norm_one_class -/-- The ball of radius `r` at `x` with respect to seminorm `p` - is the set of elements `y` with `p (y - x) < `r`. -/ -def ball (p : seminorm π•œ E) (x : E) (r : ℝ) := { y : E | p (y - x) < r } +/-! ### Seminorm ball -/ -lemma mem_ball : y ∈ ball p x r ↔ p (y - x) < r := -iff.rfl +section add_comm_group +variables [add_comm_group E] -lemma mem_ball_zero : y ∈ ball p 0 r ↔ p y < r := -by rw [mem_ball, sub_zero] +section has_scalar +variables [has_scalar π•œ E] (p : seminorm π•œ E) + +/-- The ball of radius `r` at `x` with respect to seminorm `p` is the set of elements `y` with +`p (y - x) < `r`. -/ +def ball (x : E) (r : ℝ) := { y : E | p (y - x) < r } + +variables {x y : E} {r : ℝ} + +lemma mem_ball : y ∈ ball p x r ↔ p (y - x) < r := iff.rfl -lemma ball_zero_eq : ball p 0 r = { y : E | p y < r } := -set.ext $ Ξ» x,by { rw mem_ball_zero, exact iff.rfl } +lemma mem_ball_zero : y ∈ ball p 0 r ↔ p y < r := by rw [mem_ball, sub_zero] + +lemma ball_zero_eq : ball p 0 r = { y : E | p y < r } := set.ext $ Ξ» x, p.mem_ball_zero + +end has_scalar + +section module +variables [norm_one_class π•œ] [module π•œ E] (p : seminorm π•œ E) /-- Seminorm-balls at the origin are balanced. -/ -lemma balanced_ball_zero : balanced π•œ (ball p 0 r) := +lemma balanced_ball_zero (r : ℝ): balanced π•œ (ball p 0 r) := begin rintro a ha x ⟨y, hy, hx⟩, rw [mem_ball_zero, ←hx, p.smul], @@ -311,8 +348,16 @@ begin ... < r : by rwa mem_ball_zero at hy, end +end module +end add_comm_group +end semi_normed_ring + +section normed_field +variables [normed_field π•œ] [add_comm_group E] [module π•œ E] (p : seminorm π•œ E) {A B : set E} + {a : π•œ} {r : ℝ} {x : E} + /-- Seminorm-balls at the origin are absorbent. -/ -lemma absorbent_ball_zero {r : ℝ} (hr : 0 < r) : absorbent π•œ (ball p (0 : E) r) := +lemma absorbent_ball_zero (hr : 0 < r) : absorbent π•œ (ball p (0 : E) r) := begin rw absorbent_iff_nonneg_lt, rintro x, @@ -328,10 +373,10 @@ lemma absorbent_ball (hpr : p x < r) : absorbent π•œ (ball p x r) := begin refine (p.absorbent_ball_zero $ sub_pos.2 hpr).subset (Ξ» y hy, _), rw p.mem_ball_zero at hy, - exact (p.mem_ball _ _ _).2 ((p.sub_le _ _).trans_lt $ add_lt_of_lt_sub_right hy), + exact p.mem_ball.2 ((p.sub_le _ _).trans_lt $ add_lt_of_lt_sub_right hy), end -lemma symmetric_ball_zero {x : E} (hx : x ∈ ball p 0 r) : -x ∈ ball p 0 r := +lemma symmetric_ball_zero (r : ℝ) (hx : x ∈ ball p 0 r) : -x ∈ ball p 0 r := balanced_ball_zero p r (-1) (by rw [norm_neg, norm_one]) ⟨x, hx, by rw [neg_smul, one_smul]⟩ end normed_field @@ -364,7 +409,7 @@ lemma convex_ball : convex ℝ (ball p x r) := begin convert (p.convex_on.translate_left (-x)).convex_lt r, ext y, - rw [preimage_univ, sep_univ, p.mem_ball x y r, sub_eq_add_neg], + rw [preimage_univ, sep_univ, p.mem_ball, sub_eq_add_neg], refl, end