Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
7
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in
Toggle navigation
Open sidebar
Flocq
flocq
Commits
5b5c74d4
Commit
5b5c74d4
authored
Dec 07, 2010
by
Guillaume Melquiond
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Added F2R_eq_compat.
parent
d50deb55
Changes
2
Hide whitespace changes
Inline
Side-by-side
Showing
2 changed files
with
13 additions
and
4 deletions
+13
-4
src/Core/Fcore_float_prop.v
src/Core/Fcore_float_prop.v
+9
-0
src/Core/Fcore_generic_fmt.v
src/Core/Fcore_generic_fmt.v
+4
-4
No files found.
src/Core/Fcore_float_prop.v
View file @
5b5c74d4
...
...
@@ -87,6 +87,15 @@ apply bpow_gt_0.
now
apply
Z2R_lt
.
Qed
.
Theorem
F2R_eq_compat
:
forall
e
m1
m2
:
Z
,
(
m1
=
m2
)
%
Z
->
(
F2R
(
Float
beta
m1
e
)
=
F2R
(
Float
beta
m2
e
))
%
R
.
Proof
.
intros
e
m1
m2
H
.
now
apply
(
f_equal
(
fun
m
=>
F2R
(
Float
beta
m
e
))).
Qed
.
Theorem
F2R_eq_reg
:
forall
e
m1
m2
:
Z
,
F2R
(
Float
beta
m1
e
)
=
F2R
(
Float
beta
m2
e
)
->
...
...
src/Core/Fcore_generic_fmt.v
View file @
5b5c74d4
...
...
@@ -341,7 +341,7 @@ intros (m, e) Hf.
unfold
canonic
in
Hf
.
simpl
in
Hf
.
unfold
generic_format
,
scaled_mantissa
.
rewrite
<-
Hf
.
apply
(
f_equal
(
fun
m
=>
F2R
(
Float
beta
m
e
)))
.
apply
F2R_eq_compat
.
unfold
F2R
.
simpl
.
rewrite
Rmult_assoc
,
<-
bpow_plus
,
Zplus_opp_r
,
Rmult_1_r
.
now
rewrite
Ztrunc_Z2R
.
...
...
@@ -710,7 +710,7 @@ Proof.
intros
x
.
unfold
round
.
rewrite
opp_F2R
,
canonic_exponent_opp
,
scaled_mantissa_opp
.
apply
(
f_equal
(
fun
m
=>
F2R
(
Float
beta
m
_
)))
.
apply
F2R_eq_compat
.
apply
sym_eq
.
exact
(
Zopp_involutive
_
).
Qed
.
...
...
@@ -942,7 +942,7 @@ split ; intros Hx.
(
*
*
)
replace
(
round
rndZR
x
)
with
(
round
rndDN
x
).
apply
round_DN_pt
.
apply
(
f_equal
(
fun
v
=>
F2R
(
Float
beta
v
_
)))
.
apply
F2R_eq_compat
.
apply
sym_eq
.
apply
Ztrunc_floor
.
rewrite
<-
(
Rmult_0_l
(
bpow
(
-
canonic_exponent
x
))).
...
...
@@ -951,7 +951,7 @@ apply bpow_ge_0.
(
*
*
)
replace
(
round
rndZR
x
)
with
(
round
rndUP
x
).
apply
round_UP_pt
.
apply
(
f_equal
(
fun
v
=>
F2R
(
Float
beta
v
_
)))
.
apply
F2R_eq_compat
.
apply
sym_eq
.
apply
Ztrunc_ceil
.
rewrite
<-
(
Rmult_0_l
(
bpow
(
-
canonic_exponent
x
))).
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment