aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2020-06-09 17:32:11 +0200
committerCyril SIX <cyril.six@kalray.eu>2020-06-09 17:32:11 +0200
commitd5f3ee00bb93b07f6f1a5859750cac345ed261aa (patch)
tree7ac350f9702cde48203484ab0380a4eab22f0992
parent7fd913b3e2377104f0fb2aa6cbbcd4f49f53a1f6 (diff)
downloadcompcert-kvx-d5f3ee00bb93b07f6f1a5859750cac345ed261aa.tar.gz
compcert-kvx-d5f3ee00bb93b07f6f1a5859750cac345ed261aa.zip
RTLpathSchedulerproof.all_fundef_liveness_ok is a hypothesis again
-rw-r--r--driver/Compiler.vexpand2
-rw-r--r--kvx/lib/RTLpathSchedulerproof.v14
2 files changed, 7 insertions, 9 deletions
diff --git a/driver/Compiler.vexpand b/driver/Compiler.vexpand
index 3285a012..8aaa40c6 100644
--- a/driver/Compiler.vexpand
+++ b/driver/Compiler.vexpand
@@ -298,6 +298,8 @@ EXPAND_ASM_SEMANTICS
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.
diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v
index a1a67b2f..4d7dfa2f 100644
--- a/kvx/lib/RTLpathSchedulerproof.v
+++ b/kvx/lib/RTLpathSchedulerproof.v
@@ -22,11 +22,7 @@ Hypothesis TRANSL: match_prog prog tprog.
Let pge := Genv.globalenv prog.
Let tpge := Genv.globalenv tprog.
-(* Was a Hypothesis *)
-Theorem all_fundef_liveness_ok: forall b fd,
- Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd.
-Proof.
-Admitted.
+Hypothesis all_fundef_liveness_ok: forall b fd, Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd.
Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s.
Proof.
@@ -97,7 +93,7 @@ Proof.
- constructor; eauto.
+ constructor.
+ apply transf_fundef_correct; auto.
- + eapply all_fundef_liveness_ok; eauto.
+(* + eapply all_fundef_liveness_ok; eauto. *)
Qed.
Theorem transf_final_states s1 s2 r:
@@ -153,11 +149,11 @@ Proof.
intros; exploit functions_preserved; eauto.
intros (fd' & cunit & X). eexists. intuition eauto.
eapply find_funct_liveness_ok; eauto.
- intros. eapply all_fundef_liveness_ok; eauto.
+(* intros. eapply all_fundef_liveness_ok; eauto. *)
+ rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence.
intros; exploit function_ptr_preserved; eauto.
intros (fd' & X). eexists. intuition eauto.
- intros. eapply all_fundef_liveness_ok; eauto.
+(* intros. eapply all_fundef_liveness_ok; eauto. *)
Qed.
Lemma sistate_simu f dupmap sp st st' rs m is:
@@ -271,7 +267,7 @@ Proof.
destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS.
(* exec_path *)
- try_simplify_someHyps. intros.
- exploit path_step_eqlive; eauto. { intros. eapply all_fundef_liveness_ok; eauto. }
+ exploit path_step_eqlive; eauto. (* { intros. eapply all_fundef_liveness_ok; eauto. } *)
clear STEP EQUIV rs; intros (s2 & STEP & EQLIVE).
exploit exec_path_simulation; eauto.
clear STEP; intros (path' & s' & PATH' & STEP' & MATCH').