Library FboundI

Require Export AllFloat.

Section FboundedI_Def.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.

Record FboundI : Set := BoundI {vNumInf : nat; vNumSup : nat; dExpI : nat}.

Definition FboundedI (b : FboundI) (d : float) :=
  ((- vNumInf b <= Fnum d)%Z /\ (Fnum d <= vNumSup b)%Z) /\
  (- dExpI b <= Fexp d)%Z.

Theorem isBoundedI :
 forall (b : FboundI) (p : float), {FboundedI b p} + {~ FboundedI b p}.

Theorem FboundedIFzero : forall b : FboundI, FboundedI b (Fzero (- dExpI b)).

Definition FnormalI (b : FboundI) (p : float) :=
  FboundedI b p /\
  ((vNumSup b < radix * Fnum p)%R \/ (radix * Fnum p < (- vNumInf b)%Z)%R).

Theorem FnormalINotZero :
 forall (b : FboundI) (p : float), FnormalI b p -> ~ is_Fzero p.

Theorem FnormalIUnique_aux :
 forall (b : FboundI) (p q : float),
 FnormalI b p ->
 FnormalI b q -> p = q :>R -> (Fexp p < Fexp q)%Z -> p = q :>float.

Theorem FnormalIUnique :
 forall (b : FboundI) (p q : float),
 FnormalI b p -> FnormalI b q -> p = q :>R -> p = q :>float.

Definition FsubnormalI (b : FboundI) (p : float) :=
  FboundedI b p /\
  Fexp p = (- dExpI b)%Z /\
  ((- vNumInf b)%Z <= radix * Fnum p)%R /\ (radix * Fnum p <= vNumSup b)%R.

Theorem FsubnormalIUnique :
 forall (b : FboundI) (p q : float),
 FsubnormalI b p -> FsubnormalI b q -> p = q :>R -> p = q :>float.

Theorem NormalIandSubnormalINotEq :
 forall (b : FboundI) (p q : float),
 FnormalI b p -> FsubnormalI b q -> p <> q :>R.

Definition FcanonicI (b : FboundI) (a : float) :=
  FnormalI b a \/ FsubnormalI b a.

Theorem FcanonicIUnique :
 forall (b : FboundI) (p q : float),
 FcanonicI b p -> FcanonicI b q -> p = q :>R -> p = q :>float.

Theorem Zpower_nat_S :
 forall n : nat, Zpower_nat radix (S n) = (radix * Zpower_nat radix n)%Z.

Fixpoint FNIAux (v N q : nat) {struct q} : nat :=
  match q with
  | O => 0
  | S q' =>
      match (Zpower_nat radix q' * v ?= radix * N)%Z with
      | Datatypes.Lt => q'
      | Datatypes.Eq => q'
      | _ => FNIAux v N q'
      end
  end.

Definition FNI (q N : nat) := pred (FNIAux q N (S (S N))).

Theorem FNIAuxLess :
 forall v N q : nat,
 0 < v -> v <= N -> (Zpower_nat radix (FNIAux v N q) * v <= radix * N)%Z.

Theorem FNILess :
 forall q N : nat, 0 < q -> q <= N -> (Zpower_nat radix (FNI q N) * q <= N)%Z.

Theorem FNIAuxMore :
 forall v N q : nat,
 FNIAux v N q < pred q -> (N < Zpower_nat radix (FNIAux v N q) * v)%Z.

Theorem Zlt_Zpower_nat : forall n : nat, (n < Zpower_nat radix n)%Z.

Theorem FNIMore :
 forall N q : nat,
 0 < q -> q <= N -> (N < radix * (Zpower_nat radix (FNI q N) * q))%Z.

Definition FnormalizeI (b : FboundI) (p : float) :=
  match (0 ?= Fnum p)%Z with
  | Datatypes.Eq => Float 0 (- dExpI b)
  | Datatypes.Lt =>
      Fshift radix
        (min (FNI (Zabs_nat (Fnum p)) (vNumSup b))
           (Zabs_nat (Fexp p + dExpI b))) p
  | Datatypes.Gt =>
      Fshift radix
        (min (FNI (Zabs_nat (Fnum p)) (vNumInf b))
           (Zabs_nat (Fexp p + dExpI b))) p
  end.

Theorem FnormalizeICorrect :
 forall (b : FboundI) (p : float), FnormalizeI b p = p :>R.

Theorem FnormalizeIBounded :
 forall (b : FboundI) (p : float),
 FboundedI b p -> FboundedI b (FnormalizeI b p).

Theorem FcanonicIBoundedI :
 forall (b : FboundI) (p : float), FcanonicI b p -> FboundedI b p.

Theorem FnormalizeIFcanonicI :
 forall (b : FboundI) (p : float),
 FboundedI b p -> FcanonicI b (FnormalizeI b p).

End FboundedI_Def.