Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 7eacca3

Browse files
committed
feat(analysis/normed/normed_field): limit of ∥a * x∥ as ∥x∥ → ∞ (#13819)
These lemmas should use `bornology.cobounded` but we don't have an instance `pseudo_metric_space α -> bornology α` yet.
1 parent 03fbe7d commit 7eacca3

File tree

2 files changed

+20
-2
lines changed

2 files changed

+20
-2
lines changed

src/analysis/normed/group/basic.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -166,6 +166,12 @@ by rw [sub_eq_add_neg, dist_self_add_right, norm_neg]
166166
@[simp] theorem dist_self_sub_left (g h : E) : dist (g - h) g = ∥h∥ :=
167167
by rw [dist_comm, dist_self_sub_right]
168168

169+
/-- In a (semi)normed group, negation `x ↦ -x` tends to infinity at infinity. TODO: use
170+
`bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/
171+
lemma filter.tendsto_neg_cobounded :
172+
tendsto (has_neg.neg : E → E) (comap norm at_top) (comap norm at_top) :=
173+
by simpa only [norm_neg, tendsto_comap_iff, (∘)] using tendsto_comap
174+
169175
/-- **Triangle inequality** for the norm. -/
170176
lemma norm_add_le (g h : E) : ∥g + h∥ ≤ ∥g∥ + ∥h∥ :=
171177
by simpa [dist_eq_norm] using dist_triangle g 0 (-h)

src/analysis/normed/normed_field.lean

Lines changed: 14 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -385,8 +385,6 @@ instance normed_division_ring.to_norm_one_class : norm_one_class α :=
385385
⟨mul_left_cancel₀ (mt norm_eq_zero.1 (@one_ne_zero α _ _)) $
386386
by rw [← norm_mul, mul_one, mul_one]⟩
387387

388-
export norm_one_class (norm_one)
389-
390388
@[simp] lemma nnnorm_mul (a b : α) : ∥a * b∥₊ = ∥a∥₊ * ∥b∥₊ :=
391389
nnreal.eq $ norm_mul a b
392390

@@ -422,6 +420,20 @@ nnreal.eq $ by simp
422420
@[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n :=
423421
(nnnorm_hom : α →*₀ ℝ≥0).map_zpow
424422

423+
/-- Multiplication on the left by a nonzero element of a normed division ring tends to infinity at
424+
infinity. TODO: use `bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/
425+
lemma filter.tendsto_mul_left_cobounded {a : α} (ha : a ≠ 0) :
426+
tendsto ((*) a) (comap norm at_top) (comap norm at_top) :=
427+
by simpa only [tendsto_comap_iff, (∘), norm_mul]
428+
using tendsto_const_nhds.mul_at_top (norm_pos_iff.2 ha) tendsto_comap
429+
430+
/-- Multiplication on the right by a nonzero element of a normed division ring tends to infinity at
431+
infinity. TODO: use `bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/
432+
lemma filter.tendsto_mul_right_cobounded {a : α} (ha : a ≠ 0) :
433+
tendsto (λ x, x * a) (comap norm at_top) (comap norm at_top) :=
434+
by simpa only [tendsto_comap_iff, (∘), norm_mul]
435+
using tendsto_comap.at_top_mul (norm_pos_iff.2 ha) tendsto_const_nhds
436+
425437
@[priority 100] -- see Note [lower instance priority]
426438
instance normed_division_ring.to_has_continuous_inv₀ : has_continuous_inv₀ α :=
427439
begin

0 commit comments

Comments
 (0)