aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Selection.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-10-02 16:17:51 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2016-10-02 16:17:51 +0200
commitf21a6b181dded86ef0e5c7ab94f74e5b960fd510 (patch)
tree01bb7b59e438c60d12d87d869b6c890095a977f4 /backend/Selection.v
parenta14b9578ee5297d954103e05d7b2d322816ddd8f (diff)
downloadcompcert-kvx-f21a6b181dded86ef0e5c7ab94f74e5b960fd510.tar.gz
compcert-kvx-f21a6b181dded86ef0e5c7ab94f74e5b960fd510.zip
Improve code generation for 64-bit signed integer division
Implement the 'shift right extended' trick, both in the generic implementation (backend/SplitLong) and in the IA32 port. Note that now SelectDiv depends on SelectLong, and that some work was moved from SelectLong to SelectDiv.
Diffstat (limited to 'backend/Selection.v')
-rw-r--r--backend/Selection.v6
1 files changed, 3 insertions, 3 deletions
diff --git a/backend/Selection.v b/backend/Selection.v
index 3aff446e..5cb5d119 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -27,7 +27,7 @@ Require Import Coqlib Maps.
Require Import AST Errors Integers Globalenvs Switch.
Require Cminor.
Require Import Op CminorSel.
-Require Import SelectOp SelectDiv SplitLong SelectLong.
+Require Import SelectOp SplitLong SelectLong SelectDiv.
Require Machregs.
Local Open Scope cminorsel_scope.
@@ -138,9 +138,9 @@ Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
| Cminor.Oaddl => addl arg1 arg2
| Cminor.Osubl => subl arg1 arg2
| Cminor.Omull => mull arg1 arg2
- | Cminor.Odivl => divl arg1 arg2
+ | Cminor.Odivl => divls arg1 arg2
| Cminor.Odivlu => divlu arg1 arg2
- | Cminor.Omodl => modl arg1 arg2
+ | Cminor.Omodl => modls arg1 arg2
| Cminor.Omodlu => modlu arg1 arg2
| Cminor.Oandl => andl arg1 arg2
| Cminor.Oorl => orl arg1 arg2