@@ -568,11 +568,13 @@ lemma continuous_min : continuous (λ p : α × α, min p.1 p.2) := continuous_f
568
568
569
569
lemma continuous_max : continuous (λ p : α × α, max p.1 p.2 ) := continuous_fst.max continuous_snd
570
570
571
- lemma tendsto.max {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁)) (hg : tendsto g b (𝓝 a₂)) :
571
+ lemma filter.tendsto.max {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁))
572
+ (hg : tendsto g b (𝓝 a₂)) :
572
573
tendsto (λb, max (f b) (g b)) b (𝓝 (max a₁ a₂)) :=
573
574
(continuous_max.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg)
574
575
575
- lemma tendsto.min {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁)) (hg : tendsto g b (𝓝 a₂)) :
576
+ lemma filter.tendsto.min {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (𝓝 a₁))
577
+ (hg : tendsto g b (𝓝 a₂)) :
576
578
tendsto (λb, min (f b) (g b)) b (𝓝 (min a₁ a₂)) :=
577
579
(continuous_min.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg)
578
580
@@ -1599,16 +1601,18 @@ lemma filter.tendsto.div_at_top [has_continuous_mul α] {f g : β → α} {l : f
1599
1601
(h : tendsto f l (𝓝 a)) (hg : tendsto g l at_top) : tendsto (λ x, f x / g x) l (𝓝 0 ) :=
1600
1602
by { simp only [div_eq_mul_inv], exact mul_zero a ▸ h.mul (tendsto_inv_at_top_zero.comp hg) }
1601
1603
1602
- lemma tendsto.inv_tendsto_at_top (h : tendsto f l at_top) : tendsto (f⁻¹) l (𝓝 0 ) :=
1604
+ lemma filter. tendsto.inv_tendsto_at_top (h : tendsto f l at_top) : tendsto (f⁻¹) l (𝓝 0 ) :=
1603
1605
tendsto_inv_at_top_zero.comp h
1604
1606
1605
- lemma tendsto.inv_tendsto_zero (h : tendsto f l (𝓝[set.Ioi 0 ] 0 )) : tendsto (f⁻¹) l at_top :=
1607
+ lemma filter.tendsto.inv_tendsto_zero (h : tendsto f l (𝓝[set.Ioi 0 ] 0 )) :
1608
+ tendsto (f⁻¹) l at_top :=
1606
1609
tendsto_inv_zero_at_top.comp h
1607
1610
1608
1611
/-- The function `x^(-n)` tends to `0` at `+∞` for any positive natural `n`.
1609
1612
A version for positive real powers exists as `tendsto_rpow_neg_at_top`. -/
1610
1613
lemma tendsto_pow_neg_at_top {n : ℕ} (hn : 1 ≤ n) : tendsto (λ x : α, x ^ (-(n:ℤ))) at_top (𝓝 0 ) :=
1611
- tendsto.congr (λ x, (fpow_neg x n).symm) (tendsto.inv_tendsto_at_top (tendsto_pow_at_top hn))
1614
+ tendsto.congr (λ x, (fpow_neg x n).symm)
1615
+ (filter.tendsto.inv_tendsto_at_top (tendsto_pow_at_top hn))
1612
1616
1613
1617
lemma tendsto_fpow_at_top_zero {n : ℤ} (hn : n < 0 ) :
1614
1618
tendsto (λ x : α, x^n) at_top (𝓝 0 ) :=
0 commit comments