aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/SelectOpproof.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-09-21 09:17:05 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-09-21 19:39:46 +0200
commitcae21e8816db863bf99f87469c9680e150d28960 (patch)
tree19fdf188238300245e0e5e7a9b78a94bbb53df93 /riscV/SelectOpproof.v
parent52b4f973646c3b79804fcdddeed5325ab1f3ce7d (diff)
downloadcompcert-kvx-cae21e8816db863bf99f87469c9680e150d28960.tar.gz
compcert-kvx-cae21e8816db863bf99f87469c9680e150d28960.zip
maketotal mod & div
Diffstat (limited to 'riscV/SelectOpproof.v')
-rw-r--r--riscV/SelectOpproof.v62
1 files changed, 30 insertions, 32 deletions
diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v
index 7f2014dc..436cd4f3 100644
--- a/riscV/SelectOpproof.v
+++ b/riscV/SelectOpproof.v
@@ -506,7 +506,12 @@ Theorem eval_divs_base:
Val.divs x y = Some z ->
exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v.
Proof.
- intros. unfold divs_base. exists z; split. EvalOp. auto.
+ intros. unfold divs_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_mods_base:
@@ -516,7 +521,12 @@ Theorem eval_mods_base:
Val.mods x y = Some z ->
exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v.
Proof.
- intros. unfold mods_base. exists z; split. EvalOp. auto.
+ intros. unfold mods_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_divu_base:
@@ -526,7 +536,12 @@ Theorem eval_divu_base:
Val.divu x y = Some z ->
exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v.
Proof.
- intros. unfold divu_base. exists z; split. EvalOp. auto.
+ intros. unfold divu_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_modu_base:
@@ -536,7 +551,12 @@ Theorem eval_modu_base:
Val.modu x y = Some z ->
exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v.
Proof.
- intros. unfold modu_base. exists z; split. EvalOp. auto.
+ intros. unfold modu_base. exists z; split. EvalOp.
+ 2: apply Val.lessdef_refl.
+ cbn.
+ rewrite H1.
+ cbn.
+ trivial.
Qed.
Theorem eval_shrximm:
@@ -553,34 +573,12 @@ Proof.
replace (Int.shrx i Int.zero) with i. auto.
unfold Int.shrx, Int.divs. rewrite Int.shl_zero.
change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed; auto.
- econstructor; split. EvalOp. auto.
-(*
- intros. destruct x; simpl in H0; try discriminate.
- destruct (Int.ltu n (Int.repr 31)) eqn:LTU; inv H0.
- unfold shrximm.
- predSpec Int.eq Int.eq_spec n Int.zero.
- - subst n. exists (Vint i); split; auto.
- unfold Int.shrx, Int.divs. rewrite Z.quot_1_r. rewrite Int.repr_signed. auto.
- - assert (NZ: Int.unsigned n <> 0).
- { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. }
- assert (LT: 0 <= Int.unsigned n < 31) by (apply Int.ltu_inv in LTU; assumption).
- assert (LTU2: Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true).
- { unfold Int.ltu; apply zlt_true.
- unfold Int.sub. change (Int.unsigned Int.iwordsize) with 32.
- rewrite Int.unsigned_repr. omega.
- assert (32 < Int.max_unsigned) by reflexivity. omega. }
- assert (X: eval_expr ge sp e m le
- (Eop (Oshrimm (Int.repr (Int.zwordsize - 1))) (a ::: Enil))
- (Vint (Int.shr i (Int.repr (Int.zwordsize - 1))))).
- { EvalOp. }
- assert (Y: eval_expr ge sp e m le (shrximm_inner a n)
- (Vint (Int.shru (Int.shr i (Int.repr (Int.zwordsize - 1))) (Int.sub Int.iwordsize n)))).
- { EvalOp. simpl. rewrite LTU2. auto. }
- TrivialExists.
- constructor. EvalOp. simpl; eauto. constructor.
- simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int.shrx_shr_2 by auto. reflexivity.
- change (Int.unsigned Int.iwordsize) with 32; omega.
-*)
+ econstructor; split. EvalOp.
+ cbn.
+ rewrite H0.
+ cbn.
+ reflexivity.
+ apply Val.lessdef_refl.
Qed.
Theorem eval_shl: binary_constructor_sound shl Val.shl.