aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling
diff options
context:
space:
mode:
authorLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-07-21 12:16:27 +0200
committerLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-07-21 12:16:27 +0200
commit9ce1b23f6361d8070490df1269a7bc53aa215efb (patch)
tree264365faee227986865bad3940ffb5e112c7adf0 /scheduling
parent8663f220ec0b0f10e74ae7442cdf221a83fc3943 (diff)
downloadcompcert-kvx-9ce1b23f6361d8070490df1269a7bc53aa215efb.tar.gz
compcert-kvx-9ce1b23f6361d8070490df1269a7bc53aa215efb.zip
expansions btl proofs
Diffstat (limited to 'scheduling')
-rw-r--r--scheduling/BTL_SEimpl.v6
-rw-r--r--scheduling/BTL_Scheduler.v2
-rw-r--r--scheduling/BTL_Schedulerproof.v6
-rw-r--r--scheduling/BTLmatchRTL.v4
4 files changed, 10 insertions, 8 deletions
diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v
index 3050f2e2..7711c72d 100644
--- a/scheduling/BTL_SEimpl.v
+++ b/scheduling/BTL_SEimpl.v
@@ -613,7 +613,11 @@ Proof.
{ wlp_simplify; exploit is_move_operation_correct; eauto.
intros (Hop & Hlsv); subst; simpl in *. inv REF.
simplify_SOME z; erewrite H; eauto. }
- wlp_simplify; inv REF. erewrite H0; eauto.
+ destruct (target_op_simplify _ _ _) eqn: Htarget_op_simp; wlp_simplify.
+ { destruct (eval_list_sval _ _) eqn: OKlist; try congruence.
+ rewrite <- H; exploit target_op_simplify_correct; eauto. }
+ inv REF; clear Htarget_op_simp.
+ intuition eauto.
- (* Rload *)
destruct trap; wlp_simplify; inv REF.
+ erewrite H0; eauto.
diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v
index 39672749..a3053add 100644
--- a/scheduling/BTL_Scheduler.v
+++ b/scheduling/BTL_Scheduler.v
@@ -113,7 +113,7 @@ Qed.
Definition transf_function (f: BTL.function) :=
let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in
- (*do _ <- check_symbolic_simu f tf;*)
+ do _ <- check_symbolic_simu f tf;
OK tf.
Definition transf_fundef (f: fundef) : res fundef :=
diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v
index 9b0383ae..05028f11 100644
--- a/scheduling/BTL_Schedulerproof.v
+++ b/scheduling/BTL_Schedulerproof.v
@@ -98,19 +98,17 @@ Proof.
unfold transf_function. intros H. monadInv H.
econstructor; eauto.
eapply check_symbolic_simu_input_equiv; eauto.
-Admitted. (*
eapply check_symbolic_simu_correct; eauto.
- Qed.*)
+Qed.
Lemma transf_fundef_correct f f':
transf_fundef f = OK f' -> match_fundef f f'.
Proof.
intros TRANSF; destruct f; simpl; inv TRANSF.
-Admitted. (*
+ monadInv H0. exploit transf_function_correct; eauto.
constructor; auto.
+ eapply match_External.
- Qed.*)
+Qed.
Lemma function_ptr_preserved:
forall v f,
diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v
index 439ba9cc..a59c847e 100644
--- a/scheduling/BTLmatchRTL.v
+++ b/scheduling/BTLmatchRTL.v
@@ -412,7 +412,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node)
else Error (msg "BTL.verify_block: different r in Bload")
else Error (msg "BTL.verify_block: different lr in Bload")
else Error (msg "BTL.verify_block: different addressing in Bload")
- else Error (msg "BTL.verify_block: different mchunk in Bload")
+ else Error (msg "BTL.verify_block: different chunk in Bload")
else Error (msg "BTL.verify_block: NOTRAP trapping_mode unsupported in Bload")
| _ => Error (msg "BTL.verify_block: incorrect cfg Bload")
end
@@ -426,7 +426,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node)
else Error (msg "BTL.verify_block: different r in Bstore")
else Error (msg "BTL.verify_block: different lr in Bstore")
else Error (msg "BTL.verify_block: different addressing in Bstore")
- else Error (msg "BTL.verify_block: different mchunk in Bstore")
+ else Error (msg "BTL.verify_block: different chunk in Bstore")
| _ => Error (msg "BTL.verify_block: incorrect cfg Bstore")
end
| Bseq b1 b2 =>