From 7698300cfe2d3f944ce2e1d4a60a263620487718 Mon Sep 17 00:00:00 2001 From: xleroy Date: Fri, 20 Dec 2013 13:05:53 +0000 Subject: Merge of branch value-analysis. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2381 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- powerpc/CombineOpproof.v | 80 +++++++++++++++++++++++++----------------------- 1 file changed, 42 insertions(+), 38 deletions(-) (limited to 'powerpc/CombineOpproof.v') diff --git a/powerpc/CombineOpproof.v b/powerpc/CombineOpproof.v index 0e328df..4d8fed7 100644 --- a/powerpc/CombineOpproof.v +++ b/powerpc/CombineOpproof.v @@ -21,8 +21,8 @@ Require Import Memory. Require Import Op. Require Import Registers. Require Import RTL. +Require Import CSEdomain. Require Import CombineOp. -Require Import CSE. Section COMBINE. @@ -30,9 +30,21 @@ Variable ge: genv. Variable sp: val. Variable m: mem. Variable get: valnum -> option rhs. -Variable valu: valnum -> val. -Hypothesis get_sound: forall v rhs, get v = Some rhs -> equation_holds valu ge sp m v rhs. +Variable valu: valuation. +Hypothesis get_sound: forall v rhs, get v = Some rhs -> rhs_eval_to valu ge sp m rhs (valu v). +Lemma get_op_sound: + forall v op vl, get v = Some (Op op vl) -> eval_operation ge sp op (map valu vl) m = Some (valu v). +Proof. + intros. exploit get_sound; eauto. intros REV; inv REV; auto. +Qed. + +Ltac UseGetSound := + match goal with + | [ H: get _ = Some _ |- _ ] => + let x := fresh "EQ" in (generalize (get_op_sound _ _ _ H); intros x; simpl in x; FuncInv) + end. + Lemma combine_compimm_ne_0_sound: forall x cond args, combine_compimm_ne_0 get x = Some(cond, args) -> @@ -41,12 +53,11 @@ Lemma combine_compimm_ne_0_sound: Proof. intros until args. functional induction (combine_compimm_ne_0 get x); intros EQ; inv EQ. (* of cmp *) - exploit get_sound; eauto. unfold equation_holds. simpl. intro EQ; inv EQ. + UseGetSound. rewrite <- H. destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto. (* of and *) - exploit get_sound; eauto. unfold equation_holds; simpl. - destruct args; try discriminate. destruct args; try discriminate. simpl. - intros EQ; inv EQ. destruct (valu v); simpl; auto. + UseGetSound. rewrite <- H. + destruct v; simpl; auto. Qed. Lemma combine_compimm_eq_0_sound: @@ -57,13 +68,11 @@ Lemma combine_compimm_eq_0_sound: Proof. intros until args. functional induction (combine_compimm_eq_0 get x); intros EQ; inv EQ. (* of cmp *) - exploit get_sound; eauto. unfold equation_holds. simpl. intro EQ; inv EQ. + UseGetSound. rewrite <- H. rewrite eval_negate_condition. destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto. (* of and *) - exploit get_sound; eauto. unfold equation_holds; simpl. - destruct args; try discriminate. destruct args; try discriminate. simpl. - intros EQ; inv EQ. destruct (valu v); simpl; auto. + UseGetSound. rewrite <- H. destruct v; auto. Qed. Lemma combine_compimm_eq_1_sound: @@ -74,7 +83,7 @@ Lemma combine_compimm_eq_1_sound: 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. + UseGetSound. rewrite <- H. destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto. Qed. @@ -86,7 +95,7 @@ Lemma combine_compimm_ne_1_sound: Proof. intros until args. functional induction (combine_compimm_ne_1 get x); intros EQ; inv EQ. (* of cmp *) - exploit get_sound; eauto. unfold equation_holds. simpl. intro EQ; inv EQ. + UseGetSound. rewrite <- H. rewrite eval_negate_condition. destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto. Qed. @@ -122,8 +131,7 @@ Theorem combine_addr_sound: Proof. intros. functional inversion H; subst. (* indexed - addimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intro EQ. FuncInv. - rewrite <- H0. rewrite Val.add_assoc. auto. + UseGetSound. simpl; rewrite <- H0. rewrite Val.add_assoc. auto. Qed. Theorem combine_op_sound: @@ -133,45 +141,41 @@ Theorem combine_op_sound: Proof. intros. functional inversion H; subst. (* addimm - addimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite Val.add_assoc. auto. + UseGetSound; simpl. rewrite <- H0. rewrite Val.add_assoc. auto. (* addimm - subimm *) Opaque Val.sub. - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. change (Vint (Int.add m0 n)) with (Val.add (Vint m0) (Vint n)). + UseGetSound; simpl. rewrite <- H0. + change (Vint (Int.add m0 n)) with (Val.add (Vint m0) (Vint n)). rewrite Val.sub_add_l. auto. (* subimm - addimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. + UseGetSound; simpl. rewrite <- H0. Transparent Val.sub. destruct v; simpl; auto. repeat rewrite Int.sub_add_opp. rewrite Int.add_assoc. rewrite Int.neg_add_distr. decEq. decEq. decEq. apply Int.add_commut. (* andimm - andimm *) + UseGetSound; simpl. generalize (Int.eq_spec p m0); rewrite H7; intros. - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H2. rewrite Val.and_assoc. simpl. fold p. rewrite H0. auto. - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite Val.and_assoc. auto. + rewrite <- H0. rewrite Val.and_assoc. simpl. fold p. rewrite H1. auto. + UseGetSound; simpl. + rewrite <- H0. rewrite Val.and_assoc. auto. (* andimm - rolm *) + UseGetSound; simpl. generalize (Int.eq_spec p m0); rewrite H7; intros. - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H2. destruct v; simpl; auto. unfold Int.rolm. - rewrite Int.and_assoc. fold p; rewrite H0. auto. - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. destruct v; simpl; auto. unfold Int.rolm. rewrite Int.and_assoc. auto. + rewrite <- H0. destruct v; simpl; auto. unfold Int.rolm. + rewrite Int.and_assoc. fold p. rewrite H1. auto. + UseGetSound; simpl. + rewrite <- H0. destruct v; simpl; auto. unfold Int.rolm. + rewrite Int.and_assoc. auto. (* orimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite Val.or_assoc. auto. + UseGetSound; simpl. rewrite <- H0. rewrite Val.or_assoc. auto. (* xorimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite Val.xor_assoc. auto. + UseGetSound; simpl. rewrite <- H0. rewrite Val.xor_assoc. auto. (* rolm - andimm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite <- Val.rolm_zero. rewrite Val.rolm_rolm. + UseGetSound; simpl. rewrite <- H0. + rewrite <- Val.rolm_zero. rewrite Val.rolm_rolm. rewrite (Int.add_commut Int.zero). rewrite Int.add_zero. auto. (* rolm - rolm *) - exploit get_sound; eauto. unfold equation_holds; simpl; intros. FuncInv. - rewrite <- H1. rewrite Val.rolm_rolm. auto. + UseGetSound; simpl. rewrite <- H0. rewrite Val.rolm_rolm. auto. (* cmp *) simpl. decEq; decEq. eapply combine_cond_sound; eauto. Qed. -- cgit v1.2.3