summaryrefslogtreecommitdiff
path: root/backend/Allocproof.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 15:35:09 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 15:35:09 +0000
commit4b119d6f9f0e846edccaf5c08788ca1615b22526 (patch)
tree66cf55decd8d950d0bdc1050448aa3ee448ca13a /backend/Allocproof.v
parent1fe28ba1ec3dd0657b121c4a911ee1cb046bab09 (diff)
Cil2Csyntax: added goto and labels; added assignment between structs
Kildall: simplified the interface Constprop, CSE, Allocation, Linearize: updated for the new Kildall RTL, LTL: removed the well-formedness condition on the CFG, it is no longer necessary with the new Kildall and it is problematic for validated optimizations. Maps: more efficient implementation of PTree.fold. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1124 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/Allocproof.v')
-rw-r--r--backend/Allocproof.v21
1 files changed, 9 insertions, 12 deletions
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 7e9334a..fc0a0f3 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -154,21 +154,19 @@ Proof.
Qed.
Lemma agree_succ:
- forall n s rs ls live,
+ forall n s rs ls live i,
analyze f = Some live ->
- In s (RTL.successors f n) ->
+ f.(RTL.fn_code)!n = Some i ->
+ In s (RTL.successors_instr i) ->
agree live!!n rs ls ->
agree (transfer f s live!!s) rs ls.
Proof.
intros.
- elim (RTL.fn_code_wf f n); intro.
- elim (RTL.fn_code_wf f s); intro.
apply agree_increasing with (live!!n).
eapply DS.fixpoint_solution. unfold analyze in H; eauto.
- auto. auto. auto. auto.
- unfold transfer. rewrite H3.
- red; intros. elim (Regset.empty_1 H4).
- unfold RTL.successors in H0; rewrite H2 in H0; elim H0.
+ unfold RTL.successors, Kildall.successors_list.
+ rewrite PTree.gmap. rewrite H0. simpl. auto.
+ auto.
Qed.
(** Some useful special cases of [agree_increasing]. *)
@@ -543,12 +541,11 @@ Ltac MatchStates :=
eapply match_states_intro; eauto; MatchStates
| H: (PTree.get ?pc _ = Some _) |- agree _ _ _ _ =>
eapply agree_succ with (n := pc); eauto; MatchStates
- | H: (PTree.get _ _ = Some _) |- In _ (RTL.successors _ _) =>
- unfold RTL.successors; rewrite H; auto with coqlib
+ | |- In _ (RTL.successors_instr _) =>
+ unfold RTL.successors_instr; auto with coqlib
| _ => idtac
end.
-
Lemma transl_find_function:
forall ros f args lv rs ls alloc,
RTL.find_function ge ros rs = Some f ->
@@ -659,7 +656,7 @@ Proof.
econstructor; eauto.
econstructor; eauto.
intros. eapply agree_succ with (n := pc); eauto.
- unfold RTL.successors; rewrite H; auto with coqlib.
+ simpl. auto.
eapply agree_postcall; eauto.
(* Itailcall *)