aboutsummaryrefslogtreecommitdiffstats
path: root/x86/SelectOp.vp
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2019-07-10 15:43:03 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-07-17 09:26:44 +0200
commit822a6a12316aa043eea7f6aed4d730bc10a73d7b (patch)
treebfeee08920725b1e871cc5902c4d71e862318292 /x86/SelectOp.vp
parent7cd0af8ba5d737fa9c45bb9fa454b38e9704097a (diff)
downloadcompcert-822a6a12316aa043eea7f6aed4d730bc10a73d7b.tar.gz
compcert-822a6a12316aa043eea7f6aed4d730bc10a73d7b.zip
x86_64: branchless implementation of floatofintu and intuoffloat
The implementation uses float <-> signed 64-bit integer conversion instructions, and is both efficient and branchless. Based on a suggestion by RĂ©mi Hutin.
Diffstat (limited to 'x86/SelectOp.vp')
-rw-r--r--x86/SelectOp.vp27
1 files changed, 17 insertions, 10 deletions
diff --git a/x86/SelectOp.vp b/x86/SelectOp.vp
index 378590ce..31be8c32 100644
--- a/x86/SelectOp.vp
+++ b/x86/SelectOp.vp
@@ -40,6 +40,7 @@ Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats Builtins.
Require Import Op CminorSel.
+Require Archi.
Local Open Scope cminorsel_scope.
@@ -498,21 +499,27 @@ Nondetfunction floatofint (e: expr) :=
end.
Definition intuoffloat (e: expr) :=
- Elet e
- (Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
- (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
- (intoffloat (Eletvar 1))
- (addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat.
+ if Archi.splitlong then
+ Elet e
+ (Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
+ (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
+ (intoffloat (Eletvar 1))
+ (addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat
+ else
+ Eop Olowlong (Eop Olongoffloat (e ::: Enil) ::: Enil).
Nondetfunction floatofintu (e: expr) :=
match e with
| Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_intu n)) Enil
| _ =>
- let f := Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil in
- Elet e
- (Econdition (CEcond (Ccompuimm Clt Float.ox8000_0000) (Eletvar O ::: Enil))
- (floatofint (Eletvar O))
- (addf (floatofint (addimm (Int.neg Float.ox8000_0000) (Eletvar O))) f))
+ if Archi.splitlong then
+ let f := Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil in
+ Elet e
+ (Econdition (CEcond (Ccompuimm Clt Float.ox8000_0000) (Eletvar O ::: Enil))
+ (floatofint (Eletvar O))
+ (addf (floatofint (addimm (Int.neg Float.ox8000_0000) (Eletvar O))) f))
+ else
+ Eop Ofloatoflong (Eop Ocast32unsigned (e ::: Enil) ::: Enil)
end.
Definition intofsingle (e: expr) := Eop Ointofsingle (e ::: Enil).