aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Op.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-05-11 15:20:02 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-05-11 15:20:02 +0200
commit2c428ad4e0177756db2f6dfe56831b5a44f6de5c (patch)
tree1e209b6c935e36005b31542345ac8bb593361452 /mppa_k1c/Op.v
parent3ef3e6c78026cc1d5793ccb4e905a0232ec7bb4e (diff)
downloadcompcert-kvx-2c428ad4e0177756db2f6dfe56831b5a44f6de5c.tar.gz
compcert-kvx-2c428ad4e0177756db2f6dfe56831b5a44f6de5c.zip
add with shift, beginning
Diffstat (limited to 'mppa_k1c/Op.v')
-rw-r--r--mppa_k1c/Op.v7
1 files changed, 4 insertions, 3 deletions
diff --git a/mppa_k1c/Op.v b/mppa_k1c/Op.v
index ac40c293..69620934 100644
--- a/mppa_k1c/Op.v
+++ b/mppa_k1c/Op.v
@@ -403,8 +403,8 @@ Definition eval_operation
| Ocast16signed, v1 :: nil => Some (Val.sign_ext 16 v1)
| Oadd, v1 :: v2 :: nil => Some (Val.add v1 v2)
| Oaddimm n, v1 :: nil => Some (Val.add v1 (Vint n))
- | Oaddx shift, v1 :: v2 :: nil => Some (Val.add v2 (Val.shl v1 (Vint (int_of_shift1_4 shift))))
- | Oaddximm shift n, v1 :: nil => Some (Val.add (Vint n) (Val.shl v1 (Vint (int_of_shift1_4 shift))))
+ | Oaddx s14, v1 :: v2 :: nil => Some (addx (int_of_shift1_4 s14) v1 v2)
+ | Oaddximm s14 n, v1 :: nil => Some (addx (int_of_shift1_4 s14) v1 (Vint n))
| Oneg, v1 :: nil => Some (Val.neg v1)
| Osub, v1 :: v2 :: nil => Some (Val.sub v1 v2)
| Orevsubimm n, v1 :: nil => Some (Val.sub (Vint n) v1)
@@ -1446,7 +1446,8 @@ Proof.
(* addx, addximm *)
- apply Val.add_inject; trivial.
inv H4; inv H2; simpl; try destruct (Int.ltu _ _); simpl; auto.
- - inv H4; simpl; trivial. try destruct (Int.ltu _ _); simpl; auto.
+ - inv H4; simpl; trivial.
+ destruct (Int.ltu _ _); simpl; trivial.
(* neg, sub *)
- inv H4; simpl; auto.
- apply Val.sub_inject; auto.