Library Coq.Numbers.DecimalNat


DecimalNat

Proofs that conversions between decimal numbers and nat are bijections.

Require Import Decimal DecimalFacts Arith.

Module Unsigned.

A few helper functions used during proofs

Definition hd d :=
 match d with
 | Nil => 0
 | D0 _ => 0
 | D1 _ => 1
 | D2 _ => 2
 | D3 _ => 3
 | D4 _ => 4
 | D5 _ => 5
 | D6 _ => 6
 | D7 _ => 7
 | D8 _ => 8
 | D9 _ => 9
end.

Definition tl d :=
 match d with
 | Nil => d
 | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d => d
end.

Fixpoint usize (d:uint) : nat :=
  match d with
  | Nil => 0
  | D0 d => S (usize d)
  | D1 d => S (usize d)
  | D2 d => S (usize d)
  | D3 d => S (usize d)
  | D4 d => S (usize d)
  | D5 d => S (usize d)
  | D6 d => S (usize d)
  | D7 d => S (usize d)
  | D8 d => S (usize d)
  | D9 d => S (usize d)
  end.

A direct version of to_little_uint, not tail-recursive
Fixpoint to_lu n :=
 match n with
 | 0 => Decimal.zero
 | S n => Little.succ (to_lu n)
 end.

A direct version of of_little_uint
Fixpoint of_lu (d:uint) : nat :=
  match d with
  | Nil => 0
  | D0 d => 10 * of_lu d
  | D1 d => 1 + 10 * of_lu d
  | D2 d => 2 + 10 * of_lu d
  | D3 d => 3 + 10 * of_lu d
  | D4 d => 4 + 10 * of_lu d
  | D5 d => 5 + 10 * of_lu d
  | D6 d => 6 + 10 * of_lu d
  | D7 d => 7 + 10 * of_lu d
  | D8 d => 8 + 10 * of_lu d
  | D9 d => 9 + 10 * of_lu d
  end.

Properties of to_lu
Properties of of_lu

Lemma of_lu_eqn d :
 of_lu d = hd d + 10 * of_lu (tl d).

Ltac simpl_of_lu :=
 match goal with
 | |- context [ of_lu (?f ?x) ] =>
   rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
 end.

Lemma of_lu_succ d :
 of_lu (Little.succ d) = S (of_lu d).

Lemma of_to_lu n :
 of_lu (to_lu n) = n.

Lemma of_lu_revapp d d' :
of_lu (revapp d d') =
 of_lu (rev d) + of_lu d' * 10^usize d.

Lemma of_uint_acc_spec n d :
 Nat.of_uint_acc d n = of_lu (rev d) + n * 10^usize d.

Lemma of_uint_alt d : Nat.of_uint d = of_lu (rev d).

First main bijection result

Lemma of_to (n:nat) : Nat.of_uint (Nat.to_uint n) = n.

The other direction

Lemma to_lu_tenfold n : n<>0 ->
 to_lu (10 * n) = D0 (to_lu n).

Lemma of_lu_0 d : of_lu d = 0 <-> nztail d = Nil.

Lemma to_of_lu_tenfold d :
 to_lu (of_lu d) = lnorm d ->
 to_lu (10 * of_lu d) = lnorm (D0 d).

Lemma to_of_lu d : to_lu (of_lu d) = lnorm d.

Second bijection result
Some consequences
Conversion from/to signed decimal numbers