Library Stdlib.Numbers.HexadecimalR


HexadecimalR

Proofs that conversions between hexadecimal numbers and R are bijections.

From Stdlib Require Import PeanoNat.
From Stdlib Require Import Decimal DecimalFacts.
From Stdlib Require Import Hexadecimal HexadecimalFacts HexadecimalPos HexadecimalZ.
From Stdlib Require Import HexadecimalQ Rdefinitions.

Lemma of_IQmake_to_hexadecimal num den :
  match IQmake_to_hexadecimal num den with
  | None => True
  | Some (HexadecimalExp _ _ _) => False
  | Some (Hexadecimal i f) =>
    of_hexadecimal (Hexadecimal i f) = IRQ (QArith_base.Qmake num den)
  end.

Lemma of_to (q:IR) : forall d, to_hexadecimal q = Some d -> of_hexadecimal d = q.

Lemma to_of (d:hexadecimal) : to_hexadecimal (of_hexadecimal d) = Some (dnorm d).

Some consequences