Library Coq.Numbers.DecimalQ
Require Import Decimal DecimalFacts DecimalPos DecimalN DecimalZ QArith.
Lemma of_IQmake_to_decimal num den :
match IQmake_to_decimal num den with
| None => True
| Some (DecimalExp _ _ _) => False
| Some (Decimal i f) => of_decimal (Decimal i f) = IQmake (IZ_of_Z num) den
end.
Lemma IZ_of_Z_IZ_to_Z z z' : IZ_to_Z z = Some z' -> IZ_of_Z z' = z.
Lemma of_IQmake_to_decimal' num den :
match IQmake_to_decimal' num den with
| None => True
| Some (DecimalExp _ _ _) => False
| Some (Decimal i f) => of_decimal (Decimal i f) = IQmake num den
end.
Lemma of_to (q:IQ) : forall d, to_decimal q = Some d -> of_decimal d = q.
Definition dnorm (d:decimal) : decimal :=
let norm_i i f :=
match i with
| Pos i => Pos (unorm i)
| Neg i => match nzhead (app i f) with Nil => Pos zero | _ => Neg (unorm i) end
end in
match d with
| Decimal i f => Decimal (norm_i i f) f
| DecimalExp i f e =>
match norm e with
| Pos zero => Decimal (norm_i i f) f
| e => DecimalExp (norm_i i f) f e
end
end.
Lemma dnorm_spec_i d :
let (i, f) :=
match d with Decimal i f => (i, f) | DecimalExp i f _ => (i, f) end in
let i' := match dnorm d with Decimal i _ => i | DecimalExp i _ _ => i end in
match i with
| Pos i => i' = Pos (unorm i)
| Neg i =>
(i' = Neg (unorm i) /\ (nzhead i <> Nil \/ nzhead f <> Nil))
\/ (i' = Pos zero /\ (nzhead i = Nil /\ nzhead f = Nil))
end.
Lemma dnorm_spec_f d :
let f := match d with Decimal _ f => f | DecimalExp _ f _ => f end in
let f' := match dnorm d with Decimal _ f => f | DecimalExp _ f _ => f end in
f' = f.
Lemma dnorm_spec_e d :
match d, dnorm d with
| Decimal _ _, Decimal _ _ => True
| DecimalExp _ _ e, Decimal _ _ => norm e = Pos zero
| DecimalExp _ _ e, DecimalExp _ _ e' => e' = norm e /\ e' <> Pos zero
| Decimal _ _, DecimalExp _ _ _ => False
end.
Lemma dnorm_involutive d : dnorm (dnorm d) = dnorm d.
Lemma IZ_to_Z_IZ_of_Z z : IZ_to_Z (IZ_of_Z z) = Some z.
Lemma dnorm_i_exact i f :
(nb_digits f < nb_digits (unorm (app (abs i) f)))%nat ->
match i with
| Pos i => Pos (unorm i)
| Neg i =>
match nzhead (app i f) with
| Nil => Pos zero
| _ => Neg (unorm i)
end
end = norm i.
Lemma dnorm_i_exact' i f :
(nb_digits (unorm (app (abs i) f)) <= nb_digits f)%nat ->
match i with
| Pos i => Pos (unorm i)
| Neg i =>
match nzhead (app i f) with
| Nil => Pos zero
| _ => Neg (unorm i)
end
end =
match norm (app_int i f) with
| Pos _ => Pos zero
| Neg _ => Neg zero
end.
Lemma to_of (d:decimal) : to_decimal (of_decimal d) = Some (dnorm d).
Some consequences
Lemma to_decimal_inj q q' :
to_decimal q <> None -> to_decimal q = to_decimal q' -> q = q'.
Lemma to_decimal_surj d : exists q, to_decimal q = Some (dnorm d).
Lemma of_decimal_dnorm d : of_decimal (dnorm d) = of_decimal d.
Lemma of_inj d d' : of_decimal d = of_decimal d' -> dnorm d = dnorm d'.
Lemma of_iff d d' : of_decimal d = of_decimal d' <-> dnorm d = dnorm d'.