Documentation

Mathlib.Analysis.Real.OfDigits

Representation of reals in positional system #

This file defines Real.ofDigits and Real.digits functions which allows to work with the representations of reals as sequences of digits in positional system.

Main Definitions #

Main Statements #

noncomputable def Real.ofDigitsTerm {b : } (digits : Fin b) :

ofDigits takes a sequence of digits (d₀, d₁, ...) in base b and returns the real numnber 0.d₀d₁d₂... = ∑ᵢ(dᵢ/bⁱ). This auxiliary definition ofDigitsTerm sends the sequence to the function sending i to dᵢ/bⁱ.

Equations
Instances For
    theorem Real.ofDigitsTerm_nonneg {b : } {digits : Fin b} {n : } :
    0 ofDigitsTerm digits n
    theorem Real.ofDigitsTerm_le {b : } {digits : Fin b} {n : } :
    ofDigitsTerm digits n (b - 1) * (b ^ (n + 1))⁻¹
    theorem Real.summable_ofDigitsTerm {b : } {digits : Fin b} :
    noncomputable def Real.ofDigits {b : } (digits : Fin b) :

    ofDigits d is the real number 0.d₀d₁d₂... in base b. We allow repeating representations like 0.999... here.

    Equations
    Instances For
      theorem Real.ofDigits_nonneg {b : } (digits : Fin b) :
      0 ofDigits digits
      theorem Real.ofDigits_le_one {b : } (digits : Fin b) :
      ofDigits digits 1
      theorem Real.ofDigits_eq_sum_add_ofDigits {b : } (a : Fin b) (n : ) :
      ofDigits a = iFinset.range n, ofDigitsTerm a i + (b ^ n)⁻¹ * ofDigits fun (i : ) => a (i + n)
      theorem Real.abs_ofDigits_sub_ofDigits_le {b : } {x y : Fin b} {n : } (hxy : i < n, x i = y i) :
      noncomputable def Real.digits (x : ) (b : ) [NeZero b] :
      Fin b

      Converts a real number x from the interval [0, 1) into sequence of its digits in base b.

      Equations
      Instances For
        theorem Real.ofDigits_digits_sum_eq {x : } {b : } [NeZero b] (hx : x Set.Ico 0 1) (n : ) :
        b ^ n * iFinset.range n, ofDigitsTerm (x.digits b) i = b ^ n * x⌋₊
        theorem Real.le_sum_ofDigitsTerm_digits {x : } {b : } [NeZero b] (hb : 1 < b) (hx : x Set.Ico 0 1) (n : ) :
        x - (↑b)⁻¹ ^ n iFinset.range n, ofDigitsTerm (x.digits b) i
        theorem Real.sum_ofDigitsTerm_digits_le {x : } {b : } [NeZero b] (hx : x Set.Ico 0 1) (n : ) :
        iFinset.range n, ofDigitsTerm (x.digits b) i x
        theorem Real.hasSum_ofDigitsTerm_digits (x : ) {b : } [NeZero b] (hb : 1 < b) (hx : x Set.Ico 0 1) :
        theorem Real.ofDigits_digits {b : } [NeZero b] {x : } (hb : 1 < b) (hx : x Set.Ico 0 1) :
        ofDigits (x.digits b) = x