From 5b67f8284c3a98581f4da9b065a738fc534480c4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 15:56:44 +0200 Subject: archi pour la verif du scheduler --- scheduling/BTL_Scheduler.v | 40 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) (limited to 'scheduling/BTL_Scheduler.v') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index e69de29b..8cf9635c 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -0,0 +1,40 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +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). + +(* 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 := { + 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; +}. + +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_fundef (f: fundef) : res fundef := + transf_partial_fundef (fun f => proj1_sig (transf_function f)) f. + +Definition transf_program (p: program) : res program := + transform_partial_program transf_fundef p. -- cgit