summaryrefslogtreecommitdiff
path: root/powerpc/Op.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-05-26 07:32:01 +0000
commit132e36fa0be63eb5672eda9168403d3fb74af2fa (patch)
tree33955e0ccb4210271c82326b941523e6e4b2c289 /powerpc/Op.v
parent9ea00d39bb32c1f188f1af2745c3368da6a349c1 (diff)
CSE: add recognition of some combined operators, conditions, and addressing modes (cf. CombineOp.v)
Memory model: cleaning up Memdata Inlining and new Constprop: updated for ARM. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1902 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc/Op.v')
-rw-r--r--powerpc/Op.v24
1 files changed, 11 insertions, 13 deletions
diff --git a/powerpc/Op.v b/powerpc/Op.v
index 76c426b..986ea8c 100644
--- a/powerpc/Op.v
+++ b/powerpc/Op.v
@@ -420,20 +420,18 @@ Definition negate_condition (cond: condition): condition :=
end.
Lemma eval_negate_condition:
- forall cond vl m b,
- eval_condition cond vl m = Some b ->
- eval_condition (negate_condition cond) vl m = Some (negb b).
+ forall cond vl m,
+ eval_condition (negate_condition cond) vl m = option_map negb (eval_condition cond vl m).
Proof.
- intros.
- destruct cond; simpl in H; FuncInv; simpl.
- rewrite Val.negate_cmp_bool; rewrite H; auto.
- rewrite Val.negate_cmpu_bool; rewrite H; auto.
- rewrite Val.negate_cmp_bool; rewrite H; auto.
- rewrite Val.negate_cmpu_bool; rewrite H; auto.
- rewrite H; auto.
- destruct (Val.cmpf_bool c v v0); simpl in H; inv H. rewrite negb_elim; auto.
- rewrite H0; auto.
- rewrite <- H0. rewrite negb_elim; auto.
+ intros. destruct cond; simpl.
+ repeat (destruct vl; auto). apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmp_bool.
+ 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.
Qed.
(** Shifting stack-relative references. This is used in [Stacking]. *)