Library FmaErr
Require Export AllFloat.
Require Export Veltkamp.
Section GenericA.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Variable a b e x y:float.
Hypothesis eLea: (Rabs e <= /2*Fulp bo radix p a)%R.
Hypothesis eLeb: (Rabs e <= /2*Fulp bo radix p b)%R.
Hypothesis xDef: Closest bo radix (a+b)%R x.
Hypothesis yDef: Closest bo radix (a+b+e)%R y.
Hypothesis Nx: Fnormal radix bo x.
Hypothesis Ny: Fnormal radix bo y.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Ca: Fcanonic radix bo a.
Hypothesis Fexpb: (- dExp bo < Fexp b)%Z.
Let Unmoins := (1 - (powerRZ radix (Zsucc (-p)))/2)%R.
Let Unplus := (1 + (powerRZ radix (Zsucc (-p)))/2)%R.
Lemma UnMoinsPos: (0 < Unmoins)%R.
Lemma ClosestRoundeLeNormal: forall (z : R) (f : float),
Closest bo radix z f ->
Fnormal radix bo f ->
(Rabs f <= (Rabs z) / Unmoins)%R.
Lemma ClosestRoundeGeNormal: forall (z : R) (f : float),
Closest bo radix z f ->
Fnormal radix bo f ->
(Rabs z <= (Rabs f) * Unplus)%R.
Lemma abeLeab: (Rabs b <= Rabs a)%R -> (2*powerRZ radix (Fexp b) <= Rabs (a+b))%R
-> (Rabs (a+b) <= Rabs (a+b+e) *4/3)%R.
Lemma xLe2y_aux1: (Rabs b <= Rabs a)%R -> (powerRZ radix (Fexp b) = Rabs (a+b))%R
-> (Rabs x <= 2*Rabs y)%R.
Lemma xLe2y_aux2 : (Rabs b <= Rabs a)%R -> (Rabs x <= 2*Rabs y)%R.
Lemma yLe2x_aux : (Rabs b <= Rabs a)%R -> ~(FtoRradix x=0)%R
-> (Rabs y <= 2*Rabs x)%R.
End GenericA.
Section GenericB.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Variable a b e x y:float.
Hypothesis eLea: (Rabs e <= /2*Fulp bo radix p a)%R.
Hypothesis eLeb: (Rabs e <= /2*Fulp bo radix p b)%R.
Hypothesis xDef: Closest bo radix (a+b)%R x.
Hypothesis yDef: Closest bo radix (a+b+e)%R y.
Hypothesis Nx: Fnormal radix bo x.
Hypothesis Ny: Fnormal radix bo y.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Ca: Fcanonic radix bo a.
Hypothesis Fexpb: (- dExp bo < Fexp b)%Z.
Hypothesis Fexpa: (- dExp bo < Fexp a)%Z.
Hypothesis dsd: ((0<= y)%R -> (0<= x)%R) /\ ((y <= 0)%R -> (x <= 0)%R).
Lemma xLe2y : (Rabs x <= 2*Rabs y)%R.
Lemma yLe2x: ~(FtoRradix x=0)%R -> (Rabs y <= 2*Rabs x)%R.
Lemma Subexact: exists v:float, (FtoRradix v=x-y)%R /\ (Fbounded bo v) /\
(Fexp v=Zmin (Fexp x) (Fexp y))%Z.
End GenericB.
Section GenericC.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 1 < p.
Hypothesis Evenradix: (Even radix).
Lemma powerRZSumRle:forall (e1 e2:Z),
(e2<= e1)%Z ->
(powerRZ radix e1 + powerRZ radix e2 <= powerRZ radix (e1+1))%R.
Lemma LSB_Pred: forall x y:float,
(Rabs x < Rabs y)%R -> (LSB radix x <= LSB radix y)%Z
-> (Rabs x <= Rabs y - powerRZ radix (LSB radix x))%R.
Variables x1 x2 y f:float.
Hypothesis x1Def: Closest bo radix (x1+x2)%R x1.
Hypothesis fDef : Closest bo radix (x1+x2+y)%R f.
Hypothesis yLe: (MSB radix y < LSB radix x2)%Z.
Hypothesis Nx1: Fnormal radix bo x1.
Hypothesis x1Pos: (0 < x1)%R.
Hypothesis x2NonZero: ~(FtoRradix x2 =0)%R.
Hypothesis x1Exp: (-dExp bo < Fexp x1)%Z.
Lemma Midpoint_aux_aux:
(FtoRradix x1= f) \/ (exists v:float, (FtoRradix v=x2)%R /\ (Fexp x1 -2 <= Fexp v)%Z).
End GenericC.
Section GenericD.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 1 < p.
Hypothesis Evenradix: (Even radix).
Variables x1 x2 y f:float.
Hypothesis x1Def: Closest bo radix (x1+x2)%R x1.
Hypothesis fDef : Closest bo radix (x1+x2+y)%R f.
Hypothesis yLe: (MSB radix y < LSB radix x2)%Z.
Hypothesis Nx1: Fnormal radix bo x1.
Hypothesis x2NonZero: ~(FtoRradix x2 =0)%R.
Hypothesis x1Exp: (-dExp bo < Fexp x1)%Z.
Lemma Midpoint_aux:
(FtoRradix x1= f) \/ (exists v:float, (FtoRradix v=x2)%R /\ (Fexp x1 -2 <= Fexp v)%Z).
End GenericD.
Section Be2Zero.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Theorem TwoSumProp: forall (a b x y:float),
(Fbounded bo a) ->
(Closest bo radix (a+b)%R x) -> (FtoRradix y=a+b-x)%R
-> (Rabs y <= Rabs b)%R.
Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.
Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.
Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.
Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Lemma gatCorrect: exists v:float, (FtoRradix v=be1-r1)%R /\ (Fbounded bo v)
/\ (Fexp v=Zmin (Fexp be1) (Fexp r1))%Z.
Hypothesis Be2Zero: (FtoRradix be2=0)%R.
Theorem FmaErr_aux1: (a*x+y=r1+ga+al2)%R.
End Be2Zero.
Section Be2NonZero.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
(P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.
Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.
Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.
Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z.
Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis be2Bounded: Fbounded bo be2.
Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).
Lemma Expr1 : (Fexp r1 <= Fexp be1+1)%Z.
Lemma Expbe1: (Fexp be1 <= Fexp r1+1)%Z.
Theorem BoundedL: forall (r:R) (x0:float) (e:Z),
(e <=Fexp x0)%Z -> (-dExp bo <= e)%Z -> (FtoRradix x0=r)%R ->
(Rabs r < powerRZ radix (e+p))%R ->
(exists x':float, (FtoRradix x'=r) /\ (Fbounded bo x') /\ Fexp x'=e).
Lemma Zmin_Zlt : forall z1 z2 z3 : Z,
(z1 < z2)%Z -> (z1 < z3)%Z -> (z1 < Zmin z2 z3)%Z.
Hypothesis Be2NonZero: ~(FtoRradix be2=0)%R.
Lemma be2MuchSmaller:
~(FtoRradix al2=0)%R -> ~(FtoRradix u2=0)%R ->
(MSB radix al2 < LSB radix be2)%Z.
Lemma gaCorrect: exists v:float, (FtoRradix v=be1-r1+be2)%R /\ (Fbounded bo v).
Theorem FmaErr_aux2: (a*x+y=r1+ga+al2)%R.
End Be2NonZero.
Section Final.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
(P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.
Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.
Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.
Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z.
Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis be2Bounded: Fbounded bo be2.
Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).
Theorem FmaErr_aux: (a*x+y=r1+ga+al2)%R.
End Final.
Section Final2.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Lemma ClosestZero1: forall (r:R) (f g:float), (Closest bo radix r f) -> (FtoRradix f=0)%R ->
(r=g)%R -> (-dExp bo <= Fexp g)%Z -> (r=0)%R.
Lemma ClosestZero2: forall (r:R) (x:float),
(Closest bo radix r x) -> (r=0)%R -> (FtoRradix x=0)%R.
Lemma LeExpRound:forall f g:float, Closest bo radix f g
-> exists g':float, Fbounded bo g' /\ FtoRradix g'=g /\ (Fexp f <= Fexp g')%Z.
Lemma LeExpRound2:forall (n:Z) (f g:float), Closest bo radix f g
-> (n <= Fexp f)%Z
-> exists g':float, Fbounded bo g' /\ FtoRradix g'=g /\ (n <= Fexp g')%Z.
Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
(P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.
Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.
Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.
Hypothesis Nbe1: Fcanonic radix bo be1.
Hypothesis Nr1 : Fcanonic radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z \/ (FtoRradix al1=0)%R.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z \/ (FtoRradix u1=0)%R.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z\/ (FtoRradix be1=0)%R.
Hypothesis Exp4: (Fnormal radix bo r1) \/ (FtoRradix r1=0)%R.
Hypothesis Exp5: (-dExp bo <= Fexp a+Fexp x)%Z.
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).
Theorem FmaErr: (a*x+y=r1+ga+al2)%R.
Theorem Fma_FTS: (exists ga_e:float, exists al2_e:float,
(FtoRradix ga_e=ga)%R /\ (FtoRradix al2_e=al2)%R
/\ (Fbounded bo ga_e) /\ (Fbounded bo al2_e) /\
(Fexp al2_e <= Fexp ga_e)%Z).
End Final2.
Section Final_Even.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.
Let FtoRradix := FtoR radix.
Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hint Resolve radixMoreThanZERO: zarith.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).
Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.
Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.
Hypothesis Nbe1: Fcanonic radix bo be1.
Hypothesis Nr1 : Fcanonic radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z \/ (FtoRradix al1=0)%R.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z \/ (FtoRradix u1=0)%R.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z\/ (FtoRradix be1=0)%R.
Hypothesis Exp4: (Fnormal radix bo r1) \/ (FtoRradix r1=0)%R.
Hypothesis Exp5: (-dExp bo <= Fexp a+Fexp x)%Z.
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis r1DefE: (EvenClosest bo radix p (a*x+y)%R r1).
Hypothesis be1DefE:(EvenClosest bo radix p (u1+al1)%R be1).
Theorem FmaErr_Even: (a*x+y=r1+ga+al2)%R.
Theorem Fma_FTS_Even: (exists ga_e:float, exists al2_e:float,
(FtoRradix ga_e=ga)%R /\ (FtoRradix al2_e=al2)%R
/\ (Fbounded bo ga_e) /\ (Fbounded bo al2_e) /\
(Fexp al2_e <= Fexp ga_e)%Z).
End Final_Even.