aboutsummaryrefslogtreecommitdiffstats
path: root/backend/SplitLongproof.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/SplitLongproof.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/SplitLongproof.v')
-rw-r--r--backend/SplitLongproof.v22
1 files changed, 20 insertions, 2 deletions
diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v
index a10ee3f7..31f5db67 100644
--- a/backend/SplitLongproof.v
+++ b/backend/SplitLongproof.v
@@ -48,7 +48,9 @@ Axiom i64_helpers_correct :
/\ (forall x y z, Val.modlu x y = Some z -> external_implements "__i64_umod" sig_ll_l (x::y::nil) z)
/\ (forall x y, external_implements "__i64_shl" sig_li_l (x::y::nil) (Val.shll x y))
/\ (forall x y, external_implements "__i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y))
- /\ (forall x y, external_implements "__i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)).
+ /\ (forall x y, external_implements "__i64_sar" sig_li_l (x::y::nil) (Val.shrl x y))
+ /\ (forall x y, external_implements "__i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y))
+ /\ (forall x y, external_implements "__i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)).
Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop :=
(prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))).
@@ -66,7 +68,9 @@ Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F)
/\ helper_declared p i64_umod "__i64_umod" sig_ll_l
/\ helper_declared p i64_shl "__i64_shl" sig_li_l
/\ helper_declared p i64_shr "__i64_shr" sig_li_l
- /\ helper_declared p i64_sar "__i64_sar" sig_li_l.
+ /\ helper_declared p i64_sar "__i64_sar" sig_li_l
+ /\ helper_declared p i64_umulh "__i64_umulh" sig_ll_l
+ /\ helper_declared p i64_smulh "__i64_smulh" sig_ll_l.
(** * Correctness of the instruction selection functions for 64-bit operators *)
@@ -823,6 +827,20 @@ Proof.
- apply eval_mull_base; auto.
Qed.
+Theorem eval_mullhu:
+ forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)).
+Proof.
+ unfold mullhu; intros; red; intros. econstructor; split; eauto.
+ eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper; eauto. UseHelper.
+Qed.
+
+Theorem eval_mullhs:
+ forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)).
+Proof.
+ unfold mullhs; intros; red; intros. econstructor; split; eauto.
+ eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper; eauto. UseHelper.
+Qed.
+
Theorem eval_shrxlimm:
forall le a n x z,
Archi.ptr64 = false ->