aboutsummaryrefslogtreecommitdiffstats
path: root/driver/Compiler.vexpand
diff options
context:
space:
mode:
Diffstat (limited to 'driver/Compiler.vexpand')
-rw-r--r--driver/Compiler.vexpand10
1 files changed, 9 insertions, 1 deletions
diff --git a/driver/Compiler.vexpand b/driver/Compiler.vexpand
index 00db9b36..a751b232 100644
--- a/driver/Compiler.vexpand
+++ b/driver/Compiler.vexpand
@@ -54,7 +54,7 @@ Require Import Compopts.
Parameter print_Clight: Clight.program -> unit.
Parameter print_Cminor: Cminor.program -> unit.
Parameter print_RTL: Z -> RTL.program -> unit.
-Parameter print_LTL: LTL.program -> unit.
+Parameter print_LTL: Z -> LTL.program -> unit.
Parameter print_Mach: Mach.program -> unit.
Local Open Scope string_scope.
@@ -298,6 +298,14 @@ EXPAND_ASM_SEMANTICS
eapply RTLgenproof.transf_program_correct; eassumption.
EXPAND_RTL_FORWARD_SIMULATIONS
eapply compose_forward_simulations.
+ eapply RTLpathLivegenproof.transf_program_correct; eassumption.
+ pose proof RTLpathLivegenproof.all_fundef_liveness_ok as X.
+ refine (modusponens _ _ (X _ _ _) _); eauto. intro.
+ eapply compose_forward_simulations.
+ eapply RTLpathSchedulerproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply RTLpathproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
eapply Allocationproof.transf_program_correct; eassumption.
eapply compose_forward_simulations.
eapply Tunnelingproof.transf_program_correct; eassumption.