Library FexpAdd

Require Export Fexp2.
Section FexpAdd.
Variable b : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let TMTO : (1 < radix)%Z := TwoMoreThanOne.
Hint Resolve TMTO: zarith.

Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.
Hypothesis Ngd : (1 <= pPred (vNum b) * (1 - / radix))%R.
Hypothesis Ngd2 : (6%nat <= pPred (vNum b) * (1 - / radix * / radix))%R.

Inductive NearEqual : list float -> list float -> Prop :=
  | IsEqual : forall x : list float, NearEqual x x
  | OneMoreR :
      forall (x : list float) (e : float),
      Fbounded b e -> NearEqual x (e :: x).
Hint Resolve IsEqual OneMoreR.

Definition Step (x y i x' y' : float) (input output output' : list float) :=
  Fbounded b x' /\
  Fbounded b y' /\
  NearEqual output output' /\
  (x + (y + (sum (i :: input) + sum output)))%R =
  (x' + (y' + (sum input + sum output')))%R :>R /\
  (Fexp i <= Fexp y')%Z /\
  (Fexp y' <= Fexp x')%Z /\
  IsExp b (i :: input) /\
  (x' = 0%R :>R \/
   y' = 0%R :>R \/
   (Rabs y' <=
    pPred (vNum b) * (Float 1%nat (Fexp x') - Float 1%nat (hdexp b input)))%R /\
   (Rabs y' <=
    pPred (vNum b) * (Float 1%nat (Fexp x') - Float 1%nat (Fexp i)))%R) /\
  (output' = output \/
   (Rabs (x' + y') <= 3%nat * radix * / pPred (vNum b) * Rabs (hd b output'))%R) /\
  (output' = nil \/
   (Rabs (hd b input) <=
    3%nat * (radix * (Rabs (hd b output') * / (pPred (vNum b) - 1%nat))))%R /\
   (input = nil \/
    (Float (pPred (vNum b)) (Fexp (hd b input)) <=
     3%nat * (radix * (Rabs (hd b output') * / (pPred (vNum b) - 1%nat))))%R /\
    (Rabs (sum input) <=
     length input *
     (3%nat * (radix * (Rabs (hd b output') * / (pPred (vNum b) - 1%nat)))))%R)).

Theorem Rle_mult_pos :
 forall r1 r2 : R, (0 <= r1)%R -> (0 <= r2)%R -> (0 <= r1 * r2)%R.

Theorem AddStep :
 forall (x y i : float) (input output : list float),
 Fbounded b x ->
 Fbounded b y ->
 IsExp b (i :: input) ->
 (Fexp i <= Fexp y)%Z ->
 (Fexp y <= Fexp x)%Z ->
 (FtoR radix x = 0%R :>R \/ FtoR radix y = 0%R :>R \/ FtoR radix i = 0%R :>R) \/
 (Rabs (FtoR radix y) <=
  pPred (vNum b) *
  (FtoR radix (Float 1%nat (Fexp x)) - FtoR radix (Float 1%nat (Fexp i))))%R ->
 output = nil \/
 (Rabs i <=
  3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat))))%R /\
 (Float (pPred (vNum b)) (Fexp i) <=
  3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat))))%R /\
 (Rabs (sum (i :: input)) <=
  length (i :: input) *
  (3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat)))))%R ->
 exists x' : float,
   (exists y' : float,
      (exists output' : list float, Step x y i x' y' input output output')).
Variable input : list float.

Inductive IsRleEpsExp : list float -> Prop :=
  | IsRleEpsExpNil : IsRleEpsExp nil
  | IsRleEpsExpSingle :
      forall x : float, Fbounded b x -> IsRleEpsExp (x :: nil)
  | IsRleEpsExpTop :
      forall (x y : float) (L : list float),
      Fbounded b x ->
      Fbounded b y ->
      (Rabs x <=
       (6%nat * length input + 6%nat) *
       / (pPred (vNum b) - 1%nat - 6%nat * length input) *
       Rabs y)%R -> IsRleEpsExp (y :: L) -> IsRleEpsExp (x :: y :: L).

Theorem endof_Rle_length :
 forall (P Q : list float) (k : float),
 endof input (P ++ k :: Q) -> (length P <= length input - 1)%R.

Theorem FexpAdd_aux :
 (6%nat * length input * / (pPred (vNum b) - 1%nat) < 1)%R ->
 forall (L output : list float) (x y : float),
 endof input L ->
 IsExp b L ->
 IsRleEpsExp output ->
 output = nil \/
 (Rabs (x + y) <=
  (pPred (vNum b) - 1%nat) * / pPred (vNum b) *
  ((6%nat * length input + 6%nat) *
   / (pPred (vNum b) - 1%nat - 6%nat * length input)) *
  Rabs (hd b output))%R ->
 output = nil \/
 (exists L1 : list float,
    (exists x1 : float,
       (exists y1 : float,
          IsExp b (L1 ++ L) /\
          endof input (L1 ++ L) /\
          (x1 + (y1 + sum L1))%R = (x + y)%R :>R /\
          (Rabs (sum L1) <=
           length L1 *
           (3%nat *
            (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat)))))%R /\
          (Rabs (x1 + y1) <=
           3%nat * radix * / pPred (vNum b) * Rabs (hd b output))%R))) ->
 Fbounded b x ->
 Fbounded b y ->
 (hdexp b L <= Fexp y)%Z ->
 (Fexp y <= Fexp x)%Z ->
 (FtoR radix x = 0%R :>R \/ FtoR radix y = 0%R :>R) \/
 (Rabs (FtoR radix y) <=
  pPred (vNum b) *
  (FtoR radix (Float 1%nat (Fexp x)) - FtoR radix (Float 1%nat (hdexp b L))))%R ->
 output = nil \/
 L = nil \/
 (Rabs (hd b L) <=
  3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat))))%R /\
 (Float (pPred (vNum b)) (Fexp (hd b L)) <=
  3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat))))%R /\
 (Rabs (sum L) <=
  length L *
  (3%nat * (radix * (Rabs (hd b output) * / (pPred (vNum b) - 1%nat)))))%R ->
 exists x' : float,
   (exists y' : float,
      (exists output' : list float,
         Fbounded b x' /\
         Fbounded b y' /\
         (x + (y + (sum output + sum L)))%R = (x' + (y' + sum output'))%R :>R /\
         length output' <= length L + length output /\
         (Fexp y' <= Fexp x')%Z /\
         IsRleEpsExp output' /\
         (output = nil \/
          (exists L1 : list float,
             (exists x1 : float,
                (exists y1 : float,
                   IsExp b (L1 ++ L) /\
                   endof input (L1 ++ L) /\
                   (x1 + (y1 + sum L1))%R = (x + y)%R :>R /\
                   (Rabs (sum L1) <=
                    length L1 *
                    (3%nat *
                     (radix *
                      (Rabs (hd b output) * / (pPred (vNum b) - 1%nat)))))%R /\
                   (Rabs (x1 + y1) <=
                    3%nat * radix * / pPred (vNum b) * Rabs (hd b output))%R)))) /\
         endof output' output /\
         (output' = nil \/
          (Rabs (x' + y') <=
           (pPred (vNum b) - 1%nat) * / pPred (vNum b) *
           ((6%nat * length input + 6%nat) *
            / (pPred (vNum b) - 1%nat - 6%nat * length input)) *
           Rabs (hd b output'))%R))).

Theorem FexpAdd_aux2 :
 forall L : list float,
 L = input ->
 IsExp b input ->
 (6%nat * length input * / (pPred (vNum b) - 1%nat) < 1)%R ->
 exists output : list float,
   (input <> nil -> length output <= S (length input)) /\
   sum input = sum output :>R /\ IsRleEpsExp output.

Theorem FexpAdd_main :
 IsExp b input ->
 (6%nat * length input * / (pPred (vNum b) - 1%nat) < 1)%R ->
 exists output : list float,
   (input <> nil -> length output <= S (length input)) /\
   sum input = sum output :>R /\ IsRleEpsExp output.
End FexpAdd.