summaryrefslogtreecommitdiff
path: root/arm/CombineOpproof.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-07-01 07:51:12 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-07-01 07:51:12 +0000
commit4abaa3fa312fcbe7ee9665853f52a5d37e703864 (patch)
tree8fb982f2dd6fca8990708aa30a04e537c38ab395 /arm/CombineOpproof.v
parent9b8f0f6c4683dd00ee0e3422b84c9cc34510011e (diff)
Recombine x = cmp(...); if (x == 1) ...
and x = cmp(...); if (x != 1) ... git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1946 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'arm/CombineOpproof.v')
-rw-r--r--arm/CombineOpproof.v29
1 files changed, 29 insertions, 0 deletions
diff --git a/arm/CombineOpproof.v b/arm/CombineOpproof.v
index 9d297ac..0d2e6dd 100644
--- a/arm/CombineOpproof.v
+++ b/arm/CombineOpproof.v
@@ -58,6 +58,31 @@ Proof.
destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
Qed.
+Lemma combine_compimm_eq_1_sound:
+ forall x cond args,
+ combine_compimm_eq_1 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Ceq (valu x) (Vint Int.one) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Ceq (valu x) (Vint Int.one).
+Proof.
+ intros until args. functional induction (combine_compimm_eq_1 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ exploit get_sound; eauto. unfold equation_holds. simpl. intro EQ; inv EQ.
+ destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
+Lemma combine_compimm_ne_1_sound:
+ forall x cond args,
+ combine_compimm_ne_1 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Cne (valu x) (Vint Int.one) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Cne (valu x) (Vint Int.one).
+Proof.
+ intros until args. functional induction (combine_compimm_eq_1 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ exploit get_sound; eauto. unfold equation_holds. simpl. intro EQ; inv EQ.
+ rewrite eval_negate_condition.
+ destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
Theorem combine_cond_sound:
forall cond args cond' args',
combine_cond get cond args = Some(cond', args') ->
@@ -68,10 +93,14 @@ Proof.
simpl; eapply combine_compimm_ne_0_sound; eauto.
(* compimm eq zero *)
simpl; eapply combine_compimm_eq_0_sound; eauto.
+ (* compimm eq one *)
+ simpl; eapply combine_compimm_eq_1_sound; eauto.
(* compuimm ne zero *)
simpl; eapply combine_compimm_ne_0_sound; eauto.
(* compuimm eq zero *)
simpl; eapply combine_compimm_eq_0_sound; eauto.
+ (* compuimm eq one *)
+ simpl; eapply combine_compimm_eq_1_sound; eauto.
Qed.
Theorem combine_addr_sound: