aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/SelectOp.vp
diff options
context:
space:
mode:
authorBernhard Schommer <bschommer@users.noreply.github.com>2019-05-28 10:33:34 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-05-28 10:33:34 +0200
commite10555313645cf3c35f244f42afa5a03fba2bac1 (patch)
tree7b1e42b23ac78f5866681f12402d7c7aeceaecd3 /powerpc/SelectOp.vp
parentc36514ac4b05f78dd2e02fab3f8886cab8234925 (diff)
downloadcompcert-kvx-e10555313645cf3c35f244f42afa5a03fba2bac1.tar.gz
compcert-kvx-e10555313645cf3c35f244f42afa5a03fba2bac1.zip
Provide a float select operation for PowerPC. (#173)
The FP select for PowerPC stores both addresses in two subsequent stack slots and loads them using an offset created from the result of the comparison.
Diffstat (limited to 'powerpc/SelectOp.vp')
-rw-r--r--powerpc/SelectOp.vp3
1 files changed, 3 insertions, 0 deletions
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index 41bc0efc..24aeb531 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -43,6 +43,7 @@ Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import CminorSel.
+Require Archi.
Local Open Scope cminorsel_scope.
@@ -521,6 +522,8 @@ Definition floatofsingle (e: expr) := Eop Ofloatofsingle (e ::: Enil).
Definition select (ty: typ) (cond: condition) (args: exprlist) (e1 e2: expr) :=
if match ty with
| Tint => true
+ | Tfloat => true
+ | Tsingle => true
| Tlong => Archi.ppc64
| _ => false
end