Library IEEE
Require Export Bvector.
Require Export Zbinary.
Require Export AllFloat.
Section IEEEdefs.
Lemma Zpower_nat_is_pos :
forall (base : Z) (exp : nat),
(0 < base)%Z -> (0 < Zpower_nat base exp)%Z.
Let radix := 2%Z.
Record IEEE_precision : Set := Precision {precision : nat; nbexp : nat}.
Variable p : IEEE_precision.
Hypothesis Prec_Greater_Than_One : 1 < precision p.
Hypothesis Exp_Pos : 1 < nbexp p.
Hint Resolve Prec_Greater_Than_One: arith.
Hint Resolve Exp_Pos: arith.
Hint Resolve Zpower_nat_is_pos: zarith.
Record IEEE_Float : Set := IEEE_construct
{Sign : bool;
IEEE_Frac : vector bool (pred (precision p));
IEEE_Exp : vector bool (nbexp p)}.
Let exp_max := Zpred (Zpower_nat radix (pred (nbexp p))).
Let biais1 := exp_max.
Let biais2 := Zpred (precision p).
Let mant_max := (Zpower_nat radix (pred (precision p)) - 1%Z)%R.
Let b :=
Bound
(P_of_succ_nat (pred (Zabs_nat (Zpower_nat radix (precision p)))))
(Npos (P_of_succ_nat (pred (Zabs_nat (Zpred (Zpred exp_max + precision p)))))).
Definition IEEE_normal (f : IEEE_Float) :=
binary_value (nbexp p) (IEEE_Exp f) <> 0%Z /\
binary_value (nbexp p) (IEEE_Exp f) <> Zpred (Zpower_nat radix (nbexp p)).
Definition IEEE_subnormal (f : IEEE_Float) :=
binary_value (nbexp p) (IEEE_Exp f) = 0%Z /\
binary_value (pred (precision p)) (IEEE_Frac f) <> 0%Z.
Definition IEEE_canonic (f : IEEE_Float) := IEEE_normal f \/ IEEE_subnormal f.
Definition IEEE_shift_mantissa (f : float) :=
Float (radix * Fnum f) (Zpred (Fexp f)).
Lemma vNum_eq_Zpower : Zpos (vNum b) = Zpower_nat radix (precision p).
Theorem FtoR_Zpred_exp :
forall (r : Z) (f : float),
(1 < r)%Z -> FtoR r (Float (r * Fnum f) (Zpred (Fexp f))) = FtoR r f.
Theorem FNorm_or_Subnorm :
forall f : float,
Fbounded b f ->
{Fnormal radix b (Fnormalize radix b (precision p) f)} +
{Fsubnormal radix b (Fnormalize radix b (precision p) f)}.
Theorem Expo_zero_or_pos :
forall f : IEEE_Float,
{binary_value (nbexp p) (IEEE_Exp f) = 0%Z} +
{(0 < binary_value (nbexp p) (IEEE_Exp f))%Z}.
Definition IEEE_subnormalToR (x : IEEE_Float) :=
match Sign x with
| false =>
(binary_value (pred (precision p)) (IEEE_Frac x) *
powerRZ radix (Zsucc (- (biais1 + biais2))))%R
| true =>
((- binary_value (pred (precision p)) (IEEE_Frac x))%Z *
powerRZ radix (Zsucc (- (biais1 + biais2))))%R
end.
Definition IEEE_normalToR (x : IEEE_Float) :=
match Sign x with
| false =>
((Zpower_nat radix (pred (precision p)) +
binary_value (pred (precision p)) (IEEE_Frac x))%Z *
powerRZ radix
(binary_value (nbexp p) (IEEE_Exp x) - (biais1 + biais2)))%R
| true =>
(-
((Zpower_nat radix (pred (precision p)) +
binary_value (pred (precision p)) (IEEE_Frac x))%Z *
powerRZ radix
(binary_value (nbexp p) (IEEE_Exp x) - (biais1 + biais2))))%R
end.
Definition IEEE_To_R (x : IEEE_Float) :=
match Expo_zero_or_pos x with
| left _ => IEEE_subnormalToR x
| right _ => IEEE_normalToR x
end.
Definition InDomain (f : float) :=
(Fexp (Fnormalize radix b (precision p) f) + biais2 <= exp_max)%Z /\
Fbounded b f.
Definition FNormTo_IEEE (f : float) :=
match Z_lt_ge_dec (Fnum f) 0 with
| left _ =>
IEEE_construct true
(Z_to_binary (pred (precision p))
(- Fnum (Fnormalize radix b (precision p) f) -
Zpower_nat radix (pred (precision p))))
(Z_to_binary (nbexp p)
(biais1 + Fexp (Fnormalize radix b (precision p) f) + biais2))
| right _ =>
IEEE_construct false
(Z_to_binary (pred (precision p))
(Fnum (Fnormalize radix b (precision p) f) -
Zpower_nat radix (pred (precision p))))
(Z_to_binary (nbexp p)
(biais1 + Fexp (Fnormalize radix b (precision p) f) + biais2))
end.
Definition FSubnormTo_IEEE (f : float) :=
match Z_lt_ge_dec (Fnum f) 0 with
| left _ =>
IEEE_construct true
(Z_to_binary (pred (precision p))
(- Fnum (Fnormalize radix b (precision p) f)))
(Z_to_binary (nbexp p) 0)
| right _ =>
IEEE_construct false
(Z_to_binary (pred (precision p))
(Fnum (Fnormalize radix b (precision p) f)))
(Z_to_binary (nbexp p) 0)
end.
Definition Float_To_IEEE (f : float) H :=
match FNorm_or_Subnorm f H with
| left _ => FNormTo_IEEE f
| right _ => FSubnormTo_IEEE f
end.
Theorem float_equals_IEEE :
forall (f : float) (H : Fbounded b f),
InDomain f -> FtoR radix f = IEEE_To_R (Float_To_IEEE f H).
Theorem Float_To_IEEE_fn :
forall (f g : float) (Ff : Fbounded b f) (Fg : Fbounded b g),
FtoR radix f = FtoR radix g -> Float_To_IEEE f Ff = Float_To_IEEE g Fg.
Definition IEEE_Pos_To_Float (x : IEEE_Float) :=
match Expo_zero_or_pos x with
| left _ =>
Float (binary_value (pred (precision p)) (IEEE_Frac x)) (- dExp b)
| right _ =>
Float
(Zpower_nat radix (pred (precision p)) +
binary_value (pred (precision p)) (IEEE_Frac x))
(binary_value (nbexp p) (IEEE_Exp x) - (biais1 + biais2))
end.
Definition IEEE_To_Float (x : IEEE_Float) :=
match Sign x with
| false => IEEE_Pos_To_Float x
| true => Fopp (IEEE_Pos_To_Float x)
end.
Theorem binary_value_Zle :
forall (n : nat) (B : Bvector n),
(binary_value n B <= Zpred (two_power_nat n))%Z.
Theorem IEEE_To_Float_Bounded :
forall x : IEEE_Float, Fbounded b (IEEE_To_Float x).
Theorem IEEE_To_Float_To_IEEE :
forall x : IEEE_Float,
x <>
IEEE_construct true (Z_to_binary (pred (precision p)) 0)
(Z_to_binary (nbexp p) 0) ->
Float_To_IEEE (IEEE_To_Float x) (IEEE_To_Float_Bounded x) = x.
Theorem Float_To_IEEE_To_Float :
forall (f : float) H,
InDomain f ->
IEEE_To_Float (Float_To_IEEE f H) = Fnormalize radix b (precision p) f.
End IEEEdefs.