mathlib documentation

data.​real.​cau_seq

data.​real.​cau_seq

Cauchy sequences

A basic theory of Cauchy sequences, used in the construction of the reals and p-adic numbers. Where applicable, lemmas that will be reused in other contexts have been stated in extra generality.

There are other "versions" of Cauchyness in the library, in particular Cauchy filters in topology. This is a concrete implementation that is useful for simplicity and computability reasons.

Important definitions

Tags

sequence, cauchy, abs val, absolute value

@[class]
structure is_absolute_value {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] :
(β → α) → Prop
  • abv_nonneg : ∀ (x : β), 0 f x
  • abv_eq_zero : ∀ {x : β}, f x = 0 x = 0
  • abv_add : ∀ (x y : β), f (x + y) f x + f y
  • abv_mul : ∀ (x y : β), f (x * y) = f x * f y

A function f is an absolute value if it is nonnegative, zero only at 0, additive, and multiplicative.

Instances
theorem is_absolute_value.​abv_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] :
abv 0 = 0

theorem is_absolute_value.​abv_one' {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] :
1 0abv 1 = 1

theorem is_absolute_value.​abv_one {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [domain β] (abv : β → α) [is_absolute_value abv] :
abv 1 = 1

theorem is_absolute_value.​abv_pos {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] {a : β} :
0 < abv a a 0

theorem is_absolute_value.​abv_neg {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] (a : β) :
abv (-a) = abv a

theorem is_absolute_value.​abv_sub {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] (a b : β) :
abv (a - b) = abv (b - a)

theorem is_absolute_value.​abv_inv {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] (abv : β → α) [is_absolute_value abv] (a : β) :
abv a⁻¹ = (abv a)⁻¹

theorem is_absolute_value.​abv_div {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] (abv : β → α) [is_absolute_value abv] (a b : β) :
abv (a / b) = abv a / abv b

theorem is_absolute_value.​abv_sub_le {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] (a b c : β) :
abv (a - c) abv (a - b) + abv (b - c)

theorem is_absolute_value.​sub_abv_le_abv_sub {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] (a b : β) :
abv a - abv b abv (a - b)

theorem is_absolute_value.​abs_abv_sub_le_abv_sub {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] (a b : β) :
abs (abv a - abv b) abv (a - b)

theorem is_absolute_value.​abv_pow {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [domain β] (abv : β → α) [is_absolute_value abv] (a : β) (n : ) :
abv (a ^ n) = abv a ^ n

theorem exists_forall_ge_and {α : Type u_1} [linear_order α] {P Q : α → Prop} :
(∃ (i : α), ∀ (j : α), j iP j)(∃ (i : α), ∀ (j : α), j iQ j)(∃ (i : α), ∀ (j : α), j iP j Q j)

theorem rat_add_continuous_lemma {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] {ε : α} :
0 < ε(∃ (δ : α) (H : δ > 0), ∀ {a₁ a₂ b₁ b₂ : β}, abv (a₁ - b₁) < δabv (a₂ - b₂) < δabv (a₁ + a₂ - (b₁ + b₂)) < ε)

theorem rat_mul_continuous_lemma {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] {ε K₁ K₂ : α} :
0 < ε(∃ (δ : α) (H : δ > 0), ∀ {a₁ a₂ b₁ b₂ : β}, abv a₁ < K₁abv b₂ < K₂abv (a₁ - b₁) < δabv (a₂ - b₂) < δabv (a₁ * a₂ - b₁ * b₂) < ε)

theorem rat_inv_continuous_lemma {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] (abv : β → α) [is_absolute_value abv] {ε K : α} :
0 < ε0 < K(∃ (δ : α) (H : δ > 0), ∀ {a b : β}, K abv aK abv babv (a - b) < δabv (a⁻¹ - b⁻¹) < ε)

def is_cau_seq {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] :
(β → α)( → β) → Prop

A sequence is Cauchy if the distance between its entries tends to zero.

Equations
@[nolint]
theorem is_cau_seq.​cauchy₂ {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : → β} (hf : is_cau_seq abv f) {ε : α} :
0 < ε(∃ (i : ), ∀ (j k : ), j ik iabv (f j - f k) < ε)

theorem is_cau_seq.​cauchy₃ {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : → β} (hf : is_cau_seq abv f) {ε : α} :
0 < ε(∃ (i : ), ∀ (j : ), j i∀ (k : ), k jabv (f k - f j) < ε)

def cau_seq {α : Type u_1} [discrete_linear_ordered_field α] (β : Type u_2) [ring β] :
(β → α)Type u_2

cau_seq β abv is the type of β-valued Cauchy sequences, with respect to the absolute value function abv.

Equations
@[instance]
def cau_seq.​has_coe_to_fun {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} :

Equations
@[simp]
theorem cau_seq.​mk_to_fun {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} (f : → β) (hf : is_cau_seq abv f) :
f, hf⟩ = f

theorem cau_seq.​ext {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} {f g : cau_seq β abv} :
(∀ (i : ), f i = g i)f = g

theorem cau_seq.​is_cau {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} (f : cau_seq β abv) :

theorem cau_seq.​cauchy {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} (f : cau_seq β abv) {ε : α} :
0 < ε(∃ (i : ), ∀ (j : ), j iabv (f j - f i) < ε)

def cau_seq.​of_eq {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} (f : cau_seq β abv) (g : → β) :
(∀ (i : ), f i = g i)cau_seq β abv

Given a Cauchy sequence f, create a Cauchy sequence from a sequence g with the same values as f.

Equations
@[nolint]
theorem cau_seq.​cauchy₂ {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) {ε : α} :
0 < ε(∃ (i : ), ∀ (j k : ), j ik iabv (f j - f k) < ε)

theorem cau_seq.​cauchy₃ {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) {ε : α} :
0 < ε(∃ (i : ), ∀ (j : ), j i∀ (k : ), k jabv (f k - f j) < ε)

theorem cau_seq.​bounded {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) :
∃ (r : α), ∀ (i : ), abv (f i) < r

theorem cau_seq.​bounded' {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) (x : α) :
∃ (r : α) (H : r > x), ∀ (i : ), abv (f i) < r

@[instance]
def cau_seq.​has_add {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
has_add (cau_seq β abv)

Equations
@[simp]
theorem cau_seq.​add_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f g : cau_seq β abv) (i : ) :
(f + g) i = f i + g i

def cau_seq.​const {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] (abv : β → α) [is_absolute_value abv] :
β → cau_seq β abv

The constant Cauchy sequence.

Equations
@[simp]
theorem cau_seq.​const_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (x : β) (i : ) :
(cau_seq.const abv x) i = x

theorem cau_seq.​const_inj {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {x y : β} :

@[instance]
def cau_seq.​has_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
has_zero (cau_seq β abv)

Equations
@[instance]
def cau_seq.​has_one {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
has_one (cau_seq β abv)

Equations
@[instance]
def cau_seq.​inhabited {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :

Equations
@[simp]
theorem cau_seq.​zero_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (i : ) :
0 i = 0

@[simp]
theorem cau_seq.​one_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (i : ) :
1 i = 1

theorem cau_seq.​const_add {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (x y : β) :

@[instance]
def cau_seq.​has_mul {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
has_mul (cau_seq β abv)

Equations
@[simp]
theorem cau_seq.​mul_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f g : cau_seq β abv) (i : ) :
(f * g) i = f i * g i

theorem cau_seq.​const_mul {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (x y : β) :

@[instance]
def cau_seq.​has_neg {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
has_neg (cau_seq β abv)

Equations
@[simp]
theorem cau_seq.​neg_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) (i : ) :
(-f) i = -f i

theorem cau_seq.​const_neg {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (x : β) :

@[instance]
def cau_seq.​ring {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
ring (cau_seq β abv)

Equations
@[instance]
def cau_seq.​comm_ring {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [comm_ring β] {abv : β → α} [is_absolute_value abv] :

Equations
theorem cau_seq.​const_sub {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (x y : β) :

@[simp]
theorem cau_seq.​sub_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f g : cau_seq β abv) (i : ) :
(f - g) i = f i - g i

def cau_seq.​lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} :
cau_seq β abv → Prop

lim_zero f holds when f approaches 0.

Equations
theorem cau_seq.​add_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} :
f.lim_zerog.lim_zero(f + g).lim_zero

theorem cau_seq.​mul_lim_zero_right {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) {g : cau_seq β abv} :
g.lim_zero(f * g).lim_zero

theorem cau_seq.​mul_lim_zero_left {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} (g : cau_seq β abv) :
f.lim_zero(f * g).lim_zero

theorem cau_seq.​neg_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} :

theorem cau_seq.​sub_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} :
f.lim_zerog.lim_zero(f - g).lim_zero

theorem cau_seq.​lim_zero_sub_rev {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} :
(f - g).lim_zero(g - f).lim_zero

theorem cau_seq.​zero_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :

theorem cau_seq.​const_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {x : β} :

@[instance]
def cau_seq.​equiv {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] :
setoid (cau_seq β abv)

Equations
theorem cau_seq.​equiv_def₃ {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} (h : f g) {ε : α} :
0 < ε(∃ (i : ), ∀ (j : ), j i∀ (k : ), k jabv (f k - g j) < ε)

theorem cau_seq.​lim_zero_congr {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} :
f g(f.lim_zero g.lim_zero)

theorem cau_seq.​abv_pos_of_not_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} :
¬f.lim_zero(∃ (K : α) (H : K > 0) (i : ), ∀ (j : ), j iK abv (f j))

theorem cau_seq.​of_near {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (f : → β) (g : cau_seq β abv) :
(∀ (ε : α), ε > 0(∃ (i : ), ∀ (j : ), j iabv (f j - g j) < ε))is_cau_seq abv f

theorem cau_seq.​not_lim_zero_of_not_congr_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} :
¬f 0¬f.lim_zero

theorem cau_seq.​mul_equiv_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] (g : cau_seq β abv) {f : cau_seq β abv} :
f 0g * f 0

theorem cau_seq.​mul_not_equiv_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {f g : cau_seq β abv} :
¬f 0¬g 0¬f * g 0

theorem cau_seq.​const_equiv {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [ring β] {abv : β → α} [is_absolute_value abv] {x y : β} :

theorem cau_seq.​mul_equiv_zero' {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [comm_ring β] {abv : β → α} [is_absolute_value abv] (g : cau_seq β abv) {f : cau_seq β abv} :
f 0f * g 0

theorem cau_seq.​one_not_equiv_zero {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [integral_domain β] (abv : β → α) [is_absolute_value abv] :

theorem cau_seq.​inv_aux {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} (hf : ¬f.lim_zero) (ε : α) :
ε > 0(∃ (i : ), ∀ (j : ), j iabv ((f j)⁻¹ - (f i)⁻¹) < ε)

def cau_seq.​inv {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] {abv : β → α} [is_absolute_value abv] (f : cau_seq β abv) :
¬f.lim_zerocau_seq β abv

Given a Cauchy sequence f with nonzero limit, create a Cauchy sequence with values equal to the inverses of the values of f.

Equations
@[simp]
theorem cau_seq.​inv_apply {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} (hf : ¬f.lim_zero) (i : ) :
(f.inv hf) i = (f i)⁻¹

theorem cau_seq.​inv_mul_cancel {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] {abv : β → α} [is_absolute_value abv] {f : cau_seq β abv} (hf : ¬f.lim_zero) :
f.inv hf * f 1

theorem cau_seq.​const_inv {α : Type u_1} [discrete_linear_ordered_field α] {β : Type u_2} [field β] {abv : β → α} [is_absolute_value abv] {x : β} (hx : x 0) :

def cau_seq.​pos {α : Type u_1} [discrete_linear_ordered_field α] :
cau_seq α abs → Prop

The entries of a positive Cauchy sequence eventually have a positive lower bound.

Equations
theorem cau_seq.​const_pos {α : Type u_1} [discrete_linear_ordered_field α] {x : α} :

theorem cau_seq.​add_pos {α : Type u_1} [discrete_linear_ordered_field α] {f g : cau_seq α abs} :
f.posg.pos(f + g).pos

theorem cau_seq.​pos_add_lim_zero {α : Type u_1} [discrete_linear_ordered_field α] {f g : cau_seq α abs} :
f.posg.lim_zero(f + g).pos

theorem cau_seq.​mul_pos {α : Type u_1} [discrete_linear_ordered_field α] {f g : cau_seq α abs} :
f.posg.pos(f * g).pos

theorem cau_seq.​trichotomy {α : Type u_1} [discrete_linear_ordered_field α] (f : cau_seq α abs) :

@[instance]

Equations
@[instance]

Equations
theorem cau_seq.​lt_of_lt_of_eq {α : Type u_1} [discrete_linear_ordered_field α] {f g h : cau_seq α abs} :
f < gg hf < h

theorem cau_seq.​lt_of_eq_of_lt {α : Type u_1} [discrete_linear_ordered_field α] {f g h : cau_seq α abs} :
f gg < hf < h

theorem cau_seq.​lt_trans {α : Type u_1} [discrete_linear_ordered_field α] {f g h : cau_seq α abs} :
f < gg < hf < h

theorem cau_seq.​lt_irrefl {α : Type u_1} [discrete_linear_ordered_field α] {f : cau_seq α abs} :
¬f < f

theorem cau_seq.​le_of_eq_of_le {α : Type u_1} [discrete_linear_ordered_field α] {f g h : cau_seq α abs} :
f gg hf h

theorem cau_seq.​le_of_le_of_eq {α : Type u_1} [discrete_linear_ordered_field α] {f g h : cau_seq α abs} :
f gg hf h

@[instance]

Equations
theorem cau_seq.​le_antisymm {α : Type u_1} [discrete_linear_ordered_field α] {f g : cau_seq α abs} :
f gg ff g

theorem cau_seq.​lt_total {α : Type u_1} [discrete_linear_ordered_field α] (f g : cau_seq α abs) :
f < g f g g < f

theorem cau_seq.​le_total {α : Type u_1} [discrete_linear_ordered_field α] (f g : cau_seq α abs) :
f g g f

theorem cau_seq.​const_lt {α : Type u_1} [discrete_linear_ordered_field α] {x y : α} :

theorem cau_seq.​le_of_exists {α : Type u_1} [discrete_linear_ordered_field α] {f g : cau_seq α abs} :
(∃ (i : ), ∀ (j : ), j if j g j)f g

theorem cau_seq.​exists_gt {α : Type u_1} [discrete_linear_ordered_field α] (f : cau_seq α abs) :
∃ (a : α), f < cau_seq.const abs a

theorem cau_seq.​exists_lt {α : Type u_1} [discrete_linear_ordered_field α] (f : cau_seq α abs) :
∃ (a : α), cau_seq.const abs a < f