Documentation

Mathlib.Order.Filter.Archimedean

Filter.atTop filter and archimedean (semi)rings/fields #

In this file we prove that for a linear ordered archimedean semiring R and a function f : α → ℕ, the function Nat.cast ∘ f : α → R tends to Filter.atTop along a filter l if and only if so does f. We also prove that Nat.cast : ℕ → R tends to Filter.atTop along Filter.atTop, as well as version of these two results for (and a ring R) and (and a field R).

@[simp]
theorem Nat.comap_cast_atTop {R : Type u_2} [StrictOrderedSemiring R] [Archimedean R] :
Filter.comap Nat.cast Filter.atTop = Filter.atTop
theorem tendsto_natCast_atTop_iff {α : Type u_1} {R : Type u_2} [StrictOrderedSemiring R] [Archimedean R] {f : α} {l : Filter α} :
Filter.Tendsto (fun (n : α) => (f n)) l Filter.atTop Filter.Tendsto f l Filter.atTop
theorem tendsto_natCast_atTop_atTop {R : Type u_2} [OrderedSemiring R] [Archimedean R] :
Filter.Tendsto Nat.cast Filter.atTop Filter.atTop
theorem Filter.Eventually.natCast_atTop {R : Type u_2} [OrderedSemiring R] [Archimedean R] {p : RProp} (h : ∀ᶠ (x : R) in Filter.atTop, p x) :
∀ᶠ (n : ) in Filter.atTop, p n
@[simp]
theorem Int.comap_cast_atTop {R : Type u_2} [StrictOrderedRing R] [Archimedean R] :
Filter.comap Int.cast Filter.atTop = Filter.atTop
@[simp]
theorem Int.comap_cast_atBot {R : Type u_2} [StrictOrderedRing R] [Archimedean R] :
Filter.comap Int.cast Filter.atBot = Filter.atBot
theorem tendsto_intCast_atTop_iff {α : Type u_1} {R : Type u_2} [StrictOrderedRing R] [Archimedean R] {f : α} {l : Filter α} :
Filter.Tendsto (fun (n : α) => (f n)) l Filter.atTop Filter.Tendsto f l Filter.atTop
theorem tendsto_intCast_atBot_iff {α : Type u_1} {R : Type u_2} [StrictOrderedRing R] [Archimedean R] {f : α} {l : Filter α} :
Filter.Tendsto (fun (n : α) => (f n)) l Filter.atBot Filter.Tendsto f l Filter.atBot
theorem tendsto_intCast_atTop_atTop {R : Type u_2} [StrictOrderedRing R] [Archimedean R] :
Filter.Tendsto Int.cast Filter.atTop Filter.atTop
theorem Filter.Eventually.intCast_atTop {R : Type u_2} [StrictOrderedRing R] [Archimedean R] {p : RProp} (h : ∀ᶠ (x : R) in Filter.atTop, p x) :
∀ᶠ (n : ) in Filter.atTop, p n
theorem Filter.Eventually.intCast_atBot {R : Type u_2} [StrictOrderedRing R] [Archimedean R] {p : RProp} (h : ∀ᶠ (x : R) in Filter.atBot, p x) :
∀ᶠ (n : ) in Filter.atBot, p n
@[simp]
theorem Rat.comap_cast_atTop {R : Type u_2} [LinearOrderedField R] [Archimedean R] :
Filter.comap Rat.cast Filter.atTop = Filter.atTop
@[simp]
theorem Rat.comap_cast_atBot {R : Type u_2} [LinearOrderedField R] [Archimedean R] :
Filter.comap Rat.cast Filter.atBot = Filter.atBot
theorem tendsto_ratCast_atTop_iff {α : Type u_1} {R : Type u_2} [LinearOrderedField R] [Archimedean R] {f : α} {l : Filter α} :
Filter.Tendsto (fun (n : α) => (f n)) l Filter.atTop Filter.Tendsto f l Filter.atTop
theorem tendsto_ratCast_atBot_iff {α : Type u_1} {R : Type u_2} [LinearOrderedField R] [Archimedean R] {f : α} {l : Filter α} :
Filter.Tendsto (fun (n : α) => (f n)) l Filter.atBot Filter.Tendsto f l Filter.atBot
theorem Filter.Eventually.ratCast_atTop {R : Type u_2} [LinearOrderedField R] [Archimedean R] {p : RProp} (h : ∀ᶠ (x : R) in Filter.atTop, p x) :
∀ᶠ (n : ) in Filter.atTop, p n
theorem Filter.Eventually.ratCast_atBot {R : Type u_2} [LinearOrderedField R] [Archimedean R] {p : RProp} (h : ∀ᶠ (x : R) in Filter.atBot, p x) :
∀ᶠ (n : ) in Filter.atBot, p n
theorem atTop_hasCountableBasis_of_archimedean {R : Type u_2} [StrictOrderedSemiring R] [Archimedean R] :
Filter.HasCountableBasis Filter.atTop (fun (x : ) => True) fun (n : ) => Set.Ici n
theorem atBot_hasCountableBasis_of_archimedean {R : Type u_2} [LinearOrderedRing R] [Archimedean R] :
Filter.HasCountableBasis Filter.atBot (fun (x : ) => True) fun (m : ) => Set.Iic m
theorem Filter.Tendsto.const_mul_atTop' {α : Type u_1} {R : Type u_2} {l : Filter α} {f : αR} {r : R} [LinearOrderedSemiring R] [Archimedean R] (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => r * f x) l Filter.atTop

If a function tends to infinity along a filter, then this function multiplied by a positive constant (on the left) also tends to infinity. The archimedean assumption is convenient to get a statement that works on , and , although not necessary (a version in ordered fields is given in Filter.Tendsto.const_mul_atTop).

theorem Filter.Tendsto.atTop_mul_const' {α : Type u_1} {R : Type u_2} {l : Filter α} {f : αR} {r : R} [LinearOrderedSemiring R] [Archimedean R] (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x * r) l Filter.atTop

If a function tends to infinity along a filter, then this function multiplied by a positive constant (on the right) also tends to infinity. The archimedean assumption is convenient to get a statement that works on , and , although not necessary (a version in ordered fields is given in Filter.Tendsto.atTop_mul_const).

theorem Filter.Tendsto.atTop_mul_neg_const' {α : Type u_1} {R : Type u_2} {l : Filter α} {f : αR} {r : R} [LinearOrderedRing R] [Archimedean R] (hr : r < 0) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x * r) l Filter.atBot

See also Filter.Tendsto.atTop_mul_neg_const for a version of this lemma for LinearOrderedFields which does not require the Archimedean assumption.

theorem Filter.Tendsto.atBot_mul_const' {α : Type u_1} {R : Type u_2} {l : Filter α} {f : αR} {r : R} [LinearOrderedRing R] [Archimedean R] (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atBot) :
Filter.Tendsto (fun (x : α) => f x * r) l Filter.atBot

See also Filter.Tendsto.atBot_mul_const for a version of this lemma for LinearOrderedFields which does not require the Archimedean assumption.

theorem Filter.Tendsto.atBot_mul_neg_const' {α : Type u_1} {R : Type u_2} {l : Filter α} {f : αR} {r : R} [LinearOrderedRing R] [Archimedean R] (hr : r < 0) (hf : Filter.Tendsto f l Filter.atBot) :
Filter.Tendsto (fun (x : α) => f x * r) l Filter.atTop

See also Filter.Tendsto.atBot_mul_neg_const for a version of this lemma for LinearOrderedFields which does not require the Archimedean assumption.

theorem Filter.Tendsto.atTop_nsmul_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedCancelAddCommMonoid R] [Archimedean R] {f : α} (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atTop
theorem Filter.Tendsto.atTop_nsmul_neg_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedAddCommGroup R] [Archimedean R] {f : α} (hr : r < 0) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atBot
theorem Filter.Tendsto.atTop_zsmul_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedAddCommGroup R] [Archimedean R] {f : α} (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atTop
theorem Filter.Tendsto.atTop_zsmul_neg_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedAddCommGroup R] [Archimedean R] {f : α} (hr : r < 0) (hf : Filter.Tendsto f l Filter.atTop) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atBot
theorem Filter.Tendsto.atBot_zsmul_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedAddCommGroup R] [Archimedean R] {f : α} (hr : 0 < r) (hf : Filter.Tendsto f l Filter.atBot) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atBot
theorem Filter.Tendsto.atBot_zsmul_neg_const {α : Type u_1} {R : Type u_2} {l : Filter α} {r : R} [LinearOrderedAddCommGroup R] [Archimedean R] {f : α} (hr : r < 0) (hf : Filter.Tendsto f l Filter.atBot) :
Filter.Tendsto (fun (x : α) => f x r) l Filter.atTop