aboutsummaryrefslogtreecommitdiffstats
path: root/ia32/ConstpropOp.vp
diff options
context:
space:
mode:
Diffstat (limited to 'ia32/ConstpropOp.vp')
-rw-r--r--ia32/ConstpropOp.vp226
1 files changed, 197 insertions, 29 deletions
diff --git a/ia32/ConstpropOp.vp b/ia32/ConstpropOp.vp
index a3de748c..c35d3def 100644
--- a/ia32/ConstpropOp.vp
+++ b/ia32/ConstpropOp.vp
@@ -2,7 +2,7 @@
(* *)
(* The Compcert verified compiler *)
(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* Xavier Leroy, INRIA Paris *)
(* *)
(* Copyright Institut National de Recherche en Informatique et en *)
(* Automatique. All rights reserved. This file is distributed *)
@@ -13,15 +13,32 @@
(** Strength reduction for operators and conditions.
This is the machine-dependent part of [Constprop]. *)
-Require Import Coqlib.
-Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import Registers.
+Require Import Coqlib Compopts.
+Require Import AST Integers Floats.
+Require Import Op Registers.
Require Import ValueDomain.
+(** * Converting known values to constants *)
+
+Parameter symbol_is_external: ident -> bool. (**r See [SelectOp] *)
+
+Definition Olea_ptr (a: addressing) := if Archi.ptr64 then Oleal a else Olea a.
+
+Definition const_for_result (a: aval) : option operation :=
+ match a with
+ | I n => Some(Ointconst n)
+ | L n => if Archi.ptr64 then Some(Olongconst n) else None
+ | F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None
+ | FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None
+ | Ptr(Gl id ofs) =>
+ if symbol_is_external id then
+ if Ptrofs.eq ofs Ptrofs.zero then Some (Oindirectsymbol id) else None
+ else
+ Some (Olea_ptr (Aglobal id ofs))
+ | Ptr(Stk ofs) => Some(Olea_ptr (Ainstack ofs))
+ | _ => None
+ end.
+
(** * Operator strength reduction *)
(** We now define auxiliary functions for strength reduction of
@@ -40,6 +57,14 @@ Nondetfunction cond_strength_reduction
(Ccompuimm (swap_comparison c) n1, r2 :: nil)
| Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
(Ccompuimm c n2, r1 :: nil)
+ | Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Ccomplimm (swap_comparison c) n1, r2 :: nil)
+ | Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Ccomplimm c n2, r1 :: nil)
+ | Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Ccompluimm (swap_comparison c) n1, r2 :: nil)
+ | Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Ccompluimm c n2, r1 :: nil)
| _, _, _ =>
(cond, args)
end.
@@ -61,53 +86,120 @@ Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
make_cmp_base c args vl
end.
-Nondetfunction addr_strength_reduction
+(** For addressing modes, we need to distinguish
+- reductions that produce pointers (i.e. that produce [Aglobal], [Ainstack], [Abased] and [Abasedscaled] addressing modes), which are valid only if the pointer size is right;
+- other reductions (producing [Aindexed] or [Aindexed2] modes), which are valid independently of the pointer size.
+*)
+
+Nondetfunction addr_strength_reduction_32_generic
+ (addr: addressing) (args: list reg) (vl: list aval) :=
+ match addr, args, vl with
+ | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
+ (Aindexed (Int.signed n1 + ofs), r2 :: nil)
+ | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
+ (Aindexed (Int.signed n2 + ofs), r1 :: nil)
+ | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
+ (Aindexed (Int.signed n2 * sc + ofs), r1 :: nil)
+ | Aindexed2scaled sc ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
+ (Ascaled sc (Int.signed n1 + ofs), r2 :: nil)
+ | _, _ =>
+ (addr, args)
+ end.
+
+Nondetfunction addr_strength_reduction_32
(addr: addressing) (args: list reg) (vl: list aval) :=
+
+ if Archi.ptr64 then addr_strength_reduction_32_generic addr args vl else
+
match addr, args, vl with
| Aindexed ofs, r1 :: nil, Ptr(Gl symb n) :: nil =>
- (Aglobal symb (Int.add n ofs), nil)
+ (Aglobal symb (Ptrofs.add n (Ptrofs.repr ofs)), nil)
| Aindexed ofs, r1 :: nil, Ptr(Stk n) :: nil =>
- (Ainstack (Int.add n ofs), nil)
+ (Ainstack (Ptrofs.add n (Ptrofs.repr ofs)), nil)
| Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Int.add (Int.add n1 n2) ofs), nil)
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int n2)) (Ptrofs.repr ofs)), nil)
| Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Gl symb n2) :: nil =>
- (Aglobal symb (Int.add (Int.add n1 n2) ofs), nil)
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n2 (Ptrofs.of_int n1)) (Ptrofs.repr ofs)), nil)
| Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Stk n1) :: I n2 :: nil =>
- (Ainstack (Int.add (Int.add n1 n2) ofs), nil)
+ (Ainstack (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int n2)) (Ptrofs.repr ofs)), nil)
| Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: Ptr(Stk n2) :: nil =>
- (Ainstack (Int.add (Int.add n1 n2) ofs), nil)
+ (Ainstack (Ptrofs.add (Ptrofs.add (Ptrofs.of_int n1) n2) (Ptrofs.repr ofs)), nil)
| Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abased symb (Int.add n1 ofs), r2 :: nil)
+ (Abased symb (Ptrofs.add n1 (Ptrofs.repr ofs)), r2 :: nil)
| Aindexed2 ofs, r1 :: r2 :: nil, v1 :: Ptr(Gl symb n2) :: nil =>
- (Abased symb (Int.add n2 ofs), r1 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
- (Aindexed (Int.add n1 ofs), r2 :: nil)
- | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.add n2 ofs), r1 :: nil)
+ (Abased symb (Ptrofs.add n2 (Ptrofs.repr ofs)), r1 :: nil)
| Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: I n2 :: nil =>
- (Aglobal symb (Int.add (Int.add n1 (Int.mul n2 sc)) ofs), nil)
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int (Int.mul n2 (Int.repr sc)))) (Ptrofs.repr ofs)), nil)
+
| Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: v2 :: nil =>
- (Abasedscaled sc symb (Int.add n1 ofs), r2 :: nil)
- | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
- (Aindexed (Int.add (Int.mul n2 sc) ofs), r1 :: nil)
+ (Abasedscaled sc symb (Ptrofs.add n1 (Ptrofs.repr ofs)), r2 :: nil)
| Abased id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Int.add ofs n1), nil)
+ (Aglobal id (Ptrofs.add ofs (Ptrofs.of_int n1)), nil)
| Abasedscaled sc id ofs, r1 :: nil, I n1 :: nil =>
- (Aglobal id (Int.add ofs (Int.mul sc n1)), nil)
+ (Aglobal id (Ptrofs.add ofs (Ptrofs.of_int (Int.mul n1 (Int.repr sc)))), nil)
| _, _ =>
+ addr_strength_reduction_32_generic addr args vl
+ end.
+
+Nondetfunction addr_strength_reduction_64_generic
+ (addr: addressing) (args: list reg) (vl: list aval) :=
+ match addr, args, vl with
+ | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Aindexed (Int64.signed n1 + ofs), r2 :: nil)
+ | Aindexed2 ofs, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Aindexed (Int64.signed n2 + ofs), r1 :: nil)
+ | Aindexed2scaled sc ofs, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Aindexed (Int64.signed n2 * sc + ofs), r1 :: nil)
+ | Aindexed2scaled sc ofs, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Ascaled sc (Int64.signed n1 + ofs), r2 :: nil)
+ | _, _ =>
(addr, args)
end.
+Nondetfunction addr_strength_reduction_64
+ (addr: addressing) (args: list reg) (vl: list aval) :=
+
+ if negb Archi.ptr64 then addr_strength_reduction_64_generic addr args vl else
+
+ match addr, args, vl with
+
+ | Aindexed ofs, r1 :: nil, Ptr(Gl symb n) :: nil =>
+ (Aglobal symb (Ptrofs.add n (Ptrofs.repr ofs)), nil)
+ | Aindexed ofs, r1 :: nil, Ptr(Stk n) :: nil =>
+ (Ainstack (Ptrofs.add n (Ptrofs.repr ofs)), nil)
+
+ | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: L n2 :: nil =>
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 n2)) (Ptrofs.repr ofs)), nil)
+ | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: Ptr(Gl symb n2) :: nil =>
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n2 (Ptrofs.of_int64 n1)) (Ptrofs.repr ofs)), nil)
+ | Aindexed2 ofs, r1 :: r2 :: nil, Ptr(Stk n1) :: L n2 :: nil =>
+ (Ainstack (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 n2)) (Ptrofs.repr ofs)), nil)
+ | Aindexed2 ofs, r1 :: r2 :: nil, L n1 :: Ptr(Stk n2) :: nil =>
+ (Ainstack (Ptrofs.add (Ptrofs.add (Ptrofs.of_int64 n1) n2) (Ptrofs.repr ofs)), nil)
+
+ | Aindexed2scaled sc ofs, r1 :: r2 :: nil, Ptr(Gl symb n1) :: L n2 :: nil =>
+ (Aglobal symb (Ptrofs.add (Ptrofs.add n1 (Ptrofs.of_int64 (Int64.mul n2 (Int64.repr sc)))) (Ptrofs.repr ofs)), nil)
+
+ | _, _ =>
+ addr_strength_reduction_64_generic addr args vl
+ end.
+
+Definition addr_strength_reduction
+ (addr: addressing) (args: list reg) (vl: list aval) :=
+ if Archi.ptr64
+ then addr_strength_reduction_64 addr args vl
+ else addr_strength_reduction_32 addr args vl.
+
Definition make_addimm (n: int) (r: reg) :=
if Int.eq n Int.zero
then (Omove, r :: nil)
- else (Olea (Aindexed n), r :: nil).
+ else (Olea (Aindexed (Int.signed n)), r :: nil).
Definition make_shlimm (n: int) (r1 r2: reg) :=
if Int.eq n Int.zero then (Omove, r1 :: nil)
@@ -173,6 +265,64 @@ Definition make_moduimm n (r1 r2: reg) :=
| None => (Omodu, r1 :: r2 :: nil)
end.
+Definition make_addlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero
+ then (Omove, r :: nil)
+ else (Oleal (Aindexed (Int64.signed n)), r :: nil).
+
+Definition make_shllimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshllimm n, r1 :: nil)
+ else (Oshll, r1 :: r2 :: nil).
+
+Definition make_shrlimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshrlimm n, r1 :: nil)
+ else (Oshrl, r1 :: r2 :: nil).
+
+Definition make_shrluimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshrluimm n, r1 :: nil)
+ else (Oshrlu, r1 :: r2 :: nil).
+
+Definition make_mullimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero then
+ (Olongconst Int64.zero, nil)
+ else if Int64.eq n Int64.one then
+ (Omove, r :: nil)
+ else
+ match Int64.is_power2' n with
+ | Some l => (Oshllimm l, r :: nil)
+ | None => (Omullimm n, r :: nil)
+ end.
+
+Definition make_andlimm (n: int64) (r: reg) (a: aval) :=
+ if Int64.eq n Int64.zero then (Olongconst Int64.zero, nil)
+ else if Int64.eq n Int64.mone then (Omove, r :: nil)
+ else (Oandlimm n, r :: nil).
+
+Definition make_orlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero then (Omove, r :: nil)
+ else if Int64.eq n Int64.mone then (Olongconst Int64.mone, nil)
+ else (Oorlimm n, r :: nil).
+
+Definition make_xorlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero then (Omove, r :: nil)
+ else if Int64.eq n Int64.mone then (Onotl, r :: nil)
+ else (Oxorlimm n, r :: nil).
+
+Definition make_divluimm n (r1 r2: reg) :=
+ match Int64.is_power2' n with
+ | Some l => (Oshrluimm l, r1 :: nil)
+ | None => (Odivlu, r1 :: r2 :: nil)
+ end.
+
+Definition make_modluimm n (r1 r2: reg) :=
+ match Int64.is_power2 n with
+ | Some l => (Oandlimm (Int64.sub n Int64.one), r1 :: nil)
+ | None => (Omodlu, r1 :: r2 :: nil)
+ end.
+
Definition make_mulfimm (n: float) (r r1 r2: reg) :=
if Float.eq_dec n (Float.of_int (Int.repr 2))
then (Oaddf, r :: r :: nil)
@@ -216,8 +366,26 @@ Nondetfunction op_strength_reduction
| Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrimm n2 r1 r2
| Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shruimm n2 r1 r2
| Olea addr, args, vl =>
- let (addr', args') := addr_strength_reduction addr args vl in
+ let (addr', args') := addr_strength_reduction_32 addr args vl in
(Olea addr', args')
+ | Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm (Int64.neg n2) r1
+ | Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_mullimm n1 r2
+ | Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_mullimm n2 r1
+ | Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divluimm n2 r1 r2
+ | Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_modluimm n2 r1 r2
+ | Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_andlimm n1 r2 v2
+ | Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_andlimm n2 r1 v1
+ | Oandlimm n, r1 :: nil, v1 :: nil => make_andlimm n r1 v1
+ | Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_orlimm n1 r2
+ | Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_orlimm n2 r1
+ | Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_xorlimm n1 r2
+ | Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_xorlimm n2 r1
+ | Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shllimm n2 r1 r2
+ | Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrlimm n2 r1 r2
+ | Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrluimm n2 r1 r2
+ | Oleal addr, args, vl =>
+ let (addr', args') := addr_strength_reduction_64 addr args vl in
+ (Oleal addr', args')
| Ocmp c, args, vl => make_cmp c args vl
| Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
| Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2