aboutsummaryrefslogtreecommitdiffstats
path: root/src
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2022-03-24 11:46:15 +0000
committerYann Herklotz <git@yannherklotz.com>2022-03-24 11:46:15 +0000
commit30baa719fb15c45b13cb869056e51ec7446c0207 (patch)
tree35be4ec314fbcdd27adeba367d3bbf5fc844cdd4 /src
parentb9e793eef135bd411c9945f0b1ba99308d9edbd5 (diff)
downloadvericert-30baa719fb15c45b13cb869056e51ec7446c0207.tar.gz
vericert-30baa719fb15c45b13cb869056e51ec7446c0207.zip
Add more documentation
Diffstat (limited to 'src')
-rw-r--r--src/hls/RTLBlock.v150
-rw-r--r--src/hls/RTLBlockInstr.v16
-rw-r--r--src/hls/RTLPargen.v50
-rw-r--r--src/hls/RTLPargenproof.v4
-rw-r--r--src/hls/Schedule.ml2
5 files changed, 73 insertions, 149 deletions
diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v
index 50bff90..98e032b 100644
--- a/src/hls/RTLBlock.v
+++ b/src/hls/RTLBlock.v
@@ -40,138 +40,74 @@ Definition fundef := @fundef bb.
Definition program := @program bb.
Definition funsig := @funsig bb.
Definition stackframe := @stackframe bb.
+Definition state := @state bb.
Definition genv := @genv bb.
+(* rtlblock-main ends here *)
-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.
-
+(* [[file:../../docs/scheduler-languages.org::rtlblock-semantics][rtlblock-semantics]] *)
Section RELSEM.
Context (ge: genv).
+(* rtlblock-semantics ends here *)
+(* [[file:../../docs/scheduler-languages.org::#step_instr_list][rtlblock-step_instr_list]] *)
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.
-
- 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.
-
- 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'.
-
- Inductive step: state -> trace -> 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.
+(* rtlblock-step_instr_list ends here *)
+
+(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-step]] *)
+ 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).
+(* rtlblock-step ends here *)
+(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-rest]] *)
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).
-(* rtlblock-main ends here *)
+(* rtlblock-rest ends here *)
diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v
index 801a5ea..bd40516 100644
--- a/src/hls/RTLBlockInstr.v
+++ b/src/hls/RTLBlockInstr.v
@@ -195,8 +195,10 @@ Section DEFINITION.
(rs: regset) (**r register state in calling function *)
(pr: predset), (**r predicate state of the calling function *)
stackframe.
+(* rtlblockinstr-type-def ends here *)
- Inductive state : Type :=
+(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *)
+ Variant state : Type :=
| State:
forall (stack: list stackframe) (**r call stack *)
(f: function) (**r current function *)
@@ -217,9 +219,11 @@ Section DEFINITION.
(v: val) (**r return value for the call *)
(m: mem), (**r memory state *)
state.
+(* rtlblockinstr-state ends here *)
+(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *)
End DEFINITION.
-(* rtlblockinstr-type-def ends here *)
+(* rtlblockinstr-state ends here *)
(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *)
Section RELSEM.
@@ -252,7 +256,9 @@ Section RELSEM.
eval_pred (Some p) i i' i
| eval_pred_none:
forall i i', eval_pred None i i' i.
+(* rtlblockinstr-semantics ends here *)
+(* [[file:../../docs/scheduler-languages.org::#step_instr][rtlblockinstr-step_instr]] *)
Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop :=
| exec_RBnop:
forall sp ist,
@@ -279,7 +285,9 @@ 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.
+(* rtlblockinstr-step_instr ends here *)
+(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-step_cf_instr]] *)
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,
@@ -324,6 +332,8 @@ 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'.
+(* rtlblockinstr-step_cf_instr ends here *)
+(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-end_RELSEM]] *)
End RELSEM.
-(* rtlblockinstr-semantics ends here *)
+(* rtlblockinstr-end_RELSEM ends here *)
diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v
index ac9ac2c..31ea51f 100644
--- a/src/hls/RTLPargen.v
+++ b/src/hls/RTLPargen.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-(* [[file:../../lit/scheduler.org::rtlpargen-main][rtlpargen-main]] *)
+(* [[file:../../docs/scheduler.org::rtlpargen-main][rtlpargen-main]] *)
Require Import compcert.backend.Registers.
Require Import compcert.common.AST.
Require Import compcert.common.Globalenvs.
@@ -35,18 +35,12 @@ 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.
+(* rtlpargen-main ends here *)
-(** ** 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. *)
-
+(* [[file:../../docs/scheduler.org::rtlpargen-generation][rtlpargen-generation]] *)
Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr :=
match l with
| nil => nil
@@ -139,19 +133,9 @@ 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).
+(* rtlpargen-generation ends here *)
-(** *** Update Function
-
-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. *)
-
+(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-update-function]] *)
Definition update (f : forest) (i : instr) : forest :=
match i with
| RBnop => f
@@ -181,28 +165,22 @@ Definition update (f : forest) (i : instr) : forest :=
(f # (Pred p))
(map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f))))
end.
+(* rtlpargen-update-function ends here *)
-(** 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. *)
-
+(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-abstract-seq]] *)
Fixpoint abstract_sequence (f : forest) (b : list instr) : forest :=
match b with
| nil => f
| i :: l => abstract_sequence (update f i) l
end.
+(* rtlpargen-abstract-seq ends here *)
-(** 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. *)
-
+(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-check-functions]] *)
Definition check_control_flow_instr (c1 c2: cf_instr) : bool :=
if cf_instr_eq c1 c2 then true else false.
+(* rtlpargen-check-functions ends here *)
-(** We define the top-level oracle that will check if two basic blocks are equivalent after a
-scheduling transformation. *)
-
+(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-top-check-functions]] *)
Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool :=
match bb with
| nil =>
@@ -241,9 +219,9 @@ Lemma check_scheduled_trees_correct2:
PTree.get x f1 = None ->
PTree.get x f2 = None.
Proof. solve_scheduled_trees_correct. Qed.
+(* rtlpargen-top-check-functions ends here *)
-(** ** Top-level Functions *)
-
+(* [[file:../../docs/scheduler.org::rtlpargen-top-level-functions][rtlpargen-top-level-functions]] *)
Parameter schedule : RTLBlock.function -> RTLPar.function.
Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function :=
@@ -261,4 +239,4 @@ Definition transl_fundef := transf_partial_fundef transl_function.
Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program :=
transform_partial_program transl_fundef p.
-(* rtlpargen-main ends here *)
+(* rtlpargen-top-level-functions ends here *)
diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v
index 5188709..6d61572 100644
--- a/src/hls/RTLPargenproof.v
+++ b/src/hls/RTLPargenproof.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-(* [[file:../../lit/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *)
+(* [[file:../../docs/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *)
Require Import compcert.backend.Registers.
Require Import compcert.common.AST.
Require Import compcert.common.Errors.
@@ -39,7 +39,7 @@ Require Import vericert.hls.Abstr.
#[local] Open Scope pred_op.
(* rtlpargenproof-imports ends here *)
-(* [[file:../../lit/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *)
+(* [[file:../../docs/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *)
(*Definition is_regs i := match i with mk_instr_state rs _ => rs end.
Definition is_mem i := match i with mk_instr_state _ m => m end.
diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml
index 36dbc2b..3a5351e 100644
--- a/src/hls/Schedule.ml
+++ b/src/hls/Schedule.ml
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-(* [[file:../../lit/scheduler.org::scheduler-main][scheduler-main]] *)
+(* [[file:../../docs/scheduler.org::scheduler-main][scheduler-main]] *)
open Printf
open Clflags
open Camlcoq