aboutsummaryrefslogtreecommitdiffstats
path: root/backend/ValueDomain.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 /backend/ValueDomain.v
parenta44893028eb1dd434c68001234ad56d030205a8e (diff)
downloadcompcert-kvx-d2af79a77ed2936ff0ed90cadf8e48637d774d4c.tar.gz
compcert-kvx-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 'backend/ValueDomain.v')
-rw-r--r--backend/ValueDomain.v15
1 files changed, 14 insertions, 1 deletions
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index bf88a450..be8bcccc 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -1866,6 +1866,18 @@ Lemma mull_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mull v w) (mull x y).
Proof (binop_long_sound Int64.mul).
+Definition mullhs := binop_long Int64.mulhs.
+
+Lemma mullhs_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mullhs v w) (mullhs x y).
+Proof (binop_long_sound Int64.mulhs).
+
+Definition mullhu := binop_long Int64.mulhu.
+
+Lemma mullhu_sound:
+ forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mullhu v w) (mullhu x y).
+Proof (binop_long_sound Int64.mulhu).
+
Definition divls (v w: aval) :=
match w, v with
| L i2, L i1 =>
@@ -4559,7 +4571,8 @@ Hint Resolve cnot_sound symbol_address_sound
divs_sound divu_sound mods_sound modu_sound shrx_sound
shll_sound shrl_sound shrlu_sound
andl_sound orl_sound xorl_sound notl_sound roll_sound rorl_sound
- negl_sound addl_sound subl_sound mull_sound
+ negl_sound addl_sound subl_sound
+ mull_sound mullhs_sound mullhu_sound
divls_sound divlu_sound modls_sound modlu_sound shrxl_sound
negf_sound absf_sound
addf_sound subf_sound mulf_sound divf_sound