aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/ExtValues.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/ExtValues.v
parent3ef3e6c78026cc1d5793ccb4e905a0232ec7bb4e (diff)
downloadcompcert-kvx-2c428ad4e0177756db2f6dfe56831b5a44f6de5c.tar.gz
compcert-kvx-2c428ad4e0177756db2f6dfe56831b5a44f6de5c.zip
add with shift, beginning
Diffstat (limited to 'mppa_k1c/ExtValues.v')
-rw-r--r--mppa_k1c/ExtValues.v24
1 files changed, 24 insertions, 0 deletions
diff --git a/mppa_k1c/ExtValues.v b/mppa_k1c/ExtValues.v
index 735d5c3c..32d84b60 100644
--- a/mppa_k1c/ExtValues.v
+++ b/mppa_k1c/ExtValues.v
@@ -13,6 +13,28 @@ Definition z_of_shift1_4 (x : shift1_4) :=
| SHIFT4 => 4
end.
+Definition shift1_4_of_z (x : Z) :=
+ if Z.eq_dec x 1 then Some SHIFT1
+ else if Z.eq_dec x 2 then Some SHIFT2
+ else if Z.eq_dec x 3 then Some SHIFT3
+ else if Z.eq_dec x 4 then Some SHIFT4
+ else None.
+
+Lemma shift1_4_of_z_correct :
+ forall z,
+ match shift1_4_of_z z with
+ | Some x => z_of_shift1_4 x = z
+ | None => True
+ end.
+Proof.
+ intro. unfold shift1_4_of_z.
+ destruct (Z.eq_dec _ _); simpl; try congruence.
+ destruct (Z.eq_dec _ _); simpl; try congruence.
+ destruct (Z.eq_dec _ _); simpl; try congruence.
+ destruct (Z.eq_dec _ _); simpl; try congruence.
+ trivial.
+Qed.
+
Definition int_of_shift1_4 (x : shift1_4) :=
Int.repr (z_of_shift1_4 x).
@@ -178,3 +200,5 @@ Proof.
apply Int64.neg_mul_distr_r.
Qed.
+Definition addx sh v1 v2 :=
+ Val.add v2 (Val.shl v1 (Vint sh)). \ No newline at end of file