aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling/BTL_Scheduler.v
diff options
context:
space:
mode:
authorLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-06-17 17:36:42 +0200
committerLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-06-17 17:36:42 +0200
commitef95bd7f4afe159bcedc3ec5732579bfc6ba08c6 (patch)
treee9c2e0d195d402913dffcc2c8bfd2c75a5804a93 /scheduling/BTL_Scheduler.v
parent870d0cf9d06f453e2ba3cbdaf0184bc1f657c04b (diff)
downloadcompcert-kvx-ef95bd7f4afe159bcedc3ec5732579bfc6ba08c6.tar.gz
compcert-kvx-ef95bd7f4afe159bcedc3ec5732579bfc6ba08c6.zip
some advance, new section to simplify context from symbolic exec
Diffstat (limited to 'scheduling/BTL_Scheduler.v')
-rw-r--r--scheduling/BTL_Scheduler.v12
1 files changed, 7 insertions, 5 deletions
diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v
index ec83b3c1..10f780e2 100644
--- a/scheduling/BTL_Scheduler.v
+++ b/scheduling/BTL_Scheduler.v
@@ -9,7 +9,6 @@ Axiom scheduler: BTL.function -> BTL.code.
Extract Constant scheduler => "BTLScheduleraux.btl_scheduler".
-(* TODO: could be useful ?
Definition equiv_input_regs (f1 f2: BTL.function): Prop :=
(forall pc, (fn_code f1)!pc = None <-> (fn_code f2)!pc = None)
/\ (forall pc ib1 ib2, (fn_code f1)!pc = Some ib1 -> (fn_code f2)!pc = Some ib2 -> ib1.(input_regs) = ib2.(input_regs)).
@@ -28,8 +27,13 @@ Lemma equiv_input_regs_pre f1 f2 tbl or:
Proof.
intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto.
Qed.
-*)
+Lemma equiv_input_regs_tr_inputs f1 f2 l oreg rs:
+ equiv_input_regs f1 f2 ->
+ tr_inputs f1 l oreg rs = tr_inputs f2 l oreg rs.
+Proof.
+ intros; unfold tr_inputs; erewrite equiv_input_regs_pre; eauto.
+Qed.
(* a specification of the verification to do on each function *)
Record match_function (f1 f2: BTL.function): Prop := {
@@ -37,7 +41,7 @@ Record match_function (f1 f2: BTL.function): Prop := {
preserv_fnparams: fn_params f1 = fn_params f2;
preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2;
preserv_entrypoint: fn_entrypoint f1 = fn_entrypoint f2;
- (* preserv_inputs: equiv_input_regs f1 f2; TODO: a-t-on besoin de ça ? *)
+ preserv_inputs: equiv_input_regs f1 f2; (* TODO: a-t-on besoin de ça ? *)
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);
}.
@@ -46,11 +50,9 @@ Local Open Scope error_monad_scope.
Definition check_symbolic_simu (f tf: BTL.function): res unit := OK tt. (* TODO: fixme *)
-(* TODO: a-t-on besoin de ça ?
Lemma check_symbolic_simu_input_equiv x f1 f2:
check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2.
Admitted.
-*)
Lemma check_symbolic_simu_correct x f1 f2:
check_symbolic_simu f1 f2 = OK x ->