aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/SelectOpproof.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-06-03 20:09:19 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-06-03 20:11:48 +0200
commit677771949bd62f7f5bbcf99bba6b6f816f07a6c2 (patch)
tree597b2ccc5867bc2bbb083c4e13f792671b2042c1 /riscV/SelectOpproof.v
parent36e64ee96ded0c94c83da6fb12202c276e66ba45 (diff)
parentb7e0d70de2ace6f0a22f9f65cc244d875ee48496 (diff)
downloadcompcert-kvx-677771949bd62f7f5bbcf99bba6b6f816f07a6c2.tar.gz
compcert-kvx-677771949bd62f7f5bbcf99bba6b6f816f07a6c2.zip
Merge branch 'if-conversion' of https://github.com/AbsInt/CompCert into mppa-if-conversion
Diffstat (limited to 'riscV/SelectOpproof.v')
-rw-r--r--riscV/SelectOpproof.v19
1 files changed, 17 insertions, 2 deletions
diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v
index 9966305c..46cc1bbc 100644
--- a/riscV/SelectOpproof.v
+++ b/riscV/SelectOpproof.v
@@ -20,6 +20,7 @@
Require Import Coqlib.
Require Import Maps.
Require Import AST.
+Require Import Zbits.
Require Import Integers.
Require Import Floats.
Require Import Values.
@@ -376,7 +377,7 @@ Proof.
change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl.
apply Val.lessdef_same. f_equal.
transitivity (Int.repr (Z.shiftr (Int.signed i * Int.signed i0) 32)).
- unfold Int.mulhs; f_equal. rewrite Int.Zshiftr_div_two_p by omega. reflexivity.
+ unfold Int.mulhs; f_equal. rewrite Zshiftr_div_two_p by omega. reflexivity.
apply Int.same_bits_eq; intros n N.
change Int.zwordsize with 32 in *.
assert (N1: 0 <= n < 64) by omega.
@@ -404,7 +405,7 @@ Proof.
change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl.
apply Val.lessdef_same. f_equal.
transitivity (Int.repr (Z.shiftr (Int.unsigned i * Int.unsigned i0) 32)).
- unfold Int.mulhu; f_equal. rewrite Int.Zshiftr_div_two_p by omega. reflexivity.
+ unfold Int.mulhu; f_equal. rewrite Zshiftr_div_two_p by omega. reflexivity.
apply Int.same_bits_eq; intros n N.
change Int.zwordsize with 32 in *.
assert (N1: 0 <= n < 64) by omega.
@@ -876,6 +877,20 @@ Proof.
red; intros. unfold floatofsingle. TrivialExists.
Qed.
+Theorem eval_select:
+ forall le ty cond al vl a1 v1 a2 v2 a b,
+ select ty cond al a1 a2 = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ eval_expr ge sp e m le a1 v1 ->
+ eval_expr ge sp e m le a2 v2 ->
+ eval_condition cond vl m = Some b ->
+ exists v,
+ eval_expr ge sp e m le a v
+ /\ Val.lessdef (Val.select (Some b) v1 v2 ty) v.
+Proof.
+ unfold select; intros; discriminate.
+Qed.
+
Theorem eval_addressing:
forall le chunk a v b ofs,
eval_expr ge sp e m le a v ->