aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Coloringproof.v
diff options
context:
space:
mode:
authorblazy <blazy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-01-08 07:53:02 +0000
committerblazy <blazy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-01-08 07:53:02 +0000
commita8c744000247af207b489d3cdd4e3d3cf60f72e1 (patch)
tree96c7ee4e244fccdb840233007604ba52d97c09e0 /backend/Coloringproof.v
parent283afabc594b385e4f17fa59647aa8cddee27f85 (diff)
downloadcompcert-a8c744000247af207b489d3cdd4e3d3cf60f72e1.tar.gz
compcert-a8c744000247af207b489d3cdd4e3d3cf60f72e1.zip
ajout branche allocation de registres
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1220 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/Coloringproof.v')
-rw-r--r--backend/Coloringproof.v19
1 files changed, 14 insertions, 5 deletions
diff --git a/backend/Coloringproof.v b/backend/Coloringproof.v
index 92ac0676..cc4234ec 100644
--- a/backend/Coloringproof.v
+++ b/backend/Coloringproof.v
@@ -24,6 +24,8 @@ Require Import Locations.
Require Import Conventions.
Require Import InterfGraph.
Require Import Coloring.
+Require Import InterfGraph_Construction.
+Require Import MyAllocation.
(** * Correctness of the interference graph *)
@@ -446,12 +448,12 @@ Proof.
intros until g2. intro.
unfold correct_interf_instr; destruct instr; auto.
destruct (is_move_operation o l).
- intros. eapply interfere_incl; eauto.
- intros. eapply interfere_incl; eauto.
- intros. eapply interfere_incl; eauto.
+ intros. apply interfere_incl with (g1 := g1); auto.
+ intros. apply interfere_incl with (g1 := g1); auto.
+ intros. apply interfere_incl with (g1 := g1); auto.
intros [A [B C]].
- split. intros. eapply interfere_mreg_incl; eauto.
- split. intros. eapply interfere_incl; eauto.
+ split. intros. apply interfere_mreg_incl with (g1 := g1); auto.
+ split. intros. apply interfere_incl with (g1 := g1); auto.
destruct s0; auto. intros. eapply interfere_mreg_incl; eauto.
destruct s0; auto. intros. eapply interfere_mreg_incl; eauto.
Qed.
@@ -812,16 +814,23 @@ Let g := interf_graph f live live0.
Let allregs := all_interf_regs g.
Let coloring := graph_coloring f g env allregs.
+
Lemma regalloc_ok:
regalloc f live live0 env = Some alloc ->
check_coloring g env allregs coloring = true /\
alloc = alloc_of_coloring coloring env allregs.
Proof.
+unfold regalloc. intro.
+inversion H. subst. clear H.
+split. apply allocation_correct.
+auto.
+(*
unfold regalloc, coloring, allregs, g.
case (check_coloring (interf_graph f live live0) env).
intro EQ; injection EQ; intro; clear EQ.
split. auto. auto.
intro; discriminate.
+*)
Qed.
Lemma regalloc_acceptable: