summaryrefslogtreecommitdiff
path: root/powerpc/CombineOp.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
commit132e36fa0be63eb5672eda9168403d3fb74af2fa (patch)
tree33955e0ccb4210271c82326b941523e6e4b2c289 /powerpc/CombineOp.v
parent9ea00d39bb32c1f188f1af2745c3368da6a349c1 (diff)
CSE: add recognition of some combined operators, conditions, and addressing modes (cf. CombineOp.v)
Memory model: cleaning up Memdata Inlining and new Constprop: updated for ARM. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1902 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc/CombineOp.v')
-rw-r--r--powerpc/CombineOp.v119
1 files changed, 119 insertions, 0 deletions
diff --git a/powerpc/CombineOp.v b/powerpc/CombineOp.v
new file mode 100644
index 0000000..243da4e
--- /dev/null
+++ b/powerpc/CombineOp.v
@@ -0,0 +1,119 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Recognition of combined operations, addressing modes and conditions
+ during the [CSE] phase. *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Op.
+Require SelectOp.
+
+Definition valnum := positive.
+
+Inductive rhs : Type :=
+ | Op: operation -> list valnum -> rhs
+ | Load: memory_chunk -> addressing -> list valnum -> rhs.
+
+Section COMBINE.
+
+Variable get: valnum -> option rhs.
+
+Function combine_compimm_ne_0 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (c, ys)
+ | Some(Op (Oandimm n) ys) => Some (Cmasknotzero n, ys)
+ | _ => None
+ end.
+
+Function combine_compimm_eq_0 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (negate_condition c, ys)
+ | Some(Op (Oandimm n) ys) => Some (Cmaskzero n, ys)
+ | _ => None
+ end.
+
+Function combine_cond (cond: condition) (args: list valnum) : option(condition * list valnum) :=
+ match cond, args with
+ | Ccompimm Cne n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_ne_0 x else None
+ | Ccompimm Ceq n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_eq_0 x else None
+ | Ccompuimm Cne n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_ne_0 x else None
+ | Ccompuimm Ceq n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_eq_0 x else None
+ | _, _ => None
+ end.
+
+Function combine_addr (addr: addressing) (args: list valnum) : option(addressing * list valnum) :=
+ match addr, args with
+ | Aindexed n, x::nil =>
+ match get x with
+ | Some(Op (Oaddimm m) ys) => Some(Aindexed (Int.add m n), ys)
+ | Some(Op Oadd ys) => if Int.eq_dec n Int.zero then Some(Aindexed2, ys) else None
+ | _ => None
+ end
+ | _, _ => None
+ end.
+
+Function combine_op (op: operation) (args: list valnum) : option(operation * list valnum) :=
+ match op, args with
+ | Oaddimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oaddimm m) ys) => Some(Oaddimm (Int.add m n), ys)
+ | Some(Op (Osubimm m) ys) => Some(Osubimm (Int.add m n), ys)
+ | _ => None
+ end
+ | Osubimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oaddimm m) ys) => Some(Osubimm (Int.sub n m), ys)
+ | _ => None
+ end
+ | Oandimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oandimm m) ys) => Some(Oandimm (Int.and m n), ys)
+ | Some(Op (Orolm amount m) ys) => Some(Orolm amount (Int.and m n), ys)
+ | _ => None
+ end
+ | Oorimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oorimm m) ys) => Some(Oorimm (Int.or m n), ys)
+ | _ => None
+ end
+ | Oxorimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oxorimm m) ys) => Some(Oxorimm (Int.xor m n), ys)
+ | _ => None
+ end
+ | Orolm amount2 mask2, x :: nil =>
+ match get x with
+ | Some(Op (Oandimm mask1) ys) =>
+ Some(Orolm (Int.modu amount2 Int.iwordsize)
+ (Int.and (Int.rol mask1 amount2) mask2), ys)
+ | Some(Op (Orolm amount1 mask1) ys) =>
+ Some(Orolm (Int.modu (Int.add amount1 amount2) Int.iwordsize)
+ (Int.and (Int.rol mask1 amount2) mask2), ys)
+ | _ => None
+ end
+ | Ocmp cond, _ =>
+ match combine_cond cond args with
+ | Some(cond', args') => Some(Ocmp cond', args')
+ | None => None
+ end
+ | _, _ => None
+ end.
+
+End COMBINE.
+
+