@@ -385,8 +385,6 @@ instance normed_division_ring.to_norm_one_class : norm_one_class α :=
385
385
⟨mul_left_cancel₀ (mt norm_eq_zero.1 (@one_ne_zero α _ _)) $
386
386
by rw [← norm_mul, mul_one, mul_one]⟩
387
387
388
- export norm_one_class (norm_one)
389
-
390
388
@[simp] lemma nnnorm_mul (a b : α) : ∥a * b∥₊ = ∥a∥₊ * ∥b∥₊ :=
391
389
nnreal.eq $ norm_mul a b
392
390
@@ -422,6 +420,20 @@ nnreal.eq $ by simp
422
420
@[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n :=
423
421
(nnnorm_hom : α →*₀ ℝ≥0 ).map_zpow
424
422
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
+
425
437
@[priority 100 ] -- see Note [lower instance priority]
426
438
instance normed_division_ring.to_has_continuous_inv₀ : has_continuous_inv₀ α :=
427
439
begin
0 commit comments