summaryrefslogtreecommitdiff
path: root/powerpc/Op.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-12-20 13:05:53 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-12-20 13:05:53 +0000
commit7698300cfe2d3f944ce2e1d4a60a263620487718 (patch)
tree74292bb5f65bc797906b5c768df2e2e937e869b6 /powerpc/Op.v
parentc511207bd0f25c4199770233175924a725526bd3 (diff)
Merge of branch value-analysis.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2381 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc/Op.v')
-rw-r--r--powerpc/Op.v10
1 files changed, 6 insertions, 4 deletions
diff --git a/powerpc/Op.v b/powerpc/Op.v
index dbc474e..3545b18 100644
--- a/powerpc/Op.v
+++ b/powerpc/Op.v
@@ -166,8 +166,8 @@ Definition eval_condition (cond: condition) (vl: list val) (m: mem): option bool
| Ccompuimm c n, v1 :: nil => Val.cmpu_bool (Mem.valid_pointer m) c v1 (Vint n)
| Ccompf c, v1 :: v2 :: nil => Val.cmpf_bool c v1 v2
| Cnotcompf c, v1 :: v2 :: nil => option_map negb (Val.cmpf_bool c v1 v2)
- | Cmaskzero n, Vint n1 :: nil => Some (Int.eq (Int.and n1 n) Int.zero)
- | Cmasknotzero n, Vint n1 :: nil => Some (negb (Int.eq (Int.and n1 n) Int.zero))
+ | Cmaskzero n, v1 :: nil => Val.maskzero_bool v1 n
+ | Cmasknotzero n, v1 :: nil => option_map negb (Val.maskzero_bool v1 n)
| _, _ => None
end.
@@ -452,8 +452,8 @@ Proof.
repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
repeat (destruct vl; auto).
repeat (destruct vl; auto). destruct (Val.cmpf_bool c v v0); auto. destruct b; auto.
- destruct vl; auto. destruct v; auto. destruct vl; auto.
- destruct vl; auto. destruct v; auto. destruct vl; auto. simpl. rewrite negb_involutive. auto.
+ repeat (destruct vl; auto).
+ repeat (destruct vl; auto). destruct (Val.maskzero_bool v i) as [[]|]; auto.
Qed.
(** Shifting stack-relative references. This is used in [Stacking]. *)
@@ -747,6 +747,8 @@ Proof.
eauto 3 using val_cmpu_bool_inject, Mem.valid_pointer_implies.
inv H3; inv H2; simpl in H0; inv H0; auto.
inv H3; inv H2; simpl in H0; inv H0; auto.
+ inv H3; try discriminate; auto.
+ inv H3; try discriminate; auto.
Qed.
Ltac TrivialExists :=