aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/SelectOp.vp
diff options
context:
space:
mode:
Diffstat (limited to 'powerpc/SelectOp.vp')
-rw-r--r--powerpc/SelectOp.vp27
1 files changed, 21 insertions, 6 deletions
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index 478ce251..50b1bdd6 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -38,12 +38,9 @@
Require Import Coqlib.
Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import CminorSel.
-Require Import OpHelpers.
+Require Import AST Integers Floats Builtins.
+Require Import Op OpHelpers CminorSel.
+Require Archi.
Local Open Scope cminorsel_scope.
@@ -517,6 +514,19 @@ Definition singleofintu (e: expr) :=
Definition singleoffloat (e: expr) := Eop Osingleoffloat (e ::: Enil).
Definition floatofsingle (e: expr) := Eop Ofloatofsingle (e ::: Enil).
+(** ** Selection *)
+
+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
+ then Some (Eop (Osel cond ty) (e1 ::: e2 ::: args))
+ else None.
+
(** ** Recognition of addressing modes for load and store operations *)
Definition can_use_Aindexed2 (chunk: memory_chunk): bool :=
@@ -560,3 +570,8 @@ Definition divf_base (e1: expr) (e2: expr) :=
Definition divfs_base (e1: expr) (e2: expr) :=
Eop Odivfs (e1 ::: e2 ::: Enil).
+
+(** Platform-specific known builtins *)
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ None.