summaryrefslogtreecommitdiff
path: root/arm/ConstpropOp.vp
diff options
context:
space:
mode:
Diffstat (limited to 'arm/ConstpropOp.vp')
-rw-r--r--arm/ConstpropOp.vp200
1 files changed, 35 insertions, 165 deletions
diff --git a/arm/ConstpropOp.vp b/arm/ConstpropOp.vp
index 47d2086..0540781 100644
--- a/arm/ConstpropOp.vp
+++ b/arm/ConstpropOp.vp
@@ -19,35 +19,16 @@ Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import Registers.
+Require Import ValueDomain.
-(** * Static analysis *)
-
-(** To each pseudo-register at each program point, the static analysis
- associates a compile-time approximation taken from the following set. *)
-
-Inductive approx : Type :=
- | Novalue: approx (** No value possible, code is unreachable. *)
- | Unknown: approx (** All values are possible,
- no compile-time information is available. *)
- | I: int -> approx (** A known integer value. *)
- | F: float -> approx (** A known floating-point value. *)
- | L: int64 -> approx (** A know 64-bit integer value. *)
- | G: ident -> int -> approx
- (** The value is the address of the given global
- symbol plus the given integer offset. *)
- | S: int -> approx. (** The value is the stack pointer plus the offset. *)
-
+(** * Operator strength reduction *)
-(** We now define the abstract interpretations of conditions and operators
- over this set of approximations. For instance, the abstract interpretation
- of the operator [Oaddf] applied to two expressions [a] and [b] is
- [F(Float.add f g)] if [a] and [b] have static approximations [F f]
- and [F g] respectively, and [Unknown] otherwise.
+(** We now define auxiliary functions for strength reduction of
+ operators and addressing modes: replacing an operator with a cheaper
+ one if some of its arguments are statically known. These are again
+ large pattern-matchings expressed in indirect style. *)
- The static approximations are defined by large pattern-matchings over
- the approximations of the results. We write these matchings in the
- indirect style described in file [SelectOp] to avoid excessive
- duplication of cases in proofs. *)
+Section STRENGTH_REDUCTION.
Definition eval_static_shift (s: shift) (n: int) : int :=
match s with
@@ -57,134 +38,8 @@ Definition eval_static_shift (s: shift) (n: int) : int :=
| Sror x => Int.ror n x
end.
-Nondetfunction eval_static_condition (cond: condition) (vl: list approx) :=
- match cond, vl with
- | Ccomp c, I n1 :: I n2 :: nil => Some(Int.cmp c n1 n2)
- | Ccompu c, I n1 :: I n2 :: nil => Some(Int.cmpu c n1 n2)
- | Ccompshift c s, I n1 :: I n2 :: nil => Some(Int.cmp c n1 (eval_static_shift s n2))
- | Ccompushift c s, I n1 :: I n2 :: nil => Some(Int.cmpu c n1 (eval_static_shift s n2))
- | Ccompimm c n, I n1 :: nil => Some(Int.cmp c n1 n)
- | Ccompuimm c n, I n1 :: nil => Some(Int.cmpu c n1 n)
- | Ccompf c, F n1 :: F n2 :: nil => Some(Float.cmp c n1 n2)
- | Cnotcompf c, F n1 :: F n2 :: nil => Some(negb(Float.cmp c n1 n2))
- | Ccompfzero c, F n1 :: nil => Some(Float.cmp c n1 Float.zero)
- | Cnotcompfzero c, F n1 :: nil => Some(negb(Float.cmp c n1 Float.zero))
- | _, _ => None
- end.
-
-Definition eval_static_condition_val (cond: condition) (vl: list approx) :=
- match eval_static_condition cond vl with
- | None => Unknown
- | Some b => I(if b then Int.one else Int.zero)
- end.
-
-Definition eval_static_intoffloat (f: float) :=
- match Float.intoffloat f with Some x => I x | None => Unknown end.
-
-Definition eval_static_intuoffloat (f: float) :=
- match Float.intuoffloat f with Some x => I x | None => Unknown end.
-
-Parameter propagate_float_constants: unit -> bool.
-
-Nondetfunction eval_static_operation (op: operation) (vl: list approx) :=
- match op, vl with
- | Omove, v1::nil => v1
- | Ointconst n, nil => I n
- | Ofloatconst n, nil => if propagate_float_constants tt then F n else Unknown
- | Oaddrsymbol s n, nil => G s n
- | Oaddrstack n, nil => S n
- | Oadd, I n1 :: I n2 :: nil => I(Int.add n1 n2)
- | Oaddshift s, I n1 :: I n2 :: nil => I(Int.add n1 (eval_static_shift s n2))
- | Oadd, G s1 n1 :: I n2 :: nil => G s1 (Int.add n1 n2)
- | Oaddshift s, G s1 n1 :: I n2 :: nil => G s1 (Int.add n1 (eval_static_shift s n2))
- | Oadd, S n1 :: I n2 :: nil => S (Int.add n1 n2)
- | Oaddshift s, S n1 :: I n2 :: nil => S (Int.add n1 (eval_static_shift s n2))
- | Oadd, I n1 :: G s2 n2 :: nil => G s2 (Int.add n1 n2)
- | Oadd, I n1 :: S n2 :: nil => S (Int.add n1 n2)
- | Oaddimm n, I n1 :: nil => I (Int.add n1 n)
- | Oaddimm n, G s1 n1 :: nil => G s1 (Int.add n1 n)
- | Oaddimm n, S n1 :: nil => S (Int.add n1 n)
- | Osub, I n1 :: I n2 :: nil => I(Int.sub n1 n2)
- | Osubshift s, I n1 :: I n2 :: nil => I(Int.sub n1 (eval_static_shift s n2))
- | Osub, G s1 n1 :: I n2 :: nil => G s1 (Int.sub n1 n2)
- | Osub, S n1 :: I n2 :: nil => S (Int.sub n1 n2)
- | Osubshift s, G s1 n1 :: I n2 :: nil => G s1 (Int.sub n1 (eval_static_shift s n2))
- | Orsubshift s, I n1 :: I n2 :: nil => I(Int.sub (eval_static_shift s n2) n1)
- | Orsubimm n, I n1 :: nil => I (Int.sub n n1)
- | Omul, I n1 :: I n2 :: nil => I(Int.mul n1 n2)
- | Omla, I n1 :: I n2 :: I n3 :: nil => I(Int.add (Int.mul n1 n2) n3)
- | Omulhs, I n1 :: I n2 :: nil => I(Int.mulhs n1 n2)
- | Omulhu, I n1 :: I n2 :: nil => I(Int.mulhu n1 n2)
- | Odiv, I n1 :: I n2 :: nil =>
- if Int.eq n2 Int.zero then Unknown else
- if Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone then Unknown
- else I(Int.divs n1 n2)
- | Odivu, I n1 :: I n2 :: nil =>
- if Int.eq n2 Int.zero then Unknown else I(Int.divu n1 n2)
- | Oand, I n1 :: I n2 :: nil => I(Int.and n1 n2)
- | Oandshift s, I n1 :: I n2 :: nil => I(Int.and n1 (eval_static_shift s n2))
- | Oandimm n, I n1 :: nil => I(Int.and n1 n)
- | Oor, I n1 :: I n2 :: nil => I(Int.or n1 n2)
- | Oorshift s, I n1 :: I n2 :: nil => I(Int.or n1 (eval_static_shift s n2))
- | Oorimm n, I n1 :: nil => I(Int.or n1 n)
- | Oxor, I n1 :: I n2 :: nil => I(Int.xor n1 n2)
- | Oxorshift s, I n1 :: I n2 :: nil => I(Int.xor n1 (eval_static_shift s n2))
- | Oxorimm n, I n1 :: nil => I(Int.xor n1 n)
- | Obic, I n1 :: I n2 :: nil => I(Int.and n1 (Int.not n2))
- | Obicshift s, I n1 :: I n2 :: nil => I(Int.and n1 (Int.not (eval_static_shift s n2)))
- | Onot, I n1 :: nil => I(Int.not n1)
- | Onotshift s, I n1 :: nil => I(Int.not (eval_static_shift s n1))
- | Oshl, I n1 :: I n2 :: nil => if Int.ltu n2 Int.iwordsize then I(Int.shl n1 n2) else Unknown
- | Oshr, I n1 :: I n2 :: nil => if Int.ltu n2 Int.iwordsize then I(Int.shr n1 n2) else Unknown
- | Oshru, I n1 :: I n2 :: nil => if Int.ltu n2 Int.iwordsize then I(Int.shru n1 n2) else Unknown
- | Oshift s, I n1 :: nil => I(eval_static_shift s n1)
- | Onegf, F n1 :: nil => F(Float.neg n1)
- | Oabsf, F n1 :: nil => F(Float.abs n1)
- | Oaddf, F n1 :: F n2 :: nil => F(Float.add n1 n2)
- | Osubf, F n1 :: F n2 :: nil => F(Float.sub n1 n2)
- | Omulf, F n1 :: F n2 :: nil => F(Float.mul n1 n2)
- | Odivf, F n1 :: F n2 :: nil => F(Float.div n1 n2)
- | Osingleoffloat, F n1 :: nil => F(Float.singleoffloat n1)
- | Ointoffloat, F n1 :: nil => eval_static_intoffloat n1
- | Ointuoffloat, F n1 :: nil => eval_static_intuoffloat n1
- | Ofloatofint, I n1 :: nil => if propagate_float_constants tt then F(Float.floatofint n1) else Unknown
- | Ofloatofintu, I n1 :: nil => if propagate_float_constants tt then F(Float.floatofintu n1) else Unknown
- | Omakelong, I n1 :: I n2 :: nil => L(Int64.ofwords n1 n2)
- | Olowlong, L n :: nil => I(Int64.loword n)
- | Ohighlong, L n :: nil => I(Int64.hiword n)
- | Ocmp c, vl => eval_static_condition_val c vl
- | _, _ => Unknown
- end.
-
-Nondetfunction eval_static_addressing (addr: addressing) (vl: list approx) :=
- match addr, vl with
- | Aindexed n, I n1::nil => I (Int.add n1 n)
- | Aindexed n, G id ofs::nil => G id (Int.add ofs n)
- | Aindexed n, S ofs::nil => S (Int.add ofs n)
- | Aindexed2, I n1::I n2::nil => I (Int.add n1 n2)
- | Aindexed2, G id ofs::I n2::nil => G id (Int.add ofs n2)
- | Aindexed2, I n1::G id ofs::nil => G id (Int.add ofs n1)
- | Aindexed2, S ofs::I n2::nil => S (Int.add ofs n2)
- | Aindexed2, I n1::S ofs::nil => S (Int.add ofs n1)
- | Aindexed2shift s, I n1::I n2::nil => I (Int.add n1 (eval_static_shift s n2))
- | Aindexed2shift s, G id ofs::I n2::nil => G id (Int.add ofs (eval_static_shift s n2))
- | Aindexed2shift s, S ofs::I n2::nil => S (Int.add ofs (eval_static_shift s n2))
- | Ainstack ofs, nil => S ofs
- | _, _ => Unknown
- end.
-
-
-(** * Operator strength reduction *)
-
-(** We now define auxiliary functions for strength reduction of
- operators and addressing modes: replacing an operator with a cheaper
- one if some of its arguments are statically known. These are again
- large pattern-matchings expressed in indirect style. *)
-
-Section STRENGTH_REDUCTION.
-
Nondetfunction cond_strength_reduction
- (cond: condition) (args: list reg) (vl: list approx) :=
+ (cond: condition) (args: list reg) (vl: list aval) :=
match cond, args, vl with
| Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
(Ccompimm (swap_comparison c) n1, r2 :: nil)
@@ -272,9 +127,12 @@ Definition make_divuimm (n: int) (r1 r2: reg) :=
| None => (Odivu, r1 :: r2 :: nil)
end.
-Definition make_andimm (n: int) (r: reg) :=
+Definition make_andimm (n: int) (r: reg) (a: aval) :=
if Int.eq n Int.zero then (Ointconst Int.zero, nil)
else if Int.eq n Int.mone then (Omove, r :: nil)
+ else if match a with Uns m => Int.eq (Int.zero_ext m (Int.not n)) Int.zero
+ | _ => false end
+ then (Omove, r :: nil)
else (Oandimm n, r :: nil).
Definition make_orimm (n: int) (r: reg) :=
@@ -292,9 +150,20 @@ Definition make_mulfimm (n: float) (r r1 r2: reg) :=
then (Oaddf, r :: r :: nil)
else (Omulf, r1 :: r2 :: nil).
+Definition make_cast8signed (r: reg) (a: aval) :=
+ if vincl a (Sgn 8) then (Omove, r :: nil) else (Ocast8signed, r :: nil).
+Definition make_cast16signed (r: reg) (a: aval) :=
+ if vincl a (Sgn 16) then (Omove, r :: nil) else (Ocast16signed, r :: nil).
+Definition make_singleoffloat (r: reg) (a: aval) :=
+ if vincl a Fsingle && generate_float_constants tt
+ then (Omove, r :: nil)
+ else (Osingleoffloat, r :: nil).
+
Nondetfunction op_strength_reduction
- (op: operation) (args: list reg) (vl: list approx) :=
+ (op: operation) (args: list reg) (vl: list aval) :=
match op, args, vl with
+ | Ocast8signed, r1 :: nil, v1 :: nil => make_cast8signed r1 v1
+ | Ocast16signed, r1 :: nil, v1 :: nil => make_cast16signed r1 v1
| Oadd, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_addimm n1 r2
| Oadd, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm n2 r1
| Oaddshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm (eval_static_shift s n2) r1
@@ -306,20 +175,21 @@ Nondetfunction op_strength_reduction
| Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_mulimm n2 r1 r2
| Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divimm n2 r1 r2
| Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divuimm n2 r1 r2
- | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_andimm n1 r2
- | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm n2 r1
- | Oandshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (eval_static_shift s n2) r1
+ | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_andimm n1 r2 v2
+ | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm n2 r1 v1
+ | Oandshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (eval_static_shift s n2) r1 v1
| Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_orimm n1 r2
| Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_orimm n2 r1
| Oorshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_orimm (eval_static_shift s n2) r1
| Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_xorimm n1 r2
| Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_xorimm n2 r1
| Oxorshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_xorimm (eval_static_shift s n2) r1
- | Obic, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (Int.not n2) r1
- | Obicshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (Int.not (eval_static_shift s n2)) r1
+ | Obic, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (Int.not n2) r1 v1
+ | Obicshift s, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm (Int.not (eval_static_shift s n2)) r1 v1
| Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shlimm n2 r1 r2
| 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
+ | Osingleoffloat, r1 :: nil, v1 :: nil => make_singleoffloat r1 v1
| Ocmp c, args, vl =>
let (c', args') := cond_strength_reduction c args vl in (Ocmp c', args')
| Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
@@ -329,21 +199,21 @@ Nondetfunction op_strength_reduction
Nondetfunction addr_strength_reduction
- (addr: addressing) (args: list reg) (vl: list approx) :=
+ (addr: addressing) (args: list reg) (vl: list aval) :=
match addr, args, vl with
- | Aindexed2, r1 :: r2 :: nil, S n1 :: I n2 :: nil =>
+ | Aindexed2, r1 :: r2 :: nil, Ptr(Stk n1) :: I n2 :: nil =>
(Ainstack (Int.add n1 n2), nil)
- | Aindexed2, r1 :: r2 :: nil, I n1 :: S n2 :: nil =>
+ | Aindexed2, r1 :: r2 :: nil, I n1 :: Ptr(Stk n2) :: nil =>
(Ainstack (Int.add n1 n2), nil)
| Aindexed2, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
(Aindexed n1, r2 :: nil)
| Aindexed2, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
(Aindexed n2, r1 :: nil)
- | Aindexed2shift s, r1 :: r2 :: nil, S n1 :: I n2 :: nil =>
+ | Aindexed2shift s, r1 :: r2 :: nil, Ptr(Stk n1) :: I n2 :: nil =>
(Ainstack (Int.add n1 (eval_static_shift s n2)), nil)
| Aindexed2shift s, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
(Aindexed (eval_static_shift s n2), r1 :: nil)
- | Aindexed n, r1 :: nil, S n1 :: nil =>
+ | Aindexed n, r1 :: nil, Ptr(Stk n1) :: nil =>
(Ainstack (Int.add n1 n), nil)
| _, _, _ =>
(addr, args)