aboutsummaryrefslogtreecommitdiffstats
path: root/ia32/Asm.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-10-04 15:52:16 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2016-10-04 15:52:16 +0200
commitd2af79a77ed2936ff0ed90cadf8e48637d774d4c (patch)
tree09300943e12a98ae80598c79d455b31725271ead /ia32/Asm.v
parenta44893028eb1dd434c68001234ad56d030205a8e (diff)
downloadcompcert-d2af79a77ed2936ff0ed90cadf8e48637d774d4c.tar.gz
compcert-d2af79a77ed2936ff0ed90cadf8e48637d774d4c.zip
Turn 64-bit integer division and modulus by constants into multiply-high
This trick was already implemented for 32-bit integer division and modulus. Here we extend it to the 64-bit case. For 32-bit target processors, the runtime library must implement 64-bit multiply-high (signed and unsigned). Tentative implementations are provided for IA32 and PowerPC, but need testing.
Diffstat (limited to 'ia32/Asm.v')
-rw-r--r--ia32/Asm.v10
1 files changed, 8 insertions, 2 deletions
diff --git a/ia32/Asm.v b/ia32/Asm.v
index 01ecb15a..9d4036ff 100644
--- a/ia32/Asm.v
+++ b/ia32/Asm.v
@@ -164,9 +164,9 @@ Inductive instruction: Type :=
| Pimull_ri (rd: ireg) (n: int)
| Pimulq_ri (rd: ireg) (n: int64)
| Pimull_r (r1: ireg)
-(* | Pimulq_r (r1: ireg) *)
+ | Pimulq_r (r1: ireg)
| Pmull_r (r1: ireg)
-(* | Pmulq_r (r1: ireg) *)
+ | Pmulq_r (r1: ireg)
| Pcltd
| Pcqto
| Pdivl (r1: ireg)
@@ -718,9 +718,15 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pimull_r r1 =>
Next (nextinstr_nf (rs#RAX <- (Val.mul rs#RAX rs#r1)
#RDX <- (Val.mulhs rs#RAX rs#r1))) m
+ | Pimulq_r r1 =>
+ Next (nextinstr_nf (rs#RAX <- (Val.mull rs#RAX rs#r1)
+ #RDX <- (Val.mullhs rs#RAX rs#r1))) m
| Pmull_r r1 =>
Next (nextinstr_nf (rs#RAX <- (Val.mul rs#RAX rs#r1)
#RDX <- (Val.mulhu rs#RAX rs#r1))) m
+ | Pmulq_r r1 =>
+ Next (nextinstr_nf (rs#RAX <- (Val.mull rs#RAX rs#r1)
+ #RDX <- (Val.mullhu rs#RAX rs#r1))) m
| Pcltd =>
Next (nextinstr_nf (rs#RDX <- (Val.shr rs#RAX (Vint (Int.repr 31))))) m
| Pcqto =>