mathlib3 documentation

number_theory.liouville.liouville_number

Liouville constants #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This file contains a construction of a family of Liouville numbers, indexed by a natural number $m$. The most important property is that they are examples of transcendental real numbers. This fact is recorded in transcendental_liouville_number.

More precisely, for a real number $m$, Liouville's constant is $$ \sum_{i=0}^\infty\frac{1}{m^{i!}}. $$ The series converges only for $1 < m$. However, there is no restriction on $m$, since, if the series does not converge, then the sum of the series is defined to be zero.

We prove that, for $m \in \mathbb{N}$ satisfying $2 \le m$, Liouville's constant associated to $m$ is a transcendental number. Classically, the Liouville number for $m = 2$ is the one called ``Liouville's constant''.

Implementation notes #

The indexing $m$ is eventually a natural number satisfying $2 ≤ m$. However, we prove the first few lemmas for $m \in \mathbb{R}$.

noncomputable def liouville_number (m : ) :

For a real number m, Liouville's constant is $$ \sum_{i=0}^\infty\frac{1}{m^{i!}}. $$ The series converges only for 1 < m. However, there is no restriction on m, since, if the series does not converge, then the sum of the series is defined to be zero.

Equations
noncomputable def liouville_number.partial_sum (m : ) (k : ) :

liouville_number.partial_sum is the sum of the first k + 1 terms of Liouville's constant, i.e. $$ \sum_{i=0}^k\frac{1}{m^{i!}}. $$

Equations
noncomputable def liouville_number.remainder (m : ) (k : ) :

liouville_number.remainder is the sum of the series of the terms in liouville_number m starting from k+1, i.e $$ \sum_{i=k+1}^\infty\frac{1}{m^{i!}}. $$

Equations

We start with simple observations.

@[protected]
theorem liouville_number.summable {m : } (hm : 1 < m) :
summable (λ (i : ), 1 / m ^ i.factorial)
theorem liouville_number.remainder_summable {m : } (hm : 1 < m) (k : ) :
summable (λ (i : ), 1 / m ^ (i + (k + 1)).factorial)
theorem liouville_number.remainder_pos {m : } (hm : 1 < m) (k : ) :

Split the sum definining a Liouville number into the first k term and the rest.

We now prove two useful inequalities, before collecting everything together.

theorem liouville_number.remainder_lt' (n : ) {m : } (m1 : 1 < m) :
liouville_number.remainder m n < (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1).factorial)

An upper estimate on the remainder. This estimate works with m ∈ ℝ satisfying 1 < m and is stronger than the estimate liouville_number.remainder_lt below. However, the latter estimate is more useful for the proof.

theorem liouville_number.aux_calc (n : ) {m : } (hm : 2 m) :
(1 - 1 / m)⁻¹ * (1 / m ^ (n + 1).factorial) 1 / (m ^ n.factorial) ^ n
theorem liouville_number.remainder_lt (n : ) {m : } (m2 : 2 m) :

An upper estimate on the remainder. This estimate works with m ∈ ℝ satisfying 2 ≤ m and is weaker than the estimate liouville_number.remainder_lt' above. However, this estimate is more useful for the proof.

Starting from here, we specialize to the case in which m is a natural number.

The sum of the k initial terms of the Liouville number to base m is a ratio of natural numbers where the denominator is m ^ k!.