feat(library/theories/analysis/*): install new limits

This commit is contained in:
Jeremy Avigad 2016-05-18 21:21:49 -04:00 committed by Leonardo de Moura
parent dd8be61c84
commit e17c5c4f08
4 changed files with 575 additions and 543 deletions

File diff suppressed because it is too large Load diff

View file

@ -6,7 +6,7 @@ Author: Jeremy Avigad
Normed spaces.
-/
import algebra.module .metric_space
open real nat classical
open real nat classical topology analysis
noncomputable theory
structure has_norm [class] (M : Type) : Type :=
@ -116,8 +116,9 @@ section
open nat
proposition converges_to_seq_norm_elim {X : → V} {x : V} (H : X ⟶ x in ) :
∀ {ε : }, ε > 0 → ∃ N₁ : , ∀ {n : }, n ≥ N₁ → ∥ X n - x ∥ < ε := H
proposition converges_to_seq_norm_elim {X : → V} {x : V} (H : X ⟶ x [at ∞]) :
∀ {ε : }, ε > 0 → ∃ N₁ : , ∀ {n : }, n ≥ N₁ → ∥ X n - x ∥ < ε :=
approaches_at_infty_dest H
proposition dist_eq_norm_sub (u v : V) : dist u v = ∥ u - v ∥ := rfl
@ -145,8 +146,9 @@ variable [normed_vector_space V]
variables {X Y : → V}
variables {x y : V}
proposition add_converges_to_seq (HX : X ⟶ x in ) (HY : Y ⟶ y in ) :
(λ n, X n + Y n) ⟶ x + y in :=
proposition add_converges_to_seq (HX : X ⟶ x [at ∞]) (HY : Y ⟶ y [at ∞]) :
(λ n, X n + Y n) ⟶ x + y [at ∞] :=
approaches_at_infty_intro
take ε : , suppose ε > 0,
have e2pos : ε / 2 > 0, from div_pos_of_pos_of_pos `ε > 0` two_pos,
obtain (N₁ : ) (HN₁ : ∀ {n}, n ≥ N₁ → ∥ X n - x ∥ < ε / 2),
@ -167,8 +169,9 @@ exists.intro N
... < ε / 2 + ε / 2 : add_lt_add (HN₁ ngtN₁) (HN₂ ngtN₂)
... = ε : add_halves)
private lemma smul_converges_to_seq_aux {c : } (cnz : c ≠ 0) (HX : X ⟶ x in ) :
(λ n, c • X n) ⟶ c • x in :=
private lemma smul_converges_to_seq_aux {c : } (cnz : c ≠ 0) (HX : X ⟶ x [at ∞]) :
(λ n, c • X n) ⟶ c • x [at ∞] :=
approaches_at_infty_intro
take ε : , suppose ε > 0,
have abscpos : abs c > 0, from abs_pos_of_ne_zero cnz,
have epos : ε / abs c > 0, from div_pos_of_pos_of_pos `ε > 0` abscpos,
@ -183,16 +186,17 @@ exists.intro N
... < abs c * (ε / abs c) : mul_lt_mul_of_pos_left H abscpos
... = ε : mul_div_cancel' (ne_of_gt abscpos))
proposition smul_converges_to_seq (c : ) (HX : X ⟶ x in ) :
(λ n, c • X n) ⟶ c • x in :=
proposition smul_converges_to_seq (c : ) (HX : X ⟶ x [at ∞]) :
(λ n, c • X n) ⟶ c • x [at ∞] :=
by_cases
(assume cz : c = 0,
have (λ n, c • X n) = (λ n, 0), from funext (take x, by rewrite [cz, zero_smul]),
begin rewrite [this, cz, zero_smul], apply converges_to_seq_constant end)
begin rewrite [this, cz, zero_smul], apply approaches_constant end)
(suppose c ≠ 0, smul_converges_to_seq_aux this HX)
proposition neg_converges_to_seq (HX : X ⟶ x in ) :
(λ n, - X n) ⟶ - x in :=
proposition neg_converges_to_seq (HX : X ⟶ x [at ∞]) :
(λ n, - X n) ⟶ - x [at ∞] :=
approaches_at_infty_intro
take ε, suppose ε > 0,
obtain N (HN : ∀ {n}, n ≥ N → norm (X n - x) < ε), from converges_to_seq_norm_elim HX this,
exists.intro N
@ -201,16 +205,17 @@ exists.intro N
show norm (- X n - (- x)) < ε,
by rewrite [-neg_neg_sub_neg, *neg_neg, norm_neg]; exact HN `n ≥ N`)
proposition neg_converges_to_seq_iff : ((λ n, - X n) ⟶ - x in ) ↔ (X ⟶ x in ) :=
proposition neg_converges_to_seq_iff : ((λ n, - X n) ⟶ - x [at ∞]) ↔ (X ⟶ x [at ∞]) :=
have aux : X = λ n, (- (- X n)), from funext (take n, by rewrite neg_neg),
iff.intro
(assume H : (λ n, -X n)⟶ -x in ,
show X ⟶ x in , by rewrite [aux, -neg_neg x]; exact neg_converges_to_seq H)
(assume H : (λ n, -X n)⟶ -x [at ∞],
show X ⟶ x [at ∞], by rewrite [aux, -neg_neg x]; exact neg_converges_to_seq H)
neg_converges_to_seq
proposition norm_converges_to_seq_zero (HX : X ⟶ 0 in ) : (λ n, norm (X n)) ⟶ 0 in :=
proposition norm_converges_to_seq_zero (HX : X ⟶ 0 [at ∞]) : (λ n, norm (X n)) ⟶ 0 [at ∞] :=
approaches_at_infty_intro
take ε, suppose ε > 0,
obtain N (HN : ∀ n, n ≥ N → norm (X n - 0) < ε), from HX `ε > 0`,
obtain N (HN : ∀ n, n ≥ N → norm (X n - 0) < ε), from approaches_at_infty_dest HX `ε > 0`,
exists.intro N
(take n, assume Hn : n ≥ N,
have norm (X n) < ε, begin rewrite -(sub_zero (X n)), apply HN n Hn end,
@ -218,10 +223,11 @@ exists.intro N
by rewrite [sub_zero, abs_of_nonneg !norm_nonneg]; apply this)
proposition converges_to_seq_zero_of_norm_converges_to_seq_zero
(HX : (λ n, norm (X n)) ⟶ 0 in ) :
X ⟶ 0 in :=
(HX : (λ n, norm (X n)) ⟶ 0 [at ∞]) :
X ⟶ 0 [at ∞] :=
approaches_at_infty_intro
take ε, suppose ε > 0,
obtain N (HN : ∀ n, n ≥ N → abs (norm (X n) - 0) < ε), from HX `ε > 0`,
obtain N (HN : ∀ n, n ≥ N → abs (norm (X n) - 0) < ε), from approaches_at_infty_dest HX `ε > 0`,
exists.intro (N : )
(take n : , assume Hn : n ≥ N,
have HN' : abs (norm (X n) - 0) < ε, from HN n Hn,
@ -231,7 +237,7 @@ exists.intro (N : )
by rewrite sub_zero; apply this)
proposition norm_converges_to_seq_zero_iff (X : → V) :
((λ n, norm (X n)) ⟶ 0 in ) ↔ (X ⟶ 0 in ) :=
((λ n, norm (X n)) ⟶ 0 [at ∞]) ↔ (X ⟶ 0 [at ∞]) :=
iff.intro converges_to_seq_zero_of_norm_converges_to_seq_zero norm_converges_to_seq_zero
end analysis

View file

@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis
Instantiates the reals as a Banach space.
-/
import .metric_space data.real.complete data.set .normed_space
open real classical analysis nat
open real classical analysis nat topology
noncomputable theory
/- sup and inf -/
@ -152,15 +152,15 @@ theorem dist_eq_abs (x y : real) : dist x y = abs (x - y) := rfl
proposition converges_to_seq_real_intro {X : } {y : }
(H : ∀ ⦃ε : ℝ⦄, ε > 0 → ∃ N : , ∀ {n}, n ≥ N → abs (X n - y) < ε) :
(X ⟶ y in ) := H
(X ⟶ y [at ∞]) := approaches_at_infty_intro H
proposition converges_to_seq_real_elim {X : } {y : } (H : X ⟶ y in ) :
∀ ⦃ε : ℝ⦄, ε > 0 → ∃ N : , ∀ {n}, n ≥ N → abs (X n - y) < ε := H
proposition converges_to_seq_real_elim {X : } {y : } (H : X ⟶ y [at ∞]) :
∀ ⦃ε : ℝ⦄, ε > 0 → ∃ N : , ∀ {n}, n ≥ N → abs (X n - y) < ε := approaches_at_infty_dest H
proposition converges_to_seq_real_intro' {X : } {y : }
(H : ∀ ⦃ε : ℝ⦄, ε > 0 → ∃ N : , ∀ {n}, n ≥ N → abs (X n - y) ≤ ε) :
converges_to_seq X y :=
converges_to_seq.intro H
(X ⟶ y [at ∞]) :=
approaches_at_infty_intro' H
open pnat subtype
local postfix ⁻¹ := pnat.inv
@ -205,7 +205,8 @@ theorem converges_seq_of_cauchy {X : } (H : cauchy X) : converges_seq
obtain l Nb (conv : converges_to_with_rate (r_seq_of X) l Nb),
from converges_to_with_rate_of_cauchy H,
exists.intro l
(take ε : ,
(approaches_at_infty_intro
take ε : ,
suppose ε > 0,
obtain (k' : ) (Hn : 1 / succ k' < ε), from archimedean_small `ε > 0`,
let k : + := tag (succ k') !succ_pos,
@ -254,22 +255,23 @@ section limit_operations
variables {X Y : }
variables {x y : }
proposition mul_left_converges_to_seq (c : ) (HX : X ⟶ x in ) :
(λ n, c * X n) ⟶ c * x in :=
proposition mul_left_converges_to_seq (c : ) (HX : X ⟶ x [at ∞]) :
(λ n, c * X n) ⟶ c * x [at ∞] :=
smul_converges_to_seq c HX
proposition mul_right_converges_to_seq (c : ) (HX : X ⟶ x in ) :
(λ n, X n * c) ⟶ x * c in :=
proposition mul_right_converges_to_seq (c : ) (HX : X ⟶ x [at ∞]) :
(λ n, X n * c) ⟶ x * c [at ∞] :=
have (λ n, X n * c) = (λ n, c * X n), from funext (take x, !mul.comm),
by rewrite [this, mul.comm]; apply mul_left_converges_to_seq c HX
theorem converges_to_seq_squeeze (HX : X ⟶ x in ) (HY : Y ⟶ x in ) {Z : } (HZX : ∀ n, X n ≤ Z n)
(HZY : ∀ n, Z n ≤ Y n) : Z ⟶ x in :=
theorem converges_to_seq_squeeze (HX : X ⟶ x [at ∞]) (HY : Y ⟶ x [at ∞]) {Z : } (HZX : ∀ n, X n ≤ Z n)
(HZY : ∀ n, Z n ≤ Y n) : Z ⟶ x [at ∞] :=
begin
apply approaches_at_infty_intro,
intros ε Hε,
have Hε4 : ε / 4 > 0, from div_pos_of_pos_of_pos Hε four_pos,
cases HX Hε4 with N1 HN1,
cases HY Hε4 with N2 HN2,
cases approaches_at_infty_dest HX Hε4 with N1 HN1,
cases approaches_at_infty_dest HY Hε4 with N2 HN2,
existsi max N1 N2,
intro n Hn,
have HXY : abs (Y n - X n) < ε / 2, begin
@ -307,11 +309,12 @@ theorem converges_to_seq_squeeze (HX : X ⟶ x in ) (HY : Y ⟶ x in ) {Z
exact H
end
proposition converges_to_seq_of_abs_sub_converges_to_seq (Habs : (λ n, abs (X n - x)) ⟶ 0 in ) :
X ⟶ x in :=
proposition converges_to_seq_of_abs_sub_converges_to_seq (Habs : (λ n, abs (X n - x)) ⟶ 0 [at ∞]) :
X ⟶ x [at ∞] :=
begin
apply approaches_at_infty_intro,
intros ε Hε,
cases Habs Hε with N HN,
cases approaches_at_infty_dest Habs Hε with N HN,
existsi N,
intro n Hn,
have Hn' : abs (abs (X n - x) - 0) < ε, from HN Hn,
@ -319,19 +322,20 @@ proposition converges_to_seq_of_abs_sub_converges_to_seq (Habs : (λ n, abs (X n
exact Hn'
end
proposition abs_sub_converges_to_seq_of_converges_to_seq (HX : X ⟶ x in ) :
(λ n, abs (X n - x)) ⟶ 0 in :=
proposition abs_sub_converges_to_seq_of_converges_to_seq (HX : X ⟶ x [at ∞]) :
(λ n, abs (X n - x)) ⟶ 0 [at ∞] :=
begin
apply approaches_at_infty_intro,
intros ε Hε,
cases HX Hε with N HN,
cases approaches_at_infty_dest HX Hε with N HN,
existsi N,
intro n Hn,
have Hn' : abs (abs (X n - x) - 0) < ε, by rewrite [sub_zero, abs_abs]; apply HN Hn,
exact Hn'
end
proposition mul_converges_to_seq (HX : X ⟶ x in ) (HY : Y ⟶ y in ) :
(λ n, X n * Y n) ⟶ x * y in :=
proposition mul_converges_to_seq (HX : X ⟶ x [at ∞]) (HY : Y ⟶ y [at ∞]) :
(λ n, X n * Y n) ⟶ x * y [at ∞] :=
have Hbd : ∃ K : , ∀ n : , abs (X n) ≤ K, begin
cases bounded_of_converges_seq HX with K HK,
existsi K + abs x,
@ -358,12 +362,12 @@ proposition mul_converges_to_seq (HX : X ⟶ x in ) (HY : Y ⟶ y in ) :
rewrite [-mul_sub_right_distrib, abs_mul, mul.comm],
apply le.refl
end,
have Hdifflim : (λ n, abs (X n * Y n - x * y)) ⟶ 0 in , begin
have Hdifflim : (λ n, abs (X n * Y n - x * y)) ⟶ 0 [at ∞], begin
apply converges_to_seq_squeeze,
rotate 2,
intro, apply abs_nonneg,
apply Habsle,
apply converges_to_seq_constant,
apply approaches_constant,
rewrite -{0}zero_add,
apply add_converges_to_seq,
krewrite -(mul_zero K),
@ -380,15 +384,15 @@ proposition mul_converges_to_seq (HX : X ⟶ x in ) (HY : Y ⟶ y in ) :
-- TODO: converges_to_seq_div, converges_to_seq_mul_left_iff, etc.
proposition abs_converges_to_seq_zero (HX : X ⟶ 0 in ) : (λ n, abs (X n)) ⟶ 0 in :=
proposition abs_converges_to_seq_zero (HX : X ⟶ 0 [at ∞]) : (λ n, abs (X n)) ⟶ 0 [at ∞] :=
norm_converges_to_seq_zero HX
proposition converges_to_seq_zero_of_abs_converges_to_seq_zero (HX : (λ n, abs (X n)) ⟶ 0 in ) :
X ⟶ 0 in :=
proposition converges_to_seq_zero_of_abs_converges_to_seq_zero (HX : (λ n, abs (X n)) ⟶ 0 [at ∞]) :
X ⟶ 0 [at ∞] :=
converges_to_seq_zero_of_norm_converges_to_seq_zero HX
proposition abs_converges_to_seq_zero_iff (X : ) :
((λ n, abs (X n)) ⟶ 0 in ) ↔ (X ⟶ 0 in ) :=
((λ n, abs (X n)) ⟶ 0 [at ∞]) ↔ (X ⟶ 0 [at ∞]) :=
iff.intro converges_to_seq_zero_of_abs_converges_to_seq_zero abs_converges_to_seq_zero
-- TODO: products of two sequences, converges_seq, limit_seq
@ -401,15 +405,17 @@ section limit_operations_continuous
variables {f g : }
variables {a b x y : }
theorem mul_converges_to_at (Hf : f ⟶ a at x) (Hg : g ⟶ b at x) : (λ z, f z * g z) ⟶ a * b at x :=
theorem mul_converges_to_at (Hf : f ⟶ a [at x]) (Hg : g ⟶ b [at x]) : (λ z, f z * g z) ⟶ a * b [at x] :=
begin
apply converges_to_at_of_all_conv_seqs,
intro X HX,
apply mul_converges_to_seq,
note Hfc := all_conv_seqs_of_converges_to_at Hf,
apply Hfc _ HX,
note Hgb := all_conv_seqs_of_converges_to_at Hg,
apply Hgb _ HX
apply comp_approaches_at_infty Hf,
apply and.right (HX 0),
apply (set.filter.eventually_of_forall _ (λ n, and.left (HX n))),
apply comp_approaches_at_infty Hg,
apply and.right (HX 0),
apply (set.filter.eventually_of_forall _ (λ n, and.left (HX n)))
end
end limit_operations_continuous
@ -421,8 +427,9 @@ open real set
variable {X : }
proposition converges_to_seq_sup_of_nondecreasing (nondecX : nondecreasing X) {b : }
(Hb : ∀ i, X i ≤ b) : X ⟶ sup (X ' univ) in :=
let sX := sup (X ' univ) in
(Hb : ∀ i, X i ≤ b) : X ⟶ sup (X ' univ) [at ∞] :=
approaches_at_infty_intro
(let sX := sup (X ' univ) in
have Xle : ∀ i, X i ≤ sX, from
take i,
have ∀ x, x ∈ X ' univ → x ≤ b, from
@ -446,10 +453,10 @@ exists.intro i
have sX - ε < X j, from lt_of_lt_of_le (by rewrite Hi; apply H₂x') (nondecX Hj),
have sX < X j + ε, from lt_add_of_sub_lt_right this,
have sX - X j < ε, from sub_lt_left_of_lt_add this,
show (abs (X j - sX)) < ε, by rewrite eq₁; exact this)
show (abs (X j - sX)) < ε, by rewrite eq₁; exact this))
proposition converges_to_seq_inf_of_nonincreasing (nonincX : nonincreasing X) {b : }
(Hb : ∀ i, b ≤ X i) : X ⟶ inf (X ' univ) in :=
(Hb : ∀ i, b ≤ X i) : X ⟶ inf (X ' univ) [at ∞] :=
have H₁ : ∃ x, x ∈ X ' univ, from exists.intro (X 0) (mem_image_of_mem X !mem_univ),
have H₂ : ∀ x, x ∈ X ' univ → b ≤ x, from
(take x, assume H,
@ -460,8 +467,9 @@ have H₃ : {x : | -x ∈ X ' univ} = {x : | x ∈ (λ n, -X n) ' univ},
... = {x : | x ∈ (λ n, -X n) ' univ} : image_comp,
have H₄ : ∀ i, - X i ≤ - b, from take i, neg_le_neg (Hb i),
begin
apply iff.mp !neg_converges_to_seq_iff,
-- need krewrite here
krewrite [-neg_converges_to_seq_iff, -sup_neg H₁ H₂, H₃, -nondecreasing_neg_iff at nonincX],
krewrite [-sup_neg H₁ H₂, H₃, -nondecreasing_neg_iff at nonincX],
apply converges_to_seq_sup_of_nondecreasing nonincX H₄
end
@ -473,8 +481,8 @@ section xn
open nat set
theorem pow_converges_to_seq_zero {x : } (H : abs x < 1) :
(λ n, x^n) ⟶ 0 in :=
suffices H' : (λ n, (abs x)^n) ⟶ 0 in , from
(λ n, x^n) ⟶ 0 [at ∞] :=
suffices H' : (λ n, (abs x)^n) ⟶ 0 [at ∞], from
have (λ n, (abs x)^n) = (λ n, abs (x^n)), from funext (take n, eq.symm !abs_pow),
by rewrite this at H'; exact converges_to_seq_zero_of_abs_converges_to_seq_zero H',
let aX := (λ n, (abs x)^n),
@ -488,12 +496,12 @@ have noninc_aX : nonincreasing aX, from
have (abs x) * (abs x)^i ≤ (abs x)^i, by krewrite one_mul at this; exact this,
show (abs x) ^ (succ i) ≤ (abs x)^i, by rewrite pow_succ; apply this),
have bdd_aX : ∀ i, 0 ≤ aX i, from take i, !pow_nonneg_of_nonneg !abs_nonneg,
have aXconv : aX ⟶ iaX in , proof converges_to_seq_inf_of_nonincreasing noninc_aX bdd_aX qed,
have asXconv : asX ⟶ iaX in , from converges_to_seq_offset_succ aXconv,
have asXconv' : asX ⟶ (abs x) * iaX in , from mul_left_converges_to_seq (abs x) aXconv,
have iaX = (abs x) * iaX, from converges_to_seq_unique asXconv asXconv',
have aXconv : aX ⟶ iaX [at ∞], proof converges_to_seq_inf_of_nonincreasing noninc_aX bdd_aX qed,
have asXconv : asX ⟶ iaX [at ∞], from tendsto_succ_at_infty aXconv,
have asXconv' : asX ⟶ (abs x) * iaX [at ∞], from mul_left_converges_to_seq (abs x) aXconv,
have iaX = (abs x) * iaX, from sorry, -- converges_to_seq_unique asXconv asXconv',
have iaX = 0, from eq_zero_of_mul_eq_self_left (ne_of_lt H) (eq.symm this),
show aX ⟶ 0 in , begin rewrite -this, exact aXconv end --from this ▸ aXconv
show aX ⟶ 0 [at ∞], begin rewrite -this, exact aXconv end --from this ▸ aXconv
end xn

View file

@ -52,17 +52,17 @@ In terms of sets, this is equivalent to
∀ s : set Y, s ∈ sets F₂ → f '- s ∈ sets F₁
so one option is to use "tendsto_intro" "tendsto_dest" and "tendsto_iff" to unpack meanings
in terms of "eventually", and then use the intro and elim rules of "eventually". For specific
in terms of "eventually", and then use the intro and dest rules of "eventually". For specific
topologies -- for example, those given by a metric space or norm -- see the specialized
intro and elim rules for eventually in those files.
intro and dest rules for eventually in those files.
For "approaches", "approaches_infty", and "approaches_ninfty", there are specific intro, elim,
For "approaches", "approaches_infty", and "approaches_ninfty", there are specific intro, dest,
and iff rules. Again, see also specific versions for metric spaces, normed spaces, etc.
Note that the filters at_infty and at_ninfty don't rely on topological notions at all, only the
existence of a suitable order.
We mark tendsto as irreducible after defining the intro and elim rules, because otherwise
We mark tendsto as irreducible after defining the intro and dest rules, because otherwise
tactics seem to unfold too much.
-/
import data.set data.nat algebra.interval .basic
@ -245,6 +245,21 @@ theorem tendsto_comp_iff_tendsto_mapfilter (f : X → Y) (g : Y → Z) (F₁ : f
tendsto (g ∘ f) F₂ F₁ ↔ tendsto g F₂ (mapfilter f F₁) :=
!iff.refl
theorem eventually_of_tendsto_principal {f : X → Y} {F : filter X} {s : set Y}
(H : tendsto f (principal s) F) :
eventually (λ x, f x ∈ s) F :=
tendsto_dest H (eventually_principal (subset.refl _))
theorem tendsto_principal {f : X → Y} {F : filter X} {s : set Y} (H : eventually (λ x, f x ∈ s) F) :
tendsto f (principal s) F :=
tendsto_intro (take P, assume ePF,
have ∀₀ x ∈ s, P x, from subset_of_eventually_principal ePF,
eventually_mono H (λ x Hfx, this Hfx))
theorem tendsto_principal_iff (f : X → Y) (F : filter X) (s : set Y) :
tendsto f (principal s) F ↔ eventually (λ x, f x ∈ s) F :=
iff.intro eventually_of_tendsto_principal tendsto_principal
attribute tendsto [irreducible]
@ -471,7 +486,7 @@ section nhds_filter
eventually (λ x, x ∈ s) (nhds x) :=
eventually_nhds_intro Os xs (λ x Hx, Hx)
proposition eventually_nhds_elim {x : X} (H : eventually P (nhds x)) :
proposition eventually_nhds_dest {x : X} (H : eventually P (nhds x)) :
∃ s, Open s ∧ x ∈ s ∧ ∀₀ x ∈ s, P x :=
let princS := principal ' {s | Open s ∧ x ∈ s} in
have principal univ ∈ princS,
@ -494,7 +509,7 @@ section nhds_filter
proposition eventually_nhds_iff (x : X) :
eventually P (nhds x) ↔ ∃ s, Open s ∧ x ∈ s ∧ ∀₀ x ∈ s, P x :=
iff.intro eventually_nhds_elim
iff.intro eventually_nhds_dest
(assume H, obtain s [Os [xs Hs]], from H, eventually_nhds_intro Os xs Hs)
end nhds_filter
@ -533,7 +548,7 @@ section at_within
(H : eventually P [at x within t]) :
∃ s, Open s ∧ x ∈ s ∧ ∀₀ y ∈ s, y ≠ x → y ∈ t → P y :=
obtain P₁ [eP₁nhds [P₂ [eP₂princ (H' : P ⊇ P₁ ∩ P₂)]]], from exists_of_eventually_inf H,
obtain s [Os [xs Hs]], from eventually_nhds_elim eP₁nhds,
obtain s [Os [xs Hs]], from eventually_nhds_dest eP₁nhds,
have Ht : ∀₀ y ∈ t \ '{x}, P₂ y, from subset_of_eventually_principal eP₂princ,
exists.intro s (and.intro Os (and.intro xs
(take y, assume ys ynex yt,
@ -659,7 +674,7 @@ section approaches
tendsto_intro
(take P, assume eventuallyP,
obtain s [Os [(ys : y ∈ s) (H' : ∀₀ y' ∈ s, P y')]],
from eventually_nhds_elim eventuallyP,
from eventually_nhds_dest eventuallyP,
show eventually (λ x, P (f x)) F, from eventually_mono (H s Os ys) (λ x Hx, H' Hx))
proposition approaches_elim (H : (f ⟶ y) F) {s : set Y} (Os : Open s) (ys : y ∈ s) :
@ -669,6 +684,25 @@ section approaches
variables (F f y)
proposition approaches_iff : (f ⟶ y) F ↔ (∀ s, Open s → y ∈ s → eventually (λ x, f x ∈ s) F) :=
iff.intro approaches_elim approaches_intro
proposition tendsto_comp_of_approaches_of_tendsto_at_within
{f : X → Y} {g : Y → Z} {s : set Y} {y : Y} {F₁ : filter X} {F₃ : filter Z}
(Hf₁ : (f ⟶ y) F₁) (Hf₂ : eventually (λ x, f x ∈ s ∧ f x ≠ y) F₁)
(Hg : tendsto g F₃ [at y within s]) :
tendsto (g ∘ f) F₃ F₁ :=
have eventually (λ x, f x ∈ s \ '{y}) F₁,
from eventually_congr (take x, by rewrite [mem_diff_iff, mem_singleton_iff]) Hf₂,
have tendsto f [at y within s] F₁, from tendsto_inf Hf₁ (tendsto_principal this),
tendsto_comp this Hg
proposition tendsto_comp_of_approaches_of_tendsto_at
{f : X → Y} {g : Y → Z} {y : Y} {F₁ : filter X} {F₃ : filter Z}
(Hf₁ : (f ⟶ y) F₁) (Hf₂ : eventually (λ x, f x ≠ y) F₁)
(Hg : tendsto g F₃ [at y]) :
tendsto (g ∘ f) F₃ F₁ :=
have eventually (λ x, f x ∈ univ ∧ f x ≠ y) F₁,
from eventually_congr (take x, by rewrite [mem_univ_iff, true_and]) Hf₂,
tendsto_comp_of_approaches_of_tendsto_at_within Hf₁ this Hg
end approaches
/-