From 4d542bc7eafadb16b845cf05d1eb4988eb55ed0f Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Tue, 20 Oct 2015 13:32:18 +0200 Subject: Updated PR by removing whitespaces. Bug 17450. --- powerpc/CombineOpproof.v | 46 +++++++++++++++++++++++----------------------- 1 file changed, 23 insertions(+), 23 deletions(-) (limited to 'powerpc/CombineOpproof.v') diff --git a/powerpc/CombineOpproof.v b/powerpc/CombineOpproof.v index 4d8fed78..4883876d 100644 --- a/powerpc/CombineOpproof.v +++ b/powerpc/CombineOpproof.v @@ -10,7 +10,7 @@ (* *) (* *********************************************************************) -(** Recognition of combined operations, addressing modes and conditions +(** Recognition of combined operations, addressing modes and conditions during the [CSE] phase. *) Require Import Coqlib. @@ -36,7 +36,7 @@ Hypothesis get_sound: forall v rhs, get v = Some rhs -> rhs_eval_to valu ge sp m 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. + intros. exploit get_sound; eauto. intros REV; inv REV; auto. Qed. Ltac UseGetSound := @@ -44,7 +44,7 @@ Ltac UseGetSound := | [ 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) -> @@ -53,11 +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 *) - UseGetSound. rewrite <- H. + UseGetSound. rewrite <- H. destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto. (* of and *) - UseGetSound. rewrite <- H. - destruct v; simpl; auto. + UseGetSound. rewrite <- H. + destruct v; simpl; auto. Qed. Lemma combine_compimm_eq_0_sound: @@ -69,10 +69,10 @@ Proof. intros until args. functional induction (combine_compimm_eq_0 get x); intros EQ; inv EQ. (* of cmp *) UseGetSound. rewrite <- H. - rewrite eval_negate_condition. + rewrite eval_negate_condition. destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto. (* of and *) - UseGetSound. rewrite <- H. destruct v; auto. + UseGetSound. rewrite <- H. destruct v; auto. Qed. Lemma combine_compimm_eq_1_sound: @@ -83,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 *) - UseGetSound. rewrite <- H. + UseGetSound. rewrite <- H. destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto. Qed. @@ -95,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 *) - UseGetSound. rewrite <- H. + UseGetSound. rewrite <- H. rewrite eval_negate_condition. destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto. Qed. @@ -131,7 +131,7 @@ Theorem combine_addr_sound: Proof. intros. functional inversion H; subst. (* indexed - addimm *) - UseGetSound. simpl; rewrite <- H0. rewrite Val.add_assoc. auto. + UseGetSound. simpl; rewrite <- H0. rewrite Val.add_assoc. auto. Qed. Theorem combine_op_sound: @@ -144,27 +144,27 @@ Proof. UseGetSound; simpl. rewrite <- H0. rewrite Val.add_assoc. auto. (* addimm - subimm *) Opaque Val.sub. - UseGetSound; simpl. rewrite <- H0. + 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 *) - UseGetSound; simpl. rewrite <- H0. + 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. + UseGetSound; simpl. + generalize (Int.eq_spec p m0); rewrite H7; intros. rewrite <- H0. rewrite Val.and_assoc. simpl. fold p. rewrite H1. auto. - UseGetSound; simpl. + UseGetSound; simpl. rewrite <- H0. rewrite Val.and_assoc. auto. (* andimm - rolm *) - UseGetSound; simpl. - generalize (Int.eq_spec p m0); rewrite H7; intros. - rewrite <- H0. destruct v; simpl; auto. unfold Int.rolm. + UseGetSound; simpl. + generalize (Int.eq_spec p m0); rewrite H7; intros. + 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. + UseGetSound; simpl. + rewrite <- H0. destruct v; simpl; auto. unfold Int.rolm. rewrite Int.and_assoc. auto. (* orimm *) UseGetSound; simpl. rewrite <- H0. rewrite Val.or_assoc. auto. @@ -172,10 +172,10 @@ Transparent Val.sub. UseGetSound; simpl. rewrite <- H0. rewrite Val.xor_assoc. auto. (* rolm - andimm *) UseGetSound; simpl. rewrite <- H0. - rewrite <- Val.rolm_zero. rewrite Val.rolm_rolm. + rewrite <- Val.rolm_zero. rewrite Val.rolm_rolm. rewrite (Int.add_commut Int.zero). rewrite Int.add_zero. auto. (* rolm - rolm *) - UseGetSound; simpl. rewrite <- H0. 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