aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling/BTL_Scheduler.v
diff options
context:
space:
mode:
authorSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-05-28 16:49:31 +0200
committerSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-05-28 16:49:31 +0200
commita97cea76d531b3ea985c284d76baa6370f7dc489 (patch)
tree328c2a130f8863233eae171a152c0acc9ba35b79 /scheduling/BTL_Scheduler.v
parent2311cb788e6dcf0103266fd6c6aa76096c283e34 (diff)
downloadcompcert-kvx-a97cea76d531b3ea985c284d76baa6370f7dc489.tar.gz
compcert-kvx-a97cea76d531b3ea985c284d76baa6370f7dc489.zip
declare a checker for the symbolic simulation
Diffstat (limited to 'scheduling/BTL_Scheduler.v')
-rw-r--r--scheduling/BTL_Scheduler.v15
1 files changed, 13 insertions, 2 deletions
diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v
index faeb58cd..43d6dd5e 100644
--- a/scheduling/BTL_Scheduler.v
+++ b/scheduling/BTL_Scheduler.v
@@ -17,9 +17,20 @@ Record match_function (f1 f2: BTL.function): Prop := {
exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2);
}.
+Local Open Scope error_monad_scope.
+
+Definition check_symbolic_simu (f tf: BTL.function): res unit := OK tt. (* TODO: fixme *)
+
+Lemma check_symbolic_simu_correct x f1 f2:
+ check_symbolic_simu f1 f2 = OK x ->
+ forall pc ib1, (fn_code f1)!pc = Some ib1 ->
+ exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2).
+Admitted.
+
Definition transf_function (f: BTL.function) :=
- (* TODO: fixme *)
- OK (BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f)).
+ 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;
+ OK tf.
Definition transf_fundef (f: fundef) : res fundef :=
transf_partial_fundef (fun f => transf_function f) f.