aboutsummaryrefslogtreecommitdiffstats
path: root/driver/Compiler.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2014-11-27 15:37:32 +0100
committerBernhard Schommer <bernhardschommer@gmail.com>2014-11-27 15:37:32 +0100
commit56690956f52349c3398b3de6f8ec3987501e9034 (patch)
tree5fc98e863bb41018084b2110f0ae950189a7b7d6 /driver/Compiler.v
parent853a40b117495ebf883593633f680cd5c92f5951 (diff)
parentc3b615f875ed2cf8418453c79c4621d2dc61b0a0 (diff)
downloadcompcert-56690956f52349c3398b3de6f8ec3987501e9034.tar.gz
compcert-56690956f52349c3398b3de6f8ec3987501e9034.zip
Merge branch 'master' into dwarf
Diffstat (limited to 'driver/Compiler.v')
-rw-r--r--driver/Compiler.v9
1 files changed, 8 insertions, 1 deletions
diff --git a/driver/Compiler.v b/driver/Compiler.v
index fb813c7c..0afa7bfb 100644
--- a/driver/Compiler.v
+++ b/driver/Compiler.v
@@ -46,6 +46,7 @@ Require Renumber.
Require Constprop.
Require CSE.
Require Deadcode.
+Require Unusedglob.
Require Allocation.
Require Tunneling.
Require Linearize.
@@ -65,6 +66,7 @@ Require Renumberproof.
Require Constpropproof.
Require CSEproof.
Require Deadcodeproof.
+Require Unusedglobproof.
Require Allocproof.
Require Tunnelingproof.
Require Linearizeproof.
@@ -135,6 +137,8 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program :=
@@ print (print_RTL 6)
@@@ partial_if Compopts.optim_redundancy (time "Redundancy elimination" Deadcode.transf_program)
@@ print (print_RTL 7)
+ @@@ time "Unused globals" Unusedglob.transform_program
+ @@ print (print_RTL 8)
@@@ time "Register allocation" Allocation.transf_program
@@ print print_LTL
@@ time "Branch tunneling" Tunneling.tunnel_program
@@ -244,7 +248,8 @@ Proof.
set (p21 := total_if optim_constprop Renumber.transf_program p2) in *.
destruct (partial_if optim_CSE CSE.transf_program p21) as [p3|] eqn:?; simpl in H; try discriminate.
destruct (partial_if optim_redundancy Deadcode.transf_program p3) as [p31|] eqn:?; simpl in H; try discriminate.
- destruct (Allocation.transf_program p31) as [p4|] eqn:?; simpl in H; try discriminate.
+ destruct (Unusedglob.transform_program p31) as [p32|] eqn:?; simpl in H; try discriminate.
+ destruct (Allocation.transf_program p32) as [p4|] eqn:?; simpl in H; try discriminate.
set (p5 := Tunneling.tunnel_program p4) in *.
destruct (Linearize.transf_program p5) as [p6|] eqn:?; simpl in H; try discriminate.
set (p7 := CleanupLabels.transf_program p6) in *.
@@ -263,6 +268,8 @@ Proof.
eapply partial_if_simulation; eauto. apply CSEproof.transf_program_correct.
apply compose_forward_simulation with (RTL.semantics p31).
eapply partial_if_simulation; eauto. apply Deadcodeproof.transf_program_correct.
+ apply compose_forward_simulation with (RTL.semantics p32).
+ apply Unusedglobproof.transf_program_correct; auto.
apply compose_forward_simulation with (LTL.semantics p4).
apply Allocproof.transf_program_correct; auto.
apply compose_forward_simulation with (LTL.semantics p5).