summaryrefslogtreecommitdiff
path: root/arm/Asmgenproof1.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-09-02 12:42:19 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-09-02 12:42:19 +0000
commit265fa07b34a813ba9d8249ddad82d71e6002c10d (patch)
tree45831b1793c7920b10969fc7cf6316c202d78e91 /arm/Asmgenproof1.v
parent94470fb6a652cb993982269fcb7a0e8319b54488 (diff)
Merge of the reuse-temps branch:
- Reload temporaries are marked as destroyed (set to Vundef) across operations in the semantics of LTL, LTLin, Linear and Mach, allowing Asmgen to reuse them. - Added IA32 port. - Cleaned up float conversions and axiomatization of floats. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1499 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'arm/Asmgenproof1.v')
-rw-r--r--arm/Asmgenproof1.v1077
1 files changed, 503 insertions, 574 deletions
diff --git a/arm/Asmgenproof1.v b/arm/Asmgenproof1.v
index fc2ce7f..c10c9df 100644
--- a/arm/Asmgenproof1.v
+++ b/arm/Asmgenproof1.v
@@ -27,7 +27,7 @@ Require Import Machconcr.
Require Import Machtyping.
Require Import Asm.
Require Import Asmgen.
-Require Conventions.
+Require Import Conventions.
(** * Correspondence between Mach registers and PPC registers *)
@@ -41,91 +41,86 @@ Proof.
destruct r1; destruct r2; simpl; intros; reflexivity || discriminate.
Qed.
-(** Characterization of PPC registers that correspond to Mach registers. *)
-
-Definition is_data_reg (r: preg) : Prop :=
- match r with
- | IR IR14 => False
- | CR _ => False
- | PC => False
- | _ => True
- end.
-
-Lemma ireg_of_is_data_reg:
- forall (r: mreg), is_data_reg (ireg_of r).
-Proof.
- destruct r; exact I.
-Qed.
-
-Lemma freg_of_is_data_reg:
- forall (r: mreg), is_data_reg (ireg_of r).
-Proof.
- destruct r; exact I.
-Qed.
-
-Lemma preg_of_is_data_reg:
- forall (r: mreg), is_data_reg (preg_of r).
-Proof.
- destruct r; exact I.
-Qed.
-
Lemma ireg_of_not_IR13:
forall r, ireg_of r <> IR13.
Proof.
- intro. case r; discriminate.
+ destruct r; simpl; congruence.
Qed.
+
Lemma ireg_of_not_IR14:
forall r, ireg_of r <> IR14.
Proof.
- intro. case r; discriminate.
+ destruct r; simpl; congruence.
Qed.
-Hint Resolve ireg_of_not_IR13 ireg_of_not_IR14: ppcgen.
+Lemma preg_of_not_IR13:
+ forall r, preg_of r <> IR13.
+Proof.
+ unfold preg_of; intros. destruct (mreg_type r).
+ generalize (ireg_of_not_IR13 r); congruence.
+ congruence.
+Qed.
-Lemma preg_of_not:
- forall r1 r2, ~(is_data_reg r2) -> preg_of r1 <> r2.
+Lemma preg_of_not_IR14:
+ forall r, preg_of r <> IR14.
Proof.
- intros; red; intro. subst r2. elim H. apply preg_of_is_data_reg.
+ unfold preg_of; intros. destruct (mreg_type r).
+ generalize (ireg_of_not_IR14 r); congruence.
+ congruence.
Qed.
-Hint Resolve preg_of_not: ppcgen.
-Lemma preg_of_not_IR13:
- forall r, preg_of r <> IR13.
+Lemma preg_of_not_PC:
+ forall r, preg_of r <> PC.
Proof.
- intro. case r; discriminate.
+ intros. unfold preg_of. destruct (mreg_type r); congruence.
Qed.
-Hint Resolve preg_of_not_IR13: ppcgen.
-(** Agreement between Mach register sets and PPC register sets. *)
+Lemma ireg_diff:
+ forall r1 r2, r1 <> r2 -> IR r1 <> IR r2.
+Proof. intros; congruence. Qed.
+
+Hint Resolve ireg_of_not_IR13 ireg_of_not_IR14
+ preg_of_not_IR13 preg_of_not_IR14
+ preg_of_not_PC ireg_diff: ppcgen.
+
+(** Agreement between Mach register sets and ARM register sets. *)
-Definition agree (ms: Mach.regset) (sp: val) (rs: Asm.regset) :=
- rs#IR13 = sp /\ forall r: mreg, ms r = rs#(preg_of r).
+Record agree (ms: Mach.regset) (sp: val) (rs: Asm.regset) : Prop := mkagree {
+ agree_sp: rs#IR13 = sp;
+ agree_sp_def: sp <> Vundef;
+ agree_mregs: forall r: mreg, Val.lessdef (ms r) (rs#(preg_of r))
+}.
Lemma preg_val:
forall ms sp rs r,
- agree ms sp rs -> ms r = rs#(preg_of r).
+ agree ms sp rs -> Val.lessdef (ms r) rs#(preg_of r).
+Proof.
+ intros. destruct H. auto.
+Qed.
+
+Lemma preg_vals:
+ forall ms sp rs, agree ms sp rs ->
+ forall l, Val.lessdef_list (map ms l) (map rs (map preg_of l)).
Proof.
- intros. elim H. auto.
+ induction l; simpl. constructor. constructor. eapply preg_val; eauto. auto.
Qed.
-
+
Lemma ireg_val:
forall ms sp rs r,
agree ms sp rs ->
mreg_type r = Tint ->
- ms r = rs#(ireg_of r).
+ Val.lessdef (ms r) rs#(ireg_of r).
Proof.
- intros. elim H; intros.
- generalize (H2 r). unfold preg_of. rewrite H0. auto.
+ intros. generalize (preg_val _ _ _ r H). unfold preg_of. rewrite H0. auto.
Qed.
Lemma freg_val:
forall ms sp rs r,
agree ms sp rs ->
mreg_type r = Tfloat ->
- ms r = rs#(freg_of r).
+ Val.lessdef (ms r) rs#(freg_of r).
Proof.
- intros. elim H; intros.
- generalize (H2 r). unfold preg_of. rewrite H0. auto.
+ intros. generalize (preg_val _ _ _ r H). unfold preg_of. rewrite H0. auto.
Qed.
Lemma sp_val:
@@ -133,76 +128,71 @@ Lemma sp_val:
agree ms sp rs ->
sp = rs#IR13.
Proof.
- intros. elim H; auto.
+ intros. destruct H; auto.
Qed.
-Lemma agree_exten_1:
- forall ms sp rs rs',
- agree ms sp rs ->
- (forall r, is_data_reg r -> rs'#r = rs#r) ->
- agree ms sp rs'.
+Hint Resolve preg_val ireg_val freg_val sp_val: ppcgen.
+
+Definition important_preg (r: preg) : bool :=
+ match r with
+ | IR IR14 => false
+ | IR _ => true
+ | FR _ => true
+ | CR _ => false
+ | PC => false
+ end.
+
+Lemma preg_of_important:
+ forall r, important_preg (preg_of r) = true.
Proof.
- unfold agree; intros. elim H; intros.
- split. rewrite H0. auto. exact I.
- intros. rewrite H0. auto. apply preg_of_is_data_reg.
+ intros. destruct r; reflexivity.
Qed.
-Lemma agree_exten_2:
+Lemma important_diff:
+ forall r r',
+ important_preg r = true -> important_preg r' = false -> r <> r'.
+Proof.
+ congruence.
+Qed.
+Hint Resolve important_diff: ppcgen.
+
+Lemma agree_exten:
forall ms sp rs rs',
agree ms sp rs ->
- (forall r, r <> PC -> r <> IR14 -> rs'#r = rs#r) ->
+ (forall r, important_preg r = true -> rs'#r = rs#r) ->
agree ms sp rs'.
Proof.
- intros. eapply agree_exten_1; eauto.
- intros. apply H0; red; intro; subst r; elim H1.
+ intros. destruct H. split.
+ rewrite H0; auto. auto.
+ intros. rewrite H0; auto. apply preg_of_important.
Qed.
(** Preservation of register agreement under various assignments. *)
Lemma agree_set_mreg:
- forall ms sp rs r v,
+ forall ms sp rs r v rs',
agree ms sp rs ->
- agree (Regmap.set r v ms) sp (rs#(preg_of r) <- v).
-Proof.
- unfold agree; intros. elim H; intros; clear H.
- split. rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_IR13.
- intros. unfold Regmap.set. case (RegEq.eq r0 r); intro.
- subst r0. rewrite Pregmap.gss. auto.
- rewrite Pregmap.gso. auto. red; intro.
- elim n. apply preg_of_injective; auto.
-Qed.
-Hint Resolve agree_set_mreg: ppcgen.
-
-Lemma agree_set_mireg:
- forall ms sp rs r v,
- agree ms sp (rs#(preg_of r) <- v) ->
- mreg_type r = Tint ->
- agree ms sp (rs#(ireg_of r) <- v).
-Proof.
- intros. unfold preg_of in H. rewrite H0 in H. auto.
-Qed.
-Hint Resolve agree_set_mireg: ppcgen.
-
-Lemma agree_set_mfreg:
- forall ms sp rs r v,
- agree ms sp (rs#(preg_of r) <- v) ->
- mreg_type r = Tfloat ->
- agree ms sp (rs#(freg_of r) <- v).
+ Val.lessdef v (rs'#(preg_of r)) ->
+ (forall r', important_preg r' = true -> r' <> preg_of r -> rs'#r' = rs#r') ->
+ agree (Regmap.set r v ms) sp rs'.
Proof.
- intros. unfold preg_of in H. rewrite H0 in H. auto.
+ intros. destruct H. split.
+ rewrite H1; auto. apply sym_not_equal. apply preg_of_not_IR13.
+ auto.
+ intros. unfold Regmap.set. destruct (RegEq.eq r0 r). congruence.
+ rewrite H1. auto. apply preg_of_important.
+ red; intros; elim n. eapply preg_of_injective; eauto.
Qed.
-Hint Resolve agree_set_mfreg: ppcgen.
Lemma agree_set_other:
forall ms sp rs r v,
agree ms sp rs ->
- ~(is_data_reg r) ->
+ important_preg r = false ->
agree ms sp (rs#r <- v).
Proof.
- intros. apply agree_exten_1 with rs.
- auto. intros. apply Pregmap.gso. red; intro; subst r0; contradiction.
+ intros. apply agree_exten with rs. auto.
+ intros. apply Pregmap.gso. congruence.
Qed.
-Hint Resolve agree_set_other: ppcgen.
Lemma agree_nextinstr:
forall ms sp rs,
@@ -210,139 +200,159 @@ Lemma agree_nextinstr:
Proof.
intros. unfold nextinstr. apply agree_set_other. auto. auto.
Qed.
-Hint Resolve agree_nextinstr: ppcgen.
-Lemma agree_set_mireg_twice:
- forall ms sp rs r v v',
- agree ms sp rs ->
- mreg_type r = Tint ->
- agree (Regmap.set r v ms) sp (rs #(ireg_of r) <- v' #(ireg_of r) <- v).
+Definition nontemp_preg (r: preg) : bool :=
+ match r with
+ | IR IR14 => false
+ | IR IR10 => false
+ | IR IR12 => false
+ | IR _ => true
+ | FR FR2 => false
+ | FR FR3 => false
+ | FR _ => true
+ | CR _ => false
+ | PC => false
+ end.
+
+Lemma nontemp_diff:
+ forall r r',
+ nontemp_preg r = true -> nontemp_preg r' = false -> r <> r'.
Proof.
- intros. replace (IR (ireg_of r)) with (preg_of r). elim H; intros.
- split. repeat (rewrite Pregmap.gso; auto with ppcgen).
- intros. case (mreg_eq r r0); intro.
- subst r0. rewrite Regmap.gss. rewrite Pregmap.gss. auto.
- assert (preg_of r <> preg_of r0).
- red; intro. elim n. apply preg_of_injective. auto.
- rewrite Regmap.gso; auto.
- repeat (rewrite Pregmap.gso; auto).
- unfold preg_of. rewrite H0. auto.
+ congruence.
Qed.
-Hint Resolve agree_set_mireg_twice: ppcgen.
-Lemma agree_set_twice_mireg:
- forall ms sp rs r v v',
- agree (Regmap.set r v' ms) sp rs ->
- mreg_type r = Tint ->
- agree (Regmap.set r v ms) sp (rs#(ireg_of r) <- v).
+Hint Resolve nontemp_diff: ppcgen.
+
+Lemma nontemp_important:
+ forall r, nontemp_preg r = true -> important_preg r = true.
Proof.
- intros. elim H; intros.
- split. rewrite Pregmap.gso. auto.
- generalize (ireg_of_not_IR13 r); congruence.
- intros. generalize (H2 r0).
- case (mreg_eq r0 r); intro.
- subst r0. repeat rewrite Regmap.gss. unfold preg_of; rewrite H0.
- rewrite Pregmap.gss. auto.
- repeat rewrite Regmap.gso; auto.
- rewrite Pregmap.gso. auto.
- replace (IR (ireg_of r)) with (preg_of r).
- red; intros. elim n. apply preg_of_injective; auto.
- unfold preg_of. rewrite H0. auto.
+ unfold nontemp_preg, important_preg; destruct r; auto. destruct i; auto.
Qed.
-Hint Resolve agree_set_twice_mireg: ppcgen.
-Lemma agree_set_commut:
- forall ms sp rs r1 r2 v1 v2,
- r1 <> r2 ->
- agree ms sp ((rs#r2 <- v2)#r1 <- v1) ->
- agree ms sp ((rs#r1 <- v1)#r2 <- v2).
+Hint Resolve nontemp_important: ppcgen.
+
+Remark undef_regs_1:
+ forall l ms r, ms r = Vundef -> Mach.undef_regs l ms r = Vundef.
Proof.
- intros. apply agree_exten_1 with ((rs#r2 <- v2)#r1 <- v1). auto.
- intros.
- case (preg_eq r r1); intro.
- subst r1. rewrite Pregmap.gss. rewrite Pregmap.gso. rewrite Pregmap.gss.
- auto. auto.
- case (preg_eq r r2); intro.
- subst r2. rewrite Pregmap.gss. rewrite Pregmap.gso. rewrite Pregmap.gss.
- auto. auto.
- repeat (rewrite Pregmap.gso; auto).
-Qed.
-Hint Resolve agree_set_commut: ppcgen.
-
-Lemma agree_nextinstr_commut:
- forall ms sp rs r v,
- agree ms sp (rs#r <- v) ->
- r <> PC ->
- agree ms sp ((nextinstr rs)#r <- v).
+ induction l; simpl; intros. auto. apply IHl. unfold Regmap.set.
+ destruct (RegEq.eq r a); congruence.
+Qed.
+
+Remark undef_regs_2:
+ forall l ms r, In r l -> Mach.undef_regs l ms r = Vundef.
+Proof.
+ induction l; simpl; intros. contradiction.
+ destruct H. subst. apply undef_regs_1. apply Regmap.gss.
+ auto.
+Qed.
+
+Remark undef_regs_3:
+ forall l ms r, ~In r l -> Mach.undef_regs l ms r = ms r.
Proof.
- intros. unfold nextinstr. apply agree_set_commut. auto.
- apply agree_set_other. auto. auto.
+ induction l; simpl; intros. auto.
+ rewrite IHl. apply Regmap.gso. intuition. intuition.
Qed.
-Hint Resolve agree_nextinstr_commut: ppcgen.
-Lemma agree_set_mireg_exten:
- forall ms sp rs r v (rs': regset),
+Lemma agree_exten_temps:
+ forall ms sp rs rs',
agree ms sp rs ->
- mreg_type r = Tint ->
- rs'#(ireg_of r) = v ->
- (forall r', r' <> PC -> r' <> ireg_of r -> r' <> IR14 -> rs'#r' = rs#r') ->
- agree (Regmap.set r v ms) sp rs'.
+ (forall r, nontemp_preg r = true -> rs'#r = rs#r) ->
+ agree (undef_temps ms) sp rs'.
Proof.
- intros. apply agree_exten_2 with (rs#(ireg_of r) <- v).
- auto with ppcgen.
- intros. unfold Pregmap.set. case (PregEq.eq r0 (ireg_of r)); intro.
- subst r0. auto. apply H2; auto.
+ intros. destruct H. split.
+ rewrite H0; auto. auto.
+ intros. unfold undef_temps.
+ destruct (In_dec mreg_eq r (int_temporaries ++ float_temporaries)).
+ rewrite undef_regs_2; auto.
+ rewrite undef_regs_3; auto. rewrite H0; auto.
+ simpl in n. destruct r; auto; intuition.
Qed.
-(** Useful properties of the PC and GPR0 registers. *)
+Lemma agree_set_undef_mreg:
+ forall ms sp rs r v rs',
+ agree ms sp rs ->
+ Val.lessdef v (rs'#(preg_of r)) ->
+ (forall r', nontemp_preg r' = true -> r' <> preg_of r -> rs'#r' = rs#r') ->
+ agree (Regmap.set r v (undef_temps ms)) sp rs'.
+Proof.
+ intros. apply agree_set_mreg with (rs'#(preg_of r) <- (rs#(preg_of r))); auto.
+ eapply agree_exten_temps; eauto.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r0 (preg_of r)).
+ congruence. auto.
+ intros. rewrite Pregmap.gso; auto.
+Qed.
+
+Lemma agree_undef_temps:
+ forall ms sp rs,
+ agree ms sp rs ->
+ agree (undef_temps ms) sp rs.
+Proof.
+ intros. eapply agree_exten_temps; eauto.
+Qed.
+
+(** Useful properties of the PC register. *)
Lemma nextinstr_inv:
- forall r rs, r <> PC -> (nextinstr rs)#r = rs#r.
+ forall r rs,
+ r <> PC ->
+ (nextinstr rs)#r = rs#r.
+Proof.
+ intros. unfold nextinstr. apply Pregmap.gso. red; intro; subst. auto.
+Qed.
+
+Lemma nextinstr_inv2:
+ forall r rs,
+ nontemp_preg r = true ->
+ (nextinstr rs)#r = rs#r.
Proof.
- intros. unfold nextinstr. apply Pregmap.gso. auto.
+ intros. apply nextinstr_inv. red; intro; subst; discriminate.
Qed.
-Hint Resolve nextinstr_inv: ppcgen.
Lemma nextinstr_set_preg:
forall rs m v,
(nextinstr (rs#(preg_of m) <- v))#PC = Val.add rs#PC Vone.
Proof.
intros. unfold nextinstr. rewrite Pregmap.gss.
- rewrite Pregmap.gso. auto. apply sym_not_eq. auto with ppcgen.
+ rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_PC.
Qed.
-Hint Resolve nextinstr_set_preg: ppcgen.
(** Connection between Mach and Asm calling conventions for external
functions. *)
Lemma extcall_arg_match:
- forall ms sp rs m l v,
+ forall ms sp rs m m' l v,
agree ms sp rs ->
Machconcr.extcall_arg ms m sp l v ->
- Asm.extcall_arg rs m l v.
+ Mem.extends m m' ->
+ exists v', Asm.extcall_arg rs m' l v' /\ Val.lessdef v v'.
Proof.
- intros. inv H0.
- rewrite (preg_val _ _ _ r H). constructor.
- rewrite (sp_val _ _ _ H) in H1.
- destruct ty; unfold load_stack in H1.
- econstructor. reflexivity. assumption.
- econstructor. reflexivity. assumption.
+ intros. inv H0.
+ exists (rs#(preg_of r)); split. constructor. eauto with ppcgen.
+ unfold load_stack in H2.
+ exploit Mem.loadv_extends; eauto. intros [v' [A B]].
+ rewrite (sp_val _ _ _ H) in A.
+ exists v'; split; auto. destruct ty; econstructor; eauto.
Qed.
Lemma extcall_args_match:
- forall ms sp rs m, agree ms sp rs ->
+ forall ms sp rs m m', agree ms sp rs -> Mem.extends m m' ->
forall ll vl,
Machconcr.extcall_args ms m sp ll vl ->
- Asm.extcall_args rs m ll vl.
+ exists vl', Asm.extcall_args rs m' ll vl' /\ Val.lessdef_list vl vl'.
Proof.
- induction 2; constructor; auto. eapply extcall_arg_match; eauto.
+ induction 3.
+ exists (@nil val); split; constructor.
+ exploit extcall_arg_match; eauto. intros [v1' [A B]].
+ exploit IHextcall_args; eauto. intros [vl' [C D]].
+ exists(v1' :: vl'); split. constructor; auto. constructor; auto.
Qed.
Lemma extcall_arguments_match:
- forall ms m sp rs sg args,
+ forall ms m sp rs sg args m',
agree ms sp rs ->
Machconcr.extcall_arguments ms m sp sg args ->
- Asm.extcall_arguments rs m sg args.
+ Mem.extends m m' ->
+ exists args', Asm.extcall_arguments rs m' sg args' /\ Val.lessdef_list args args'.
Proof.
unfold Machconcr.extcall_arguments, Asm.extcall_arguments; intros.
eapply extcall_args_match; eauto.
@@ -564,7 +574,7 @@ Proof.
exploit loadimm_correct. intros [rs' [A [B C]]].
exists (nextinstr (rs'#r1 <- (Val.and rs#r2 (Vint n)))).
split. eapply exec_straight_trans. eauto. apply exec_straight_one.
- simpl. rewrite B. rewrite C; auto with ppcgen. congruence.
+ simpl. rewrite B. rewrite C; auto with ppcgen.
auto.
split. rewrite nextinstr_inv; auto with ppcgen. apply Pregmap.gss.
intros. rewrite nextinstr_inv; auto. rewrite Pregmap.gso; auto.
@@ -756,20 +766,6 @@ Qed.
(** Translation of conditions *)
-Ltac TypeInv :=
- match goal with
- | H: (List.map ?f ?x = nil) |- _ =>
- destruct x; [clear H | simpl in H; discriminate]
- | H: (List.map ?f ?x = ?hd :: ?tl) |- _ =>
- destruct x; simpl in H;
- [ discriminate |
- injection H; clear H; let T := fresh "T" in (
- intros H T; TypeInv) ]
- | _ => idtac
- end.
-
-(** Translation of conditions. *)
-
Lemma compare_int_spec:
forall rs v1 v2,
let rs1 := nextinstr (compare_int rs v1 v2) in
@@ -783,11 +779,11 @@ Lemma compare_int_spec:
/\ rs1#CRlt = (Val.cmp Clt v1 v2)
/\ rs1#CRgt = (Val.cmp Cgt v1 v2)
/\ rs1#CRle = (Val.cmp Cle v1 v2)
- /\ forall r', is_data_reg r' -> rs1#r' = rs#r'.
+ /\ forall r', important_preg r' = true -> rs1#r' = rs#r'.
Proof.
intros. unfold rs1. intuition; try reflexivity.
- rewrite nextinstr_inv; [unfold compare_int; repeat rewrite Pregmap.gso; auto | idtac];
- red; intro; subst r'; elim H.
+ rewrite nextinstr_inv; auto with ppcgen.
+ unfold compare_int. repeat rewrite Pregmap.gso; auto with ppcgen.
Qed.
Lemma compare_float_spec:
@@ -803,302 +799,237 @@ Lemma compare_float_spec:
/\ rs'#CRlt = (Val.notbool (Val.cmpf Cge v1 v2))
/\ rs'#CRgt = (Val.cmpf Cgt v1 v2)
/\ rs'#CRle = (Val.notbool (Val.cmpf Cgt v1 v2))
- /\ forall r', is_data_reg r' -> rs'#r' = rs#r'.
+ /\ forall r', important_preg r' = true -> rs'#r' = rs#r'.
Proof.
intros. unfold rs'. intuition; try reflexivity.
- rewrite nextinstr_inv; [unfold compare_float; repeat rewrite Pregmap.gso; auto | idtac];
- red; intro; subst r'; elim H.
+ rewrite nextinstr_inv; auto with ppcgen.
+ unfold compare_float. repeat rewrite Pregmap.gso; auto with ppcgen.
Qed.
+Ltac TypeInv1 :=
+ match goal with
+ | H: (List.map ?f ?x = nil) |- _ =>
+ destruct x; inv H; TypeInv1
+ | H: (List.map ?f ?x = ?hd :: ?tl) |- _ =>
+ destruct x; simpl in H; simplify_eq H; clear H; intros; TypeInv1
+ | _ => idtac
+ end.
+
+Ltac TypeInv2 :=
+ match goal with
+ | H: (mreg_type _ = Tint) |- _ => try (rewrite H in *); clear H; TypeInv2
+ | H: (mreg_type _ = Tfloat) |- _ => try (rewrite H in *); clear H; TypeInv2
+ | _ => idtac
+ end.
+
+Ltac TypeInv := TypeInv1; simpl in *; unfold preg_of in *; TypeInv2.
+
Lemma transl_cond_correct:
- forall cond args k ms sp rs m b,
+ forall cond args k rs m b,
map mreg_type args = type_of_condition cond ->
- agree ms sp rs ->
- eval_condition cond (map ms args) = Some b ->
+ eval_condition cond (map rs (map preg_of args)) = Some b ->
exists rs',
exec_straight (transl_cond cond args k) rs m k rs' m
/\ rs'#(CR (crbit_for_cond cond)) = Val.of_bool b
- /\ agree ms sp rs'.
+ /\ forall r, important_preg r = true -> rs'#r = rs r.
Proof.
- intros.
- rewrite <- (eval_condition_weaken _ _ H1). clear H1.
- destruct cond; simpl in H; TypeInv; simpl.
+ intros until b; intros TY EV. rewrite <- (eval_condition_weaken _ _ EV). clear EV.
+ destruct cond; simpl in TY; TypeInv.
(* Ccomp *)
- generalize (compare_int_spec rs ms#m0 ms#m1).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (rs (ireg_of m1))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 ms#m1)).
- split. apply exec_straight_one. simpl.
- repeat rewrite <- (ireg_val ms sp rs); trivial.
- reflexivity.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; assumption.
+ auto.
(* Ccompu *)
- generalize (compare_int_spec rs ms#m0 ms#m1).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (rs (ireg_of m1))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 ms#m1)).
- split. apply exec_straight_one. simpl.
- repeat rewrite <- (ireg_val ms sp rs); trivial.
- reflexivity.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; assumption.
+ auto.
(* Ccompshift *)
- generalize (compare_int_spec rs ms#m0 (eval_shift_total s ms#m1)).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (eval_shift_total s (rs (ireg_of m1)))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 (eval_shift_total s ms#m1))).
- split. apply exec_straight_one. simpl.
- rewrite transl_shift_correct.
- repeat rewrite <- (ireg_val ms sp rs); trivial.
- reflexivity.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. rewrite transl_shift_correct. case c; assumption.
+ rewrite transl_shift_correct. auto.
(* Ccompushift *)
- generalize (compare_int_spec rs ms#m0 (eval_shift_total s ms#m1)).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (eval_shift_total s (rs (ireg_of m1)))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 (eval_shift_total s ms#m1))).
- split. apply exec_straight_one. simpl.
- rewrite transl_shift_correct.
- repeat rewrite <- (ireg_val ms sp rs); trivial.
- reflexivity.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. rewrite transl_shift_correct. case c; assumption.
+ rewrite transl_shift_correct. auto.
(* Ccompimm *)
destruct (is_immed_arith i).
- generalize (compare_int_spec rs ms#m0 (Vint i)).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (Vint i)).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 (Vint i))).
- split. apply exec_straight_one. simpl.
- rewrite <- (ireg_val ms sp rs); trivial. auto.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; assumption.
+ auto.
exploit (loadimm_correct IR14). intros [rs' [P [Q R]]].
- assert (AG: agree ms sp rs'). apply agree_exten_2 with rs; auto.
- generalize (compare_int_spec rs' ms#m0 (Vint i)).
+ generalize (compare_int_spec rs' (rs (ireg_of m0)) (Vint i)).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs' ms#m0 (Vint i))).
+ econstructor.
split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl.
- rewrite Q. rewrite <- (ireg_val ms sp rs'); trivial. auto.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs'; auto.
+ rewrite Q. rewrite R; eauto with ppcgen. auto.
+ split. case c; assumption.
+ intros. rewrite K; auto with ppcgen.
(* Ccompuimm *)
destruct (is_immed_arith i).
- generalize (compare_int_spec rs ms#m0 (Vint i)).
+ generalize (compare_int_spec rs (rs (ireg_of m0)) (Vint i)).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs ms#m0 (Vint i))).
- split. apply exec_straight_one. simpl.
- rewrite <- (ireg_val ms sp rs); trivial. auto.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; assumption.
+ auto.
exploit (loadimm_correct IR14). intros [rs' [P [Q R]]].
- assert (AG: agree ms sp rs'). apply agree_exten_2 with rs; auto.
- generalize (compare_int_spec rs' ms#m0 (Vint i)).
+ generalize (compare_int_spec rs' (rs (ireg_of m0)) (Vint i)).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_int rs' ms#m0 (Vint i))).
+ econstructor.
split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl.
- rewrite Q. rewrite <- (ireg_val ms sp rs'); trivial. auto.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs'; auto.
+ rewrite Q. rewrite R; eauto with ppcgen. auto.
+ split. case c; assumption.
+ intros. rewrite K; auto with ppcgen.
(* Ccompf *)
- generalize (compare_float_spec rs ms#m0 ms#m1).
+ generalize (compare_float_spec rs (rs (freg_of m0)) (rs (freg_of m1))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_float rs ms#m0 ms#m1)).
- split. apply exec_straight_one. simpl.
- repeat rewrite <- (freg_val ms sp rs); trivial. auto.
- split.
- case c; simpl; auto.
- apply agree_exten_1 with rs; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; assumption.
+ auto.
(* Cnotcompf *)
- generalize (compare_float_spec rs ms#m0 ms#m1).
+ generalize (compare_float_spec rs (rs (freg_of m0)) (rs (freg_of m1))).
intros [A [B [C [D [E [F [G [H [I [J K]]]]]]]]]].
- exists (nextinstr (compare_float rs ms#m0 ms#m1)).
- split. apply exec_straight_one. simpl.
- repeat rewrite <- (freg_val ms sp rs); trivial. auto.
- split.
- case c; simpl; auto.
+ econstructor.
+ split. apply exec_straight_one. simpl. eauto. auto.
+ split. case c; try assumption.
rewrite Val.negate_cmpf_ne. auto.
rewrite Val.negate_cmpf_eq. auto.
- apply agree_exten_1 with rs; auto.
+ auto.
Qed.
(** Translation of arithmetic operations. *)
-Ltac TranslOpSimpl :=
+Ltac Simpl :=
match goal with
- | |- exists rs' : regset,
- exec_straight ?c ?rs ?m ?k rs' ?m /\
- agree (Regmap.set ?res ?v ?ms) ?sp rs' =>
- (exists (nextinstr (rs#(ireg_of res) <- v));
- split;
- [ apply exec_straight_one;
- [ repeat (rewrite (ireg_val ms sp rs); auto);
- simpl; try rewrite transl_shift_correct; reflexivity
- | reflexivity ]
- | auto with ppcgen ])
- ||
- (exists (nextinstr (rs#(freg_of res) <- v));
- split;
- [ apply exec_straight_one;
- [ repeat (rewrite (freg_val ms sp rs); auto); reflexivity
- | reflexivity ]
- | auto with ppcgen ])
+ | [ |- nextinstr _ _ = _ ] => rewrite nextinstr_inv; [auto | auto with ppcgen]
+ | [ |- Pregmap.get ?x (Pregmap.set ?x _ _) = _ ] => rewrite Pregmap.gss; auto
+ | [ |- Pregmap.set ?x _ _ ?x = _ ] => rewrite Pregmap.gss; auto
+ | [ |- Pregmap.get _ (Pregmap.set _ _ _) = _ ] => rewrite Pregmap.gso; [auto | auto with ppcgen]
+ | [ |- Pregmap.set _ _ _ _ = _ ] => rewrite Pregmap.gso; [auto | auto with ppcgen]
end.
+Ltac TranslOpSimpl :=
+ econstructor; split;
+ [ apply exec_straight_one; [simpl; eauto | reflexivity ]
+ | split; [try rewrite transl_shift_correct; repeat Simpl | intros; repeat Simpl] ].
+
Lemma transl_op_correct:
- forall op args res k ms sp rs m v,
+ forall op args res k (rs: regset) m v,
wt_instr (Mop op args res) ->
- agree ms sp rs ->
- eval_operation ge sp op (map ms args) = Some v ->
+ eval_operation ge rs#IR13 op (map rs (map preg_of args)) = Some v ->
exists rs',
exec_straight (transl_op op args res k) rs m k rs' m
- /\ agree (Regmap.set res v ms) sp rs'.
+ /\ rs'#(preg_of res) = v
+ /\ forall r, important_preg r = true -> r <> preg_of res -> rs'#r = rs#r.
Proof.
- intros. rewrite <- (eval_operation_weaken _ _ _ _ H1). (*clear H1; clear v.*)
- inversion H.
+ intros. rewrite <- (eval_operation_weaken _ _ _ _ H0). inv H.
(* Omove *)
- simpl. exists (nextinstr (rs#(preg_of res) <- (ms r1))).
- split. caseEq (mreg_type r1); intro.
- apply exec_straight_one. simpl. rewrite (ireg_val ms sp rs); auto.
- simpl. unfold preg_of. rewrite <- H3. rewrite H6. reflexivity.
- auto with ppcgen.
- apply exec_straight_one. simpl. rewrite (freg_val ms sp rs); auto.
- simpl. unfold preg_of. rewrite <- H3. rewrite H6. reflexivity.
- auto with ppcgen.
- auto with ppcgen.
+ simpl.
+ exists (nextinstr (rs#(preg_of res) <- (rs#(preg_of r1)))).
+ split. unfold preg_of; rewrite <- H2.
+ destruct (mreg_type r1); apply exec_straight_one; auto.
+ split. Simpl. Simpl.
+ intros. Simpl. Simpl.
(* Other instructions *)
- clear H2 H3 H5.
- destruct op; simpl in H6; injection H6; clear H6; intros;
- TypeInv; simpl; try (TranslOpSimpl).
+ destruct op; simpl in H5; inv H5; TypeInv; try (TranslOpSimpl; fail).
(* Omove again *)
congruence.
(* Ointconst *)
- generalize (loadimm_correct (ireg_of res) i k rs m).
- intros [rs' [A [B C]]].
- exists rs'. split. auto.
- apply agree_set_mireg_exten with rs; auto.
+ generalize (loadimm_correct (ireg_of res) i k rs m). intros [rs' [A [B C]]].
+ exists rs'. split. auto. split. auto. intros. auto with ppcgen.
(* Oaddrstack *)
generalize (addimm_correct (ireg_of res) IR13 i k rs m).
intros [rs' [EX [RES OTH]]].
- exists rs'. split. auto.
- apply agree_set_mireg_exten with rs; auto.
- rewrite (sp_val ms sp rs). auto. auto.
+ exists rs'. split. auto. split. auto. auto with ppcgen.
(* Ocast8signed *)
- set (rs1 := nextinstr (rs#(ireg_of res) <- (Val.shl (ms m0) (Vint (Int.repr 24))))).
- set (rs2 := nextinstr (rs1#(ireg_of res) <- (Val.shr (rs1 (ireg_of res)) (Vint (Int.repr 24))))).
- exists rs2. split.
- apply exec_straight_two with rs1 m; auto.
- simpl. rewrite <- (ireg_val ms sp rs); auto.
- unfold rs2.
- replace (Val.shr (rs1 (ireg_of res)) (Vint (Int.repr 24))) with (Val.sign_ext 8 (ms m0)).
- apply agree_nextinstr. unfold rs1. apply agree_nextinstr_commut.
- apply agree_set_mireg_twice; auto with ppcgen. auto with ppcgen.
- unfold rs1. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
- destruct (ms m0); simpl; auto. rewrite Int.sign_ext_shr_shl. reflexivity.
- vm_compute; auto.
+ econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl. Simpl. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
+ destruct (rs (ireg_of m0)); simpl; auto. rewrite Int.sign_ext_shr_shl. reflexivity.
+ compute; auto.
+ intros. repeat Simpl.
(* Ocast8unsigned *)
- exists (nextinstr (rs#(ireg_of res) <- (Val.and (ms m0) (Vint (Int.repr 255))))).
- split. apply exec_straight_one. repeat (rewrite (ireg_val ms sp rs)); auto. reflexivity.
- replace (Val.zero_ext 8 (ms m0))
- with (Val.and (ms m0) (Vint (Int.repr 255))).
- auto with ppcgen.
- destruct (ms m0); simpl; auto. rewrite Int.zero_ext_and. reflexivity.
- vm_compute; auto.
+ econstructor; split.
+ eapply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. Simpl.
+ destruct (rs (ireg_of m0)); simpl; auto. rewrite Int.zero_ext_and. reflexivity.
+ compute; auto.
+ intros. repeat Simpl.
(* Ocast16signed *)
- set (rs1 := nextinstr (rs#(ireg_of res) <- (Val.shl (ms m0) (Vint (Int.repr 16))))).
- set (rs2 := nextinstr (rs1#(ireg_of res) <- (Val.shr (rs1 (ireg_of res)) (Vint (Int.repr 16))))).
- exists rs2. split.
- apply exec_straight_two with rs1 m; auto.
- simpl. rewrite <- (ireg_val ms sp rs); auto.
- unfold rs2.
- replace (Val.shr (rs1 (ireg_of res)) (Vint (Int.repr 16))) with (Val.sign_ext 16 (ms m0)).
- apply agree_nextinstr. unfold rs1. apply agree_nextinstr_commut.
- apply agree_set_mireg_twice; auto with ppcgen. auto with ppcgen.
- unfold rs1. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
- destruct (ms m0); simpl; auto. rewrite Int.sign_ext_shr_shl. reflexivity.
- vm_compute; auto.
+ econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl. Simpl. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
+ destruct (rs (ireg_of m0)); simpl; auto. rewrite Int.sign_ext_shr_shl. reflexivity.
+ compute; auto.
+ intros. repeat Simpl.
(* Ocast16unsigned *)
- set (rs1 := nextinstr (rs#(ireg_of res) <- (Val.shl (ms m0) (Vint (Int.repr 16))))).
- set (rs2 := nextinstr (rs1#(ireg_of res) <- (Val.shru (rs1 (ireg_of res)) (Vint (Int.repr 16))))).
- exists rs2. split.
- apply exec_straight_two with rs1 m; auto.
- simpl. rewrite <- (ireg_val ms sp rs); auto.
- unfold rs2.
- replace (Val.shru (rs1 (ireg_of res)) (Vint (Int.repr 16))) with (Val.zero_ext 16 (ms m0)).
- apply agree_nextinstr. unfold rs1. apply agree_nextinstr_commut.
- apply agree_set_mireg_twice; auto with ppcgen. auto with ppcgen.
- unfold rs1. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
- destruct (ms m0); simpl; auto. rewrite Int.zero_ext_shru_shl. reflexivity.
- vm_compute; auto.
+ econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl. Simpl. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
+ destruct (rs (ireg_of m0)); simpl; auto. rewrite Int.zero_ext_shru_shl. reflexivity.
+ compute; auto.
+ intros. repeat Simpl.
(* Oaddimm *)
generalize (addimm_correct (ireg_of res) (ireg_of m0) i k rs m).
intros [rs' [A [B C]]].
- exists rs'. split. auto.
- apply agree_set_mireg_exten with rs; auto.
- rewrite (ireg_val ms sp rs); auto.
+ exists rs'. split. auto. split. auto. auto with ppcgen.
(* Orsbimm *)
exploit (makeimm_correct Prsb (fun v1 v2 => Val.sub v2 v1) (ireg_of res) (ireg_of m0));
auto with ppcgen.
intros [rs' [A [B C]]].
exists rs'.
- split. eauto.
- apply agree_set_mireg_exten with rs; auto. rewrite B.
- rewrite <- (ireg_val ms sp rs); auto.
+ split. eauto. split. rewrite B. auto. auto with ppcgen.
(* Omul *)
destruct (ireg_eq (ireg_of res) (ireg_of m0) || ireg_eq (ireg_of res) (ireg_of m1)).
- set (rs1 := nextinstr (rs#IR14 <- (Val.mul (ms m0) (ms m1)))).
- set (rs2 := nextinstr (rs1#(ireg_of res) <- (rs1#IR14))).
- exists rs2; split.
- apply exec_straight_two with rs1 m; auto.
- simpl. repeat rewrite <- (ireg_val ms sp rs); auto.
- unfold rs2. unfold rs1. rewrite nextinstr_inv. rewrite Pregmap.gss.
- apply agree_nextinstr. apply agree_nextinstr_commut.
- apply agree_set_mireg; auto. apply agree_set_mreg. apply agree_set_other. auto.
- simpl; auto. auto with ppcgen. discriminate.
+ econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. repeat Simpl.
+ intros. repeat Simpl.
TranslOpSimpl.
(* Oandimm *)
generalize (andimm_correct (ireg_of res) (ireg_of m0) i k rs m
(ireg_of_not_IR14 m0)).
intros [rs' [A [B C]]].
- exists rs'. split. auto.
- apply agree_set_mireg_exten with rs; auto.
- rewrite (ireg_val ms sp rs); auto.
+ exists rs'. split. auto. split. auto. auto with ppcgen.
(* Oorimm *)
exploit (makeimm_correct Porr Val.or (ireg_of res) (ireg_of m0));
auto with ppcgen.
- intros [rs' [A [B C]]].
- exists rs'.
- split. eauto.
- apply agree_set_mireg_exten with rs; auto. rewrite B.
- rewrite <- (ireg_val ms sp rs); auto.
+ intros [rs' [A [B C]]].
+ exists rs'. split. eauto. split. auto. auto with ppcgen.
(* Oxorimm *)
exploit (makeimm_correct Peor Val.xor (ireg_of res) (ireg_of m0));
auto with ppcgen.
- intros [rs' [A [B C]]].
- exists rs'.
- split. eauto.
- apply agree_set_mireg_exten with rs; auto. rewrite B.
- rewrite <- (ireg_val ms sp rs); auto.
+ intros [rs' [A [B C]]].
+ exists rs'. split. eauto. split. auto. auto with ppcgen.
(* Oshrximm *)
- assert (exists n, ms m0 = Vint n /\ Int.ltu i (Int.repr 31) = true).
- simpl in H1. destruct (ms m0); try discriminate.
+ assert (exists n, rs (ireg_of m0) = Vint n /\ Int.ltu i (Int.repr 31) = true).
+ destruct (rs (ireg_of m0)); try discriminate.
exists i0; split; auto. destruct (Int.ltu i (Int.repr 31)); discriminate || auto.
- destruct H3 as [n [ARG1 LTU]].
+ destruct H as [n [ARG1 LTU]]. clear H0.
assert (LTU': Int.ltu i Int.iwordsize = true).
exploit Int.ltu_inv. eexact LTU. intro.
unfold Int.ltu. apply zlt_true.
- assert (Int.unsigned (Int.repr 31) < Int.unsigned Int.iwordsize). vm_compute; auto.
+ assert (Int.unsigned (Int.repr 31) < Int.unsigned Int.iwordsize). compute; auto.
omega.
- assert (RSm0: rs (ireg_of m0) = Vint n).
- rewrite <- ARG1. symmetry. eapply ireg_val; eauto.
set (islt := Int.lt n Int.zero).
set (rs1 := nextinstr (compare_int rs (Vint n) (Vint Int.zero))).
- assert (OTH1: forall r', is_data_reg r' -> rs1#r' = rs#r').
+ assert (OTH1: forall r', important_preg r' = true -> rs1#r' = rs#r').
generalize (compare_int_spec rs (Vint n) (Vint Int.zero)).
fold rs1. intros [A B]. intuition.
exploit (addimm_correct IR14 (ireg_of m0) (Int.sub (Int.shl Int.one i) Int.one)).
@@ -1107,83 +1038,47 @@ Proof.
set (rs4 := nextinstr (rs3#(ireg_of res) <- (Val.shr rs3#IR14 (Vint i)))).
exists rs4; split.
apply exec_straight_step with rs1 m.
- simpl. rewrite RSm0. auto. auto.
+ simpl. rewrite ARG1. auto. auto.
eapply exec_straight_trans. eexact EXEC2.
apply exec_straight_two with rs3 m.
simpl. rewrite OTH2. change (rs1 CRge) with (Val.cmp Cge (Vint n) (Vint Int.zero)).
unfold Val.cmp. change (Int.cmp Cge n Int.zero) with (negb islt).
- rewrite OTH2. rewrite OTH1. rewrite RSm0.
+ rewrite OTH2. rewrite OTH1. rewrite ARG1.
unfold rs3. case islt; reflexivity.
- apply ireg_of_is_data_reg. decEq; auto with ppcgen. auto with ppcgen. congruence. congruence.
+ destruct m0; reflexivity. auto with ppcgen. auto with ppcgen. discriminate. discriminate.
simpl. auto.
auto. unfold rs3. case islt; auto. auto.
- (* agreement *)
- assert (RES4: rs4#(ireg_of res) = Vint(Int.shrx n i)).
- unfold rs4. rewrite nextinstr_inv; auto. rewrite Pregmap.gss.
- rewrite Int.shrx_shr. fold islt. unfold rs3.
- repeat rewrite nextinstr_inv; auto.
- case islt. rewrite RES2. rewrite OTH1. rewrite RSm0.
- simpl. rewrite LTU'. auto.
- apply ireg_of_is_data_reg.
- rewrite Pregmap.gss. simpl. rewrite LTU'. auto. congruence.
- exact LTU. auto with ppcgen.
- assert (OTH4: forall r, is_data_reg r -> r <> ireg_of res -> rs4#r = rs#r).
- intros.
- assert (r <> PC). red; intro; subst r; elim H3.
- assert (r <> IR14). red; intro; subst r; elim H3.
- unfold rs4. rewrite nextinstr_inv; auto. rewrite Pregmap.gso; auto.
- unfold rs3. rewrite nextinstr_inv; auto.
- transitivity (rs2 r).
- case islt. auto. apply Pregmap.gso; auto.
- rewrite OTH2; auto.
- apply agree_exten_1 with (rs#(ireg_of res) <- (Val.shrx (ms m0) (Vint i))).
- auto with ppcgen.
- intros. unfold Pregmap.set. destruct (PregEq.eq r (ireg_of res)).
- subst r. rewrite ARG1. simpl. rewrite LTU'. auto.
- auto.
- (* Ointoffloat *)
- exists (nextinstr (rs#(ireg_of res) <- (Val.intoffloat (ms m0)))).
- split. apply exec_straight_one.
- repeat (rewrite (freg_val ms sp rs); auto).
- reflexivity. auto with ppcgen.
- (* Ointuoffloat *)
- exists (nextinstr (rs#(ireg_of res) <- (Val.intuoffloat (ms m0)))).
- split. apply exec_straight_one.
- repeat (rewrite (freg_val ms sp rs); auto).
- reflexivity. auto with ppcgen.
- (* Ofloatofint *)
- exists (nextinstr (rs#(freg_of res) <- (Val.floatofint (ms m0)))).
- split. apply exec_straight_one.
- repeat (rewrite (ireg_val ms sp rs); auto).
- reflexivity. auto 10 with ppcgen.
- (* Ofloatofintu *)
- exists (nextinstr (rs#(freg_of res) <- (Val.floatofintu (ms m0)))).
- split. apply exec_straight_one.
- repeat (rewrite (ireg_val ms sp rs); auto).
- reflexivity. auto 10 with ppcgen.
+ split. unfold rs4. repeat Simpl. rewrite ARG1. simpl. rewrite LTU'. rewrite Int.shrx_shr.
+ fold islt. unfold rs3. rewrite nextinstr_inv; auto with ppcgen.
+ destruct islt. rewrite RES2. change (rs1 (IR (ireg_of m0))) with (rs (IR (ireg_of m0))).
+ rewrite ARG1. simpl. rewrite LTU'. auto.
+ rewrite Pregmap.gss. simpl. rewrite LTU'. auto.
+ assumption.
+ intros. unfold rs4; repeat Simpl. unfold rs3; repeat Simpl.
+ transitivity (rs2 r). destruct islt; auto. Simpl.
+ rewrite OTH2; auto with ppcgen.
(* Ocmp *)
- assert (exists b, eval_condition c ms##args = Some b /\ v = Val.of_bool b).
- simpl in H1. destruct (eval_condition c ms##args).
- destruct b; inv H1. exists true; auto. exists false; auto.
- discriminate.
- destruct H5 as [b [EVC EQ]].
- exploit transl_cond_correct; eauto. intros [rs' [A [B C]]].
- rewrite (eval_condition_weaken _ _ EVC).
- set (rs1 := nextinstr (rs'#(ireg_of res) <- (Vint Int.zero))).
- set (rs2 := nextinstr (if b then (rs1#(ireg_of res) <- Vtrue) else rs1)).
- exists rs2; split.
- eapply exec_straight_trans. eauto.
- apply exec_straight_two with rs1 m; auto.
- simpl. replace (rs1 (crbit_for_cond c)) with (Val.of_bool b).
- unfold rs2. destruct b; auto.
- unfold rs2. destruct b; auto.
- apply agree_set_mireg_exten with rs'; auto.
- unfold rs2. rewrite nextinstr_inv; auto with ppcgen.
- destruct b. apply Pregmap.gss.
- unfold rs1. rewrite nextinstr_inv; auto with ppcgen. apply Pregmap.gss.
- intros. unfold rs2. rewrite nextinstr_inv; auto.
- transitivity (rs1 r'). destruct b; auto. rewrite Pregmap.gso; auto.
- unfold rs1. rewrite nextinstr_inv; auto. apply Pregmap.gso; auto.
+ fold preg_of in *.
+ assert (exists b, eval_condition c rs ## (preg_of ## args) = Some b /\ v = Val.of_bool b).
+ fold preg_of in H0. destruct (eval_condition c rs ## (preg_of ## args)).
+ exists b; split; auto. destruct b; inv H0; auto. congruence.
+ clear H0. destruct H as [b [EVC VBO]]. rewrite (eval_condition_weaken _ _ EVC).
+ destruct (transl_cond_correct c args
+ (Pmov (ireg_of res) (SOimm Int.zero)
+ :: Pmovc (crbit_for_cond c) (ireg_of res) (SOimm Int.one) :: k)
+ rs m b H1 EVC)
+ as [rs1 [A [B C]]].
+ set (rs2 := nextinstr (rs1#(ireg_of res) <- (Vint Int.zero))).
+ set (rs3 := nextinstr (if b then (rs2#(ireg_of res) <- Vtrue) else rs2)).
+ exists rs3.
+ split. eapply exec_straight_trans. eauto.
+ apply exec_straight_two with rs2 m; auto.
+ simpl. replace (rs2 (crbit_for_cond c)) with (Val.of_bool b).
+ unfold rs3. destruct b; auto.
+ unfold rs3. destruct b; auto.
+ split. unfold rs3. Simpl. destruct b. Simpl. unfold rs2. repeat Simpl.
+ intros. unfold rs3. Simpl. transitivity (rs2 r).
+ destruct b; auto; Simpl. unfold rs2. repeat Simpl.
Qed.
Remark val_add_add_zero:
@@ -1191,15 +1086,15 @@ Remark val_add_add_zero:
Proof.
intros. destruct v1; destruct v2; simpl; auto; rewrite Int.add_zero; auto.
Qed.
-
+
Lemma transl_load_store_correct:
forall (mk_instr_imm: ireg -> int -> instruction)
(mk_instr_gen: option (ireg -> shift_addr -> instruction))
(is_immed: int -> bool)
addr args k ms sp rs m ms' m',
(forall (r1: ireg) (rs1: regset) n k,
- eval_addressing_total sp addr (map ms args) = Val.add rs1#r1 (Vint n) ->
- agree ms sp rs1 ->
+ eval_addressing_total sp addr (map rs (map preg_of args)) = Val.add rs1#r1 (Vint n) ->
+ (forall (r: preg), r <> PC -> r <> IR14 -> rs1 r = rs r) ->
exists rs',
exec_straight (mk_instr_imm r1 n :: k) rs1 m k rs' m' /\
agree ms' sp rs') ->
@@ -1207,8 +1102,8 @@ Lemma transl_load_store_correct:
| None => True
| Some mk =>
(forall (r1: ireg) (sa: shift_addr) (rs1: regset) k,
- eval_addressing_total sp addr (map ms args) = Val.add rs1#r1 (eval_shift_addr sa rs1) ->
- agree ms sp rs1 ->
+ eval_addressing_total sp addr (map rs (map preg_of args)) = Val.add rs1#r1 (eval_shift_addr sa rs1) ->
+ (forall (r: preg), r <> PC -> r <> IR14 -> rs1 r = rs r) ->
exists rs',
exec_straight (mk r1 sa :: k) rs1 m k rs' m' /\
agree ms' sp rs')
@@ -1224,62 +1119,53 @@ Proof.
(* Aindexed *)
case (is_immed i).
(* Aindexed, small displacement *)
- apply H; eauto. simpl. rewrite (ireg_val ms sp rs); auto.
+ apply H; auto.
(* Aindexed, large displacement *)
- exploit (addimm_correct IR14 (ireg_of t)); eauto with ppcgen.
- intros [rs' [A [B C]]].
- exploit (H IR14 rs' Int.zero); eauto.
- simpl. rewrite (ireg_val ms sp rs); auto. rewrite B.
- rewrite Val.add_assoc. simpl Val.add. rewrite Int.add_zero. reflexivity.
- apply agree_exten_2 with rs; auto.
+ destruct (addimm_correct IR14 (ireg_of m0) i (mk_instr_imm IR14 Int.zero :: k) rs m)
+ as [rs' [A [B C]]].
+ exploit (H IR14 rs' Int.zero); eauto.
+ rewrite B. rewrite Val.add_assoc. simpl Val.add. rewrite Int.add_zero. reflexivity.
intros [rs'' [D E]].
exists rs''; split.
eapply exec_straight_trans. eexact A. eexact D. auto.
(* Aindexed2 *)
destruct mk_instr_gen as [mk | ].
(* binary form available *)
- apply H0; auto. simpl. repeat rewrite (ireg_val ms sp rs); auto.
+ apply H0; auto.
(* binary form not available *)
- set (rs' := nextinstr (rs#IR14 <- (Val.add (ms t) (ms t0)))).
+ set (rs' := nextinstr (rs#IR14 <- (Val.add (rs (ireg_of m0)) (rs (ireg_of m1))))).
exploit (H IR14 rs' Int.zero); eauto.
- simpl. repeat rewrite (ireg_val ms sp rs); auto.
unfold rs'. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
- repeat rewrite (ireg_val ms sp rs); auto. apply val_add_add_zero.
- unfold rs'; auto with ppcgen.
+ apply val_add_add_zero.
+ unfold rs'. intros. repeat Simpl.
intros [rs'' [A B]].
exists rs''; split.
eapply exec_straight_step with (rs2 := rs'); eauto.
- simpl. repeat rewrite <- (ireg_val ms sp rs); auto.
- auto.
+ simpl. auto. auto.
(* Aindexed2shift *)
destruct mk_instr_gen as [mk | ].
(* binary form available *)
- apply H0; auto. simpl. repeat rewrite (ireg_val ms sp rs); auto.
- rewrite transl_shift_addr_correct. auto.
+ apply H0; auto. rewrite transl_shift_addr_correct. auto.
(* binary form not available *)
- set (rs' := nextinstr (rs#IR14 <- (Val.add (ms t) (eval_shift_total s (ms t0))))).
+ set (rs' := nextinstr (rs#IR14 <- (Val.add (rs (ireg_of m0)) (eval_shift_total s (rs (ireg_of m1)))))).
exploit (H IR14 rs' Int.zero); eauto.
- simpl. repeat rewrite (ireg_val ms sp rs); auto.
unfold rs'. rewrite nextinstr_inv; auto with ppcgen. rewrite Pregmap.gss.
- repeat rewrite (ireg_val ms sp rs); auto. apply val_add_add_zero.
- unfold rs'; auto with ppcgen.
+ apply val_add_add_zero.
+ unfold rs'; intros; repeat Simpl.
intros [rs'' [A B]].
exists rs''; split.
eapply exec_straight_step with (rs2 := rs'); eauto.
- simpl. rewrite transl_shift_correct.
- repeat rewrite <- (ireg_val ms sp rs); auto.
+ simpl. rewrite transl_shift_correct. auto.
auto.
(* Ainstack *)
destruct (is_immed i).
(* Ainstack, short displacement *)
- apply H. simpl. rewrite (sp_val ms sp rs); auto. auto.
+ apply H; auto. rewrite (sp_val _ _ _ H1). auto.
(* Ainstack, large displacement *)
- exploit (addimm_correct IR14 IR13); eauto with ppcgen.
- intros [rs' [A [B C]]].
+ destruct (addimm_correct IR14 IR13 i (mk_instr_imm IR14 Int.zero :: k) rs m)
+ as [rs' [A [B C]]].
exploit (H IR14 rs' Int.zero); eauto.
- simpl. rewrite (sp_val ms sp rs); auto. rewrite B.
- rewrite Val.add_assoc. simpl Val.add. rewrite Int.add_zero. reflexivity.
- apply agree_exten_2 with rs; auto.
+ rewrite (sp_val _ _ _ H1). rewrite B. rewrite Val.add_assoc. simpl Val.add. rewrite Int.add_zero. auto.
intros [rs'' [D E]].
exists rs''; split.
eapply exec_straight_trans. eexact A. eexact D. auto.
@@ -1288,116 +1174,159 @@ Qed.
Lemma transl_load_int_correct:
forall (mk_instr: ireg -> ireg -> shift_addr -> instruction)
(is_immed: int -> bool)
- (rd: mreg) addr args k ms sp rs m chunk a v,
+ (rd: mreg) addr args k ms sp rs m m' chunk a v,
(forall (c: code) (r1 r2: ireg) (sa: shift_addr) (rs1: regset),
- exec_instr ge c (mk_instr r1 r2 sa) rs1 m =
- exec_load chunk (Val.add rs1#r2 (eval_shift_addr sa rs1)) r1 rs1 m) ->
+ exec_instr ge c (mk_instr r1 r2 sa) rs1 m' =
+ exec_load chunk (Val.add rs1#r2 (eval_shift_addr sa rs1)) r1 rs1 m') ->
agree ms sp rs ->
map mreg_type args = type_of_addressing addr ->
mreg_type rd = Tint ->
eval_addressing ge sp addr (map ms args) = Some a ->
Mem.loadv chunk m a = Some v ->
+ Mem.extends m m' ->
exists rs',
- exec_straight (transl_load_store_int mk_instr is_immed rd addr args k) rs m
- k rs' m
- /\ agree (Regmap.set rd v ms) sp rs'.
+ exec_straight (transl_load_store_int mk_instr is_immed rd addr args k) rs m'
+ k rs' m'
+ /\ agree (Regmap.set rd v (undef_temps ms)) sp rs'.
Proof.
intros. unfold transl_load_store_int.
- exploit eval_addressing_weaken. eauto. intros.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eauto.
+ intros [a' [A B]].
+ exploit Mem.loadv_extends; eauto. intros [v' [C D]].
+ exploit eval_addressing_weaken. eexact A. intros E.
apply transl_load_store_correct with ms; auto.
- intros. exists (nextinstr (rs1#(ireg_of rd) <- v)); split.
- apply exec_straight_one. rewrite H. simpl. rewrite <- H6. rewrite H5.
- unfold exec_load. rewrite H4. auto. auto.
- auto with ppcgen.
- intros. exists (nextinstr (rs1#(ireg_of rd) <- v)); split.
- apply exec_straight_one. rewrite H. simpl. rewrite <- H6. rewrite H5.
- unfold exec_load. rewrite H4. auto. auto.
- auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (Vint n) = a') by congruence.
+ exists (nextinstr (rs1#(ireg_of rd) <- v')); split.
+ apply exec_straight_one. rewrite H. unfold exec_load.
+ simpl. rewrite H8. rewrite C. auto. auto.
+ apply agree_nextinstr. eapply agree_set_undef_mreg; eauto.
+ unfold preg_of. rewrite H2. rewrite Pregmap.gss. auto.
+ unfold preg_of. rewrite H2. intros. rewrite Pregmap.gso; auto. apply H7; auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (eval_shift_addr sa rs1) = a') by congruence.
+ exists (nextinstr (rs1#(ireg_of rd) <- v')); split.
+ apply exec_straight_one. rewrite H. unfold exec_load.
+ simpl. rewrite H8. rewrite C. auto. auto.
+ apply agree_nextinstr. eapply agree_set_undef_mreg; eauto.
+ unfold preg_of. rewrite H2. rewrite Pregmap.gss. auto.
+ unfold preg_of. rewrite H2. intros. rewrite Pregmap.gso; auto. apply H7; auto with ppcgen.
Qed.
Lemma transl_load_float_correct:
forall (mk_instr: freg -> ireg -> int -> instruction)
(is_immed: int -> bool)
- (rd: mreg) addr args k ms sp rs m chunk a v,
+ (rd: mreg) addr args k ms sp rs m m' chunk a v,
(forall (c: code) (r1: freg) (r2: ireg) (n: int) (rs1: regset),
- exec_instr ge c (mk_instr r1 r2 n) rs1 m =
- exec_load chunk (Val.add rs1#r2 (Vint n)) r1 rs1 m) ->
+ exec_instr ge c (mk_instr r1 r2 n) rs1 m' =
+ exec_load chunk (Val.add rs1#r2 (Vint n)) r1 rs1 m') ->
agree ms sp rs ->
map mreg_type args = type_of_addressing addr ->
mreg_type rd = Tfloat ->
eval_addressing ge sp addr (map ms args) = Some a ->
Mem.loadv chunk m a = Some v ->
+ Mem.extends m m' ->
exists rs',
- exec_straight (transl_load_store_float mk_instr is_immed rd addr args k) rs m
- k rs' m
- /\ agree (Regmap.set rd v ms) sp rs'.
+ exec_straight (transl_load_store_float mk_instr is_immed rd addr args k) rs m'
+ k rs' m'
+ /\ agree (Regmap.set rd v (undef_temps ms)) sp rs'.
Proof.
- intros. unfold transl_load_store_float.
- exploit eval_addressing_weaken. eauto. intros.
+ intros. unfold transl_load_store_int.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eauto.
+ intros [a' [A B]].
+ exploit Mem.loadv_extends; eauto. intros [v' [C D]].
+ exploit eval_addressing_weaken. eexact A. intros E.
apply transl_load_store_correct with ms; auto.
- intros. exists (nextinstr (rs1#(freg_of rd) <- v)); split.
- apply exec_straight_one. rewrite H. rewrite <- H6. rewrite H5.
- unfold exec_load. rewrite H4. auto. auto.
- auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (Vint n) = a') by congruence.
+ exists (nextinstr (rs1#(freg_of rd) <- v')); split.
+ apply exec_straight_one. rewrite H. unfold exec_load.
+ simpl. rewrite H8. rewrite C. auto. auto.
+ apply agree_nextinstr. eapply agree_set_undef_mreg; eauto.
+ unfold preg_of. rewrite H2. rewrite Pregmap.gss. auto.
+ unfold preg_of. rewrite H2. intros. rewrite Pregmap.gso; auto. apply H7; auto with ppcgen.
Qed.
Lemma transl_store_int_correct:
forall (mk_instr: ireg -> ireg -> shift_addr -> instruction)
(is_immed: int -> bool)
- (rd: mreg) addr args k ms sp rs m chunk a m',
+ (rd: mreg) addr args k ms sp rs m1 chunk a m2 m1',
(forall (c: code) (r1 r2: ireg) (sa: shift_addr) (rs1: regset),
- exec_instr ge c (mk_instr r1 r2 sa) rs1 m =
- exec_store chunk (Val.add rs1#r2 (eval_shift_addr sa rs1)) r1 rs1 m) ->
+ exec_instr ge c (mk_instr r1 r2 sa) rs1 m1' =
+ exec_store chunk (Val.add rs1#r2 (eval_shift_addr sa rs1)) r1 rs1 m1') ->
agree ms sp rs ->
map mreg_type args = type_of_addressing addr ->
mreg_type rd = Tint ->
eval_addressing ge sp addr (map ms args) = Some a ->
- Mem.storev chunk m a (ms rd) = Some m' ->
- exists rs',
- exec_straight (transl_load_store_int mk_instr is_immed rd addr args k) rs m
- k rs' m'
- /\ agree ms sp rs'.
+ Mem.storev chunk m1 a (ms rd) = Some m2 ->
+ Mem.extends m1 m1' ->
+ exists m2',
+ Mem.extends m2 m2' /\
+ exists rs',
+ exec_straight (transl_load_store_int mk_instr is_immed rd addr args k) rs m1'
+ k rs' m2'
+ /\ agree (undef_temps ms) sp rs'.
Proof.
intros. unfold transl_load_store_int.
- exploit eval_addressing_weaken. eauto. intros.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eauto.
+ intros [a' [A B]].
+ exploit preg_val; eauto. instantiate (1 := rd). intros C.
+ exploit Mem.storev_extends; eauto. unfold preg_of; rewrite H2. intros [m2' [D E]].
+ exploit eval_addressing_weaken. eexact A. intros F.
+ exists m2'; split; auto.
apply transl_load_store_correct with ms; auto.
- intros. exists (nextinstr rs1); split.
- apply exec_straight_one. rewrite H. simpl. rewrite <- H6. rewrite H5.
- unfold exec_store. rewrite <- (ireg_val ms sp rs1); auto.
- rewrite H4. auto. auto.
- auto with ppcgen.
- intros. exists (nextinstr rs1); split.
- apply exec_straight_one. rewrite H. simpl. rewrite <- H6. rewrite H5.
- unfold exec_store. rewrite <- (ireg_val ms sp rs1); auto.
- rewrite H4. auto. auto.
- auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (Vint n) = a') by congruence.
+ exists (nextinstr rs1); split.
+ apply exec_straight_one. rewrite H. simpl. rewrite H8.
+ unfold exec_store. rewrite H7; auto with ppcgen. rewrite D. auto. auto.
+ apply agree_nextinstr. apply agree_exten_temps with rs; auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (eval_shift_addr sa rs1) = a') by congruence.
+ exists (nextinstr rs1); split.
+ apply exec_straight_one. rewrite H. simpl. rewrite H8.
+ unfold exec_store. rewrite H7; auto with ppcgen. rewrite D. auto. auto.
+ apply agree_nextinstr. apply agree_exten_temps with rs; auto with ppcgen.
Qed.
Lemma transl_store_float_correct:
forall (mk_instr: freg -> ireg -> int -> instruction)
(is_immed: int -> bool)
- (rd: mreg) addr args k ms sp rs m chunk a m',
- (forall (c: code) (r1: freg) (r2: ireg) (n: int) (rs1: regset),
- exec_instr ge c (mk_instr r1 r2 n) rs1 m =
- exec_store chunk (Val.add rs1#r2 (Vint n)) r1 rs1 m) ->
+ (rd: mreg) addr args k ms sp rs m1 chunk a m2 m1',
+ (forall (c: code) (r1: freg) (r2: ireg) (n: int) (rs1: regset) m2',
+ exec_store chunk (Val.add rs1#r2 (Vint n)) r1 rs1 m1' = OK (nextinstr rs1) m2' ->
+ exists rs2,
+ exec_instr ge c (mk_instr r1 r2 n) rs1 m1' = OK rs2 m2'
+ /\ (forall (r: preg), r <> FR3 -> rs2 r = nextinstr rs1 r)) ->
agree ms sp rs ->
map mreg_type args = type_of_addressing addr ->
mreg_type rd = Tfloat ->
eval_addressing ge sp addr (map ms args) = Some a ->
- Mem.storev chunk m a (ms rd) = Some m' ->
- exists rs',
- exec_straight (transl_load_store_float mk_instr is_immed rd addr args k) rs m
- k rs' m'
- /\ agree ms sp rs'.
+ Mem.storev chunk m1 a (ms rd) = Some m2 ->
+ Mem.extends m1 m1' ->
+ exists m2',
+ Mem.extends m2 m2' /\
+ exists rs',
+ exec_straight (transl_load_store_float mk_instr is_immed rd addr args k) rs m1'
+ k rs' m2'
+ /\ agree (undef_temps ms) sp rs'.
Proof.
intros. unfold transl_load_store_float.
- exploit eval_addressing_weaken. eauto. intros.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eauto.
+ intros [a' [A B]].
+ exploit preg_val; eauto. instantiate (1 := rd). intros C.
+ exploit Mem.storev_extends; eauto. unfold preg_of; rewrite H2. intros [m2' [D E]].
+ exploit eval_addressing_weaken. eexact A. intros F.
+ exists m2'; split; auto.
apply transl_load_store_correct with ms; auto.
- intros. exists (nextinstr rs1); split.
- apply exec_straight_one. rewrite H. simpl. rewrite <- H6. rewrite H5.
- unfold exec_store. rewrite <- (freg_val ms sp rs1); auto.
- rewrite H4. auto. auto.
- auto with ppcgen.
+ intros.
+ assert (Val.add (rs1 r1) (Vint n) = a') by congruence.
+ exploit (H fn (freg_of rd) r1 n rs1 m2').
+ unfold exec_store. rewrite H8. rewrite H7; auto with ppcgen. rewrite D. auto.
+ intros [rs2 [P Q]].
+ exists rs2; split. apply exec_straight_one. auto. rewrite Q; auto with ppcgen.
+ apply agree_exten_temps with rs; auto.
+ intros. rewrite Q; auto with ppcgen. Simpl. apply H7; auto with ppcgen.
Qed.
End STRAIGHTLINE.