From 30baa719fb15c45b13cb869056e51ec7446c0207 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 11:46:15 +0000 Subject: Add more documentation --- docs/scheduler-languages.org | 199 +++++++++++++++++++++---------------------- docs/scheduler.org | 54 ++++++++---- 2 files changed, 134 insertions(+), 119 deletions(-) (limited to 'docs') diff --git a/docs/scheduler-languages.org b/docs/scheduler-languages.org index cf03b3a..02249f8 100644 --- a/docs/scheduler-languages.org +++ b/docs/scheduler-languages.org @@ -224,8 +224,18 @@ Section DEFINITION. (rs: regset) (**r register state in calling function *) (pr: predset), (**r predicate state of the calling function *) stackframe. +#+end_src + +*** State definition +:PROPERTIES: +:CUSTOM_ID: state +:END: - Inductive state : Type := +Definition of the ~state~ type, which is used by the ~step~ functions. + +#+name: rtlblockinstr-state +#+begin_src coq + Variant state : Type := | State: forall (stack: list stackframe) (**r call stack *) (f: function) (**r current function *) @@ -246,7 +256,10 @@ Section DEFINITION. (v: val) (**r return value for the call *) (m: mem), (**r memory state *) state. +#+end_src +#+name: rtlblockinstr-state +#+begin_src coq End DEFINITION. #+end_src @@ -284,7 +297,15 @@ Section RELSEM. eval_pred (Some p) i i' i | eval_pred_none: forall i i', eval_pred None i i' i. +#+end_src + +*** Step a single instruction +:PROPERTIES: +:CUSTOM_ID: step_instr +:END: +#+name: rtlblockinstr-step_instr +#+begin_src coq Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: forall sp ist, @@ -311,7 +332,15 @@ Section RELSEM. Op.eval_condition c rs##args m = Some b -> eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist -> step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist. +#+end_src + +*** Step a control-flow instruction +:PROPERTIES: +:CUSTOM_ID: step_cf_instr +:END: +#+name: rtlblockinstr-step_cf_instr +#+begin_src coq Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := | exec_RBcall: forall s f sp rs m res fd ros sig args pc pc' pr, @@ -356,7 +385,10 @@ Section RELSEM. forall s f sp pc rs pr m cf1 cf2 st' p t, step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. +#+end_src +#+name: rtlblockinstr-end_RELSEM +#+begin_src coq End RELSEM. #+end_src @@ -394,137 +426,98 @@ Definition fundef := @fundef bb. Definition program := @program bb. Definition funsig := @funsig bb. Definition stackframe := @stackframe bb. +Definition state := @state bb. Definition genv := @genv bb. +#+end_src -Inductive state : Type := -| State: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (b: bb) (**r current block being executed *) - (sp: val) (**r stack pointer *) - (pc: node) (**r current program point in [c] *) - (rs: regset) (**r register state *) - (pr: predset) (**r predicate register state *) - (m: mem), (**r memory state *) - state -| Callstate: - forall (stack: list stackframe) (**r call stack *) - (f: fundef) (**r function to call *) - (args: list val) (**r arguments to the call *) - (m: mem), (**r memory state *) - state -| Returnstate: - forall (stack: list stackframe) (**r call stack *) - (v: val) (**r return value for the call *) - (m: mem), (**r memory state *) - state. +** Semantics + +We first describe the semantics by assuming a global program environment with type ~genv~ which was +declared earlier. +#+name: rtlblock-semantics +#+begin_src coq Section RELSEM. Context (ge: genv). +#+end_src + +*** Instruction list step +:PROPERTIES: +:CUSTOM_ID: step_instr_list +:END: + +The ~step_instr_list~ definition describes the execution of a list of instructions in one big step, +inductively traversing the list of instructions and applying the ~step_instr~ ([[#step_instr][step_instr]]). +#+name: rtlblock-step_instr_list +#+begin_src coq Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := - | exec_RBcons: - forall state i state' state'' instrs sp, - step_instr ge sp state i state' -> - step_instr_list sp state' instrs state'' -> - step_instr_list sp state (i :: instrs) state'' - | exec_RBnil: - forall state sp, - step_instr_list sp state nil state. + | exec_RBcons: + forall state i state' state'' instrs sp, + step_instr ge sp state i state' -> + step_instr_list sp state' instrs state'' -> + step_instr_list sp state (i :: instrs) state'' + | exec_RBnil: + forall state sp, + step_instr_list sp state nil state. +#+end_src - Definition find_function - (ros: reg + ident) (rs: regset) : option fundef := - match ros with - | inl r => Genv.find_funct ge rs#r - | inr symb => - match Genv.find_symbol ge symb with - | None => None - | Some b => Genv.find_funct_ptr ge b - end - end. +*** Top-level step +:PROPERTIES: +:CUSTOM_ID: rtlblock-step +:END: - Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := - | exec_RBcall: - forall s f b sp rs m res fd ros sig args pc pc' pr, - find_function ros rs = Some fd -> - funsig fd = sig -> - step_cf_instr (State s f b sp pc rs pr m) (RBcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) - | exec_RBtailcall: - forall s f b stk rs m sig ros args fd m' pc pr, - find_function ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args) - E0 (Callstate s fd rs##args m') - | exec_RBbuiltin: - forall s f b sp rs m ef args res pc' vargs t vres m' pc pr, - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - step_cf_instr (State s f b sp pc rs pr m) (RBbuiltin ef args res pc') - t (State s f b sp pc' (regmap_setres res vres rs) pr m') - | exec_RBcond: - forall s f block sp rs m cond args ifso ifnot b pc pc' pr, - eval_condition cond rs##args m = Some b -> - pc' = (if b then ifso else ifnot) -> - step_cf_instr (State s f block sp pc rs pr m) (RBcond cond args ifso ifnot) - E0 (State s f block sp pc' rs pr m) - | exec_RBjumptable: - forall s f b sp rs m arg tbl n pc pc' pr, - rs#arg = Vint n -> - list_nth_z tbl (Int.unsigned n) = Some pc' -> - step_cf_instr (State s f b sp pc rs pr m) (RBjumptable arg tbl) - E0 (State s f b sp pc' rs pr m) - | exec_RBreturn: - forall s f b stk rs m or pc m' pr, - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or) - E0 (Returnstate s (regmap_optget or Vundef rs) m') - | exec_RBgoto: - forall s f b sp pc rs pr m pc', - step_cf_instr (State s f b sp pc rs pr m) (RBgoto pc') E0 (State s f b sp pc' rs pr m) - | exec_RBpred_cf: - forall s f b sp pc rs pr m cf1 cf2 st' p t, - step_cf_instr (State s f b sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> - step_cf_instr (State s f b sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. +The step function itself then uses this big step of the list of instructions to then show a +transition from basic block to basic block. - Inductive step: state -> trace -> state -> Prop := +#+name: rtlblock-step +#+begin_src coq + Variant step: state -> trace -> state -> Prop := + | exec_bblock: + forall s f sp pc rs rs' m m' t s' bb pr pr', + f.(fn_code)!pc = Some bb -> + step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> + step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> + step (State s f sp pc rs pr m) t s' | exec_function_internal: forall s f args m m' stk, Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> step (Callstate s (Internal f) args m) - E0 (State s f - (Vptr stk Ptrofs.zero) - f.(fn_entrypoint) - (init_regs args f.(fn_params)) - (PMap.init false) - m') + E0 (State s f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + (PMap.init false) + m') | exec_function_external: forall s ef args res t m m', external_call ef ge args m t res m' -> step (Callstate s (External ef) args m) - t (Returnstate s res m') + t (Returnstate s res m') | exec_return: forall res f sp pc rs s vres m pr, step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) - E0 (State s f sp pc (rs#res <- vres) pr m). + E0 (State s f sp pc (rs#res <- vres) pr m). +#+end_src +#+name: rtlblock-rest +#+begin_src coq End RELSEM. Inductive initial_state (p: program): state -> Prop := - | initial_state_intro: forall b f m0, - let ge := Genv.globalenv p in - Genv.init_mem p = Some m0 -> - Genv.find_symbol ge p.(prog_main) = Some b -> - Genv.find_funct_ptr ge b = Some f -> - funsig f = signature_main -> - initial_state p (Callstate nil f nil m0). +| initial_state_intro: forall b f m0, + let ge := Genv.globalenv p in + Genv.init_mem p = Some m0 -> + Genv.find_symbol ge p.(prog_main) = Some b -> + Genv.find_funct_ptr ge b = Some f -> + funsig f = signature_main -> + initial_state p (Callstate nil f nil m0). Inductive final_state: state -> int -> Prop := - | final_state_intro: forall r m, - final_state (Returnstate nil (Vint r) m) r. +| final_state_intro: forall r m, + final_state (Returnstate nil (Vint r) m) r. Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). diff --git a/docs/scheduler.org b/docs/scheduler.org index 018633c..c8b00ff 100644 --- a/docs/scheduler.org +++ b/docs/scheduler.org @@ -909,18 +909,19 @@ Require Import vericert.hls.Predicate. Require Import vericert.hls.Abstr. Import NE.NonEmptyNotation. -(** * RTLPar Generation *) - #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. +#+end_src -(** ** Abstract Computations +** Abstract Computations Define the abstract computation using the [update] function, which will set each register to its symbolic value. First we need to define a few helper functions to correctly translate the -predicates. *) +predicates. +#+name: rtlpargen-generation +#+begin_src coq Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr := match l with | nil => nil @@ -1014,18 +1015,25 @@ Definition app_predicated {A: Type} (p: option pred_op) (a b: predicated A) := Definition pred_ret {A: Type} (a: A) : predicated A := NE.singleton (T, a). -(** *** Update Function +#+end_src -The [update] function will generate a new forest given an existing forest and a new instruction, -so that it can evaluate a symbolic expression by folding over a list of instructions. The main -problem is that predicates need to be merged as well, so that: +*** Update Function +:PROPERTIES: +:CUSTOM_ID: update-function +:END: + +The [update] function will generate a new forest given an existing forest and a new instruction, so +that it can evaluate a symbolic expression by folding over a list of instructions. The main problem +is that predicates need to be merged as well, so that: 1. The predicates are *independent*. 2. The expression assigned to the register should still be correct. This is done by multiplying the predicates together, and assigning the negation of the expression to -the other predicates. *) +the other predicates. +#+name: rtlpargen-update-function +#+begin_src coq Definition update (f : forest) (i : instr) : forest := match i with | RBnop => f @@ -1055,28 +1063,37 @@ Definition update (f : forest) (i : instr) : forest := (f # (Pred p)) (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) end. +#+end_src -(** Implementing which are necessary to show the correctness of the translation validation by +Implementing which are necessary to show the correctness of the translation validation by showing that there aren't any more effects in the resultant RTLPar code than in the RTLBlock code. -Get a sequence from the basic block. *) +Get a sequence from the basic block. +#+name: rtlpargen-abstract-seq +#+begin_src coq Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := match b with | nil => f | i :: l => abstract_sequence (update f i) l end. +#+end_src -(** Check equivalence of control flow instructions. As none of the basic blocks should have been +Check equivalence of control flow instructions. As none of the basic blocks should have been moved, none of the labels should be different, meaning the control-flow instructions should match -exactly. *) +exactly. +#+name: rtlpargen-check-functions +#+begin_src coq Definition check_control_flow_instr (c1 c2: cf_instr) : bool := if cf_instr_eq c1 c2 then true else false. +#+end_src -(** We define the top-level oracle that will check if two basic blocks are equivalent after a -scheduling transformation. *) +We define the top-level oracle that will check if two basic blocks are equivalent after a +scheduling transformation. +#+name: rtlpargen-top-check-functions +#+begin_src coq Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := match bb with | nil => @@ -1116,8 +1133,12 @@ Lemma check_scheduled_trees_correct2: PTree.get x f2 = None. Proof. solve_scheduled_trees_correct. Qed. -(** ** Top-level Functions *) +#+end_src + +** Top-level Functions +#+name: rtlpargen-top-level-functions +#+begin_src coq Parameter schedule : RTLBlock.function -> RTLPar.function. Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function := @@ -1136,6 +1157,7 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program := transform_partial_program transl_fundef p. #+end_src + * RTLPargenproof :PROPERTIES: :header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v -- cgit