aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/SelectLong.vp
diff options
context:
space:
mode:
Diffstat (limited to 'mppa_k1c/SelectLong.vp')
-rw-r--r--mppa_k1c/SelectLong.vp48
1 files changed, 47 insertions, 1 deletions
diff --git a/mppa_k1c/SelectLong.vp b/mppa_k1c/SelectLong.vp
index 811a8ab1..717b0120 100644
--- a/mppa_k1c/SelectLong.vp
+++ b/mppa_k1c/SelectLong.vp
@@ -23,6 +23,8 @@ Require Import AST Integers Floats.
Require Import Op CminorSel.
Require Import OpHelpers.
Require Import SelectOp SplitLong.
+Require Import ExtValues.
+Require Import DecBoolOps.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
@@ -68,7 +70,10 @@ Nondetfunction addlimm (n: int64) (e: expr) :=
if Int64.eq n Int64.zero then e else
match e with
| Eop (Olongconst m) Enil => longconst (Int64.add n m)
- | Eop (Oaddrsymbol s m) Enil => Eop (Oaddrsymbol s (Ptrofs.add (Ptrofs.of_int64 n) m)) Enil
+ | Eop (Oaddrsymbol s m) Enil =>
+ (if Compopts.optim_fglobaladdroffset tt
+ then Eop (Oaddrsymbol s (Ptrofs.add (Ptrofs.of_int64 n) m)) Enil
+ else Eop (Oaddlimm n) (e ::: Enil))
| Eop (Oaddrstack m) Enil => Eop (Oaddrstack (Ptrofs.add (Ptrofs.of_int64 n) m)) Enil
| Eop (Oaddlimm m) (t ::: Enil) => Eop (Oaddlimm(Int64.add n m)) (t ::: Enil)
| _ => Eop (Oaddlimm n) (e ::: Enil)
@@ -155,6 +160,12 @@ Nondetfunction shrluimm (e1: expr) (n: int) :=
if Int.ltu (Int.add n n1) Int64.iwordsize'
then Eop (Oshrluimm (Int.add n n1)) (t1:::Enil)
else Eop (Oshrluimm n) (e1:::Enil)
+ | Eop (Oshllimm n1) (t1:::Enil) =>
+ let stop := Z.sub Int64.zwordsize (Z.add (Int.unsigned n1) Z.one) in
+ let start := Z.sub (Z.add (Z.add (Int.unsigned n) stop) Z.one) Int64.zwordsize in
+ if is_bitfieldl stop start
+ then Eop (Oextfzl stop start) (t1:::Enil)
+ else Eop (Oshrluimm n) (e1:::Enil)
| _ =>
Eop (Oshrluimm n) (e1:::Enil)
end.
@@ -172,6 +183,12 @@ Nondetfunction shrlimm (e1: expr) (n: int) :=
if Int.ltu (Int.add n n1) Int64.iwordsize'
then Eop (Oshrlimm (Int.add n n1)) (t1:::Enil)
else Eop (Oshrlimm n) (e1:::Enil)
+ | Eop (Oshllimm n1) (t1:::Enil) =>
+ let stop := Z.sub Int64.zwordsize (Z.add (Int.unsigned n1) Z.one) in
+ let start := Z.sub (Z.add (Z.add (Int.unsigned n) stop) Z.one) Int64.zwordsize in
+ if is_bitfieldl stop start
+ then Eop (Oextfsl stop start) (t1:::Enil)
+ else Eop (Oshrlimm n) (e1:::Enil)
| _ =>
Eop (Oshrlimm n) (e1:::Enil)
end.
@@ -298,6 +315,31 @@ Nondetfunction orl (e1: expr) (e2: expr) :=
&& Int64.eq zero1 Int64.zero
then Eop (Oselectl (Ccompl0 Cne)) (v0:::v1:::y0:::Enil)
else Eop Oorl (e1:::e2:::Enil)
+ | (Eop (Oandlimm nmask) (prev:::Enil)),
+ (Eop (Oandlimm mask)
+ ((Eop (Oshllimm start) (fld:::Enil)):::Enil)) =>
+ let zstart := Int.unsigned start in
+ let zstop := int64_highest_bit mask in
+ if is_bitfieldl zstop zstart
+ then
+ let mask' := Int64.repr (zbitfield_mask zstop zstart) in
+ if and_dec (Int64.eq_dec mask mask')
+ (Int64.eq_dec nmask (Int64.not mask'))
+ then Eop (Oinsfl zstop zstart) (prev:::fld:::Enil)
+ else Eop Oorl (e1:::e2:::Enil)
+ else Eop Oorl (e1:::e2:::Enil)
+ | (Eop (Oandlimm nmask) (prev:::Enil)),
+ (Eop (Oandlimm mask) (fld:::Enil)) =>
+ let zstart := 0 in
+ let zstop := int64_highest_bit mask in
+ if is_bitfieldl zstop zstart
+ then
+ let mask' := Int64.repr (zbitfield_mask zstop zstart) in
+ if and_dec (Int64.eq_dec mask mask')
+ (Int64.eq_dec nmask (Int64.not mask'))
+ then Eop (Oinsfl zstop zstart) (prev:::fld:::Enil)
+ else Eop Oorl (e1:::e2:::Enil)
+ else Eop Oorl (e1:::e2:::Enil)
| _, _ => Eop Oorl (e1:::e2:::Enil)
end.
@@ -408,3 +450,7 @@ Definition singleoflong (e: expr) := SplitLong.singleoflong e.
Definition singleoflongu (e: expr) := SplitLong.singleoflongu e.
End SELECT.
+
+(* Local Variables: *)
+(* mode: coq *)
+(* End: *) \ No newline at end of file