aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Op.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-19 21:31:20 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-19 21:31:20 +0100
commit9adb576998f3b2017db5c062b459449e1721579a (patch)
treefd70a6ce92921404c91b7f9f494e219da4ae3945 /mppa_k1c/Op.v
parentae571e2467e977f03044d750568f6528d8d64e43 (diff)
downloadcompcert-kvx-9adb576998f3b2017db5c062b459449e1721579a.tar.gz
compcert-kvx-9adb576998f3b2017db5c062b459449e1721579a.zip
mul immediate
Diffstat (limited to 'mppa_k1c/Op.v')
-rw-r--r--mppa_k1c/Op.v5
1 files changed, 5 insertions, 0 deletions
diff --git a/mppa_k1c/Op.v b/mppa_k1c/Op.v
index 04081158..0e9a7af9 100644
--- a/mppa_k1c/Op.v
+++ b/mppa_k1c/Op.v
@@ -115,6 +115,7 @@ Inductive operation : Type :=
| Onegl (**r [rd = - r1] *)
| Osubl (**r [rd = r1 - r2] *)
| Omull (**r [rd = r1 * r2] *)
+ | Omullimm (n: int64) (**r [rd = r1 * n] *)
| Omullhs (**r [rd = high part of r1 * r2, signed] *)
| Omullhu (**r [rd = high part of r1 * r2, unsigned] *)
| Odivl (**r [rd = r1 / r2] (signed) *)
@@ -314,6 +315,7 @@ Definition eval_operation
| Onegl, v1::nil => Some (Val.negl v1)
| Osubl, v1::v2::nil => Some (Val.subl v1 v2)
| Omull, v1::v2::nil => Some (Val.mull v1 v2)
+ | Omullimm n, v1::nil => Some (Val.mull v1 (Vlong n))
| Omullhs, v1::v2::nil => Some (Val.mullhs v1 v2)
| Omullhu, v1::v2::nil => Some (Val.mullhu v1 v2)
| Odivl, v1::v2::nil => Val.divls v1 v2
@@ -503,6 +505,7 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Onegl => (Tlong :: nil, Tlong)
| Osubl => (Tlong :: Tlong :: nil, Tlong)
| Omull => (Tlong :: Tlong :: nil, Tlong)
+ | Omullimm _ => (Tlong :: nil, Tlong)
| Omullhs => (Tlong :: Tlong :: nil, Tlong)
| Omullhu => (Tlong :: Tlong :: nil, Tlong)
| Odivl => (Tlong :: Tlong :: nil, Tlong)
@@ -700,6 +703,7 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
destruct (eq_block b b0)...
(* mull, mullhs, mullhu *)
- destruct v0; destruct v1...
+ - destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
(* divl, divlu *)
@@ -1205,6 +1209,7 @@ Proof.
- apply Val.subl_inject; auto.
(* mull, mullhs, mullhu *)
- inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
(* divl, divlu *)