aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling/BTL_Scheduler.v
diff options
context:
space:
mode:
authorSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-05-28 16:33:23 +0200
committerSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-05-28 16:33:23 +0200
commit2311cb788e6dcf0103266fd6c6aa76096c283e34 (patch)
tree83d7cab703605fb293459fd5287244722517fbe6 /scheduling/BTL_Scheduler.v
parent5b67f8284c3a98581f4da9b065a738fc534480c4 (diff)
downloadcompcert-kvx-2311cb788e6dcf0103266fd6c6aa76096c283e34.tar.gz
compcert-kvx-2311cb788e6dcf0103266fd6c6aa76096c283e34.zip
remove dupmap from BTL_Scheduler !
Diffstat (limited to 'scheduling/BTL_Scheduler.v')
-rw-r--r--scheduling/BTL_Scheduler.v32
1 files changed, 10 insertions, 22 deletions
diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v
index 8cf9635c..faeb58cd 100644
--- a/scheduling/BTL_Scheduler.v
+++ b/scheduling/BTL_Scheduler.v
@@ -5,36 +5,24 @@ Require Import RTL Op Registers OptionMonad BTL.
Require Import Errors Linking BTL_SEtheory.
(** External oracle *)
-Axiom scheduler: BTL.function -> BTL.code * node * (PTree.t node).
+Axiom scheduler: BTL.function -> BTL.code.
-(* FAKE: A REVOIR avec BTL_SEtheory... *)
-Definition symbolic_simu (dupmap: PTree.t node) (f1 f2: BTL.function) (pc1 pc2: node): Prop := True.
-
-Record match_function (dupmap: PTree.t node) (f1 f2: BTL.function): Prop := {
+(* a specification of the verification to do on each function *)
+Record match_function (f1 f2: BTL.function): Prop := {
preserv_fnsig: fn_sig f1 = fn_sig f2;
preserv_fnparams: fn_params f1 = fn_params f2;
preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2;
- preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint);
- dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> symbolic_simu dupmap f1 f2 pc1 pc2;
+ preserv_entrypoint: fn_entrypoint f1 = fn_entrypoint f2;
+ symbolic_simu_ok: 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);
}.
-Definition verified_scheduler (f: BTL.function) :=
- let (tcte, dupmap) := scheduler f in
- let (tc, te) := tcte in
- let f' := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in
- OK (f', dupmap). (* TODO: fixme *)
-
-Program Definition transf_function (f: BTL.function):
- { r : res BTL.function | forall f', r = OK f' -> exists dm, match_function dm f f'} :=
- match (verified_scheduler f) with
- | Error e => Error e
- | OK (tf, dm) => OK tf
- end.
-Next Obligation.
-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)).
Definition transf_fundef (f: fundef) : res fundef :=
- transf_partial_fundef (fun f => proj1_sig (transf_function f)) f.
+ transf_partial_fundef (fun f => transf_function f) f.
Definition transf_program (p: program) : res program :=
transform_partial_program transf_fundef p.