From 3539ac357ae86f8923e98887e7ff9cc5a7a09a94 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 29 Jan 2021 15:41:59 +0000 Subject: Refactoring RTLBlock and RTLPar --- src/hls/RTLBlock.v | 149 ++++++++----------------------------- src/hls/RTLBlockInstr.v | 159 ++++++++++++++++++++++++++++++++------- src/hls/RTLPar.v | 194 +++++++++++------------------------------------- 3 files changed, 205 insertions(+), 297 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index f513779..6a3487a 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -30,132 +30,41 @@ Require Import compcert.verilog.Op. Require Import vericert.hls.RTLBlockInstr. -Definition bblock_body : Type := list instr. -Definition bblock := bblock bblock_body. - -Definition code : Type := PTree.t bblock. - -Record function: Type := mkfunction { - fn_sig: signature; - fn_params: list reg; - fn_stacksize: Z; - fn_code: code; - fn_entrypoint: node -}. - -Definition fundef := AST.fundef function. - -Definition program := AST.program fundef unit. - -Definition funsig (fd: fundef) := - match fd with - | Internal f => fn_sig f - | External ef => ef_sig ef - end. - -Inductive stackframe : Type := -| Stackframe: - forall (res: reg) (**r where to store the result *) - (f: function) (**r calling function *) - (sp: val) (**r stack pointer in calling function *) - (pc: node) (**r program point in calling function *) - (rs: regset), (**r register state in calling function *) - stackframe. - -Inductive state : Type := -| State: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (sp: val) (**r stack pointer *) - (pc: node) (**r current program point in [c] *) - (rs: regset) (**r register state *) - (m: mem), (**r memory state *) - state -| Block: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (sp: val) (**r stack pointer *) - (bb: bblock) (**r Current basic block *) - (rs: regset) (**r 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. - -Definition genv := Genv.t fundef unit. +Definition bb := list instr. + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe bb. +Definition state := @state bb. +Definition genv := @genv bb. Section RELSEM. Context (ge: genv). - 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_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. Inductive step: state -> trace -> state -> Prop := - | exec_bblock_start: - forall stack f sp pc rs m bb, - (fn_code f)!pc = Some bb -> - step (State stack f sp pc rs m) E0 (Block stack f sp bb rs m) - | exec_bblock_instr: - forall stack f sp bb cfi rs m rs' m', - step_instr_list _ ge sp (InstrState rs m) bb (InstrState rs' m') -> - step (Block stack f sp (mk_bblock bb cfi) rs m) - E0 (Block stack f sp (mk_bblock nil cfi) rs' m') - | exec_RBcall: - forall s f sp rs m res fd ros sig args pc', - find_function ros rs = Some fd -> - funsig fd = sig -> - step (Block s f sp (mk_bblock nil (RBcall sig ros args res pc')) rs m) - E0 (Callstate (Stackframe res f sp pc' rs :: s) fd rs##args m) - | exec_RBtailcall: - forall s f stk rs m sig ros args fd m', - find_function ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step (Block s f (Vptr stk Ptrofs.zero) (mk_bblock nil (RBtailcall sig ros args)) rs m) - E0 (Callstate s fd rs##args m') - | exec_RBbuiltin: - forall s f sp rs m ef args res pc' vargs t vres m', - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - step (Block s f sp (mk_bblock nil (RBbuiltin ef args res pc')) rs m) - t (State s f sp pc' (regmap_setres res vres rs) m') - | exec_RBcond: - forall s f sp rs m cond args ifso ifnot b pc', - eval_condition cond rs##args m = Some b -> - pc' = (if b then ifso else ifnot) -> - step (Block s f sp (mk_bblock nil (RBcond cond args ifso ifnot)) rs m) - E0 (State s f sp pc' rs m) - | exec_RBjumptable: - forall s f sp rs m arg tbl n pc', - rs#arg = Vint n -> - list_nth_z tbl (Int.unsigned n) = Some pc' -> - step (Block s f sp (mk_bblock nil (RBjumptable arg tbl)) rs m) - E0 (State s f sp pc' rs m) - | exec_Ireturn: - forall s f stk rs m or m', - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step (Block s f (Vptr stk Ptrofs.zero) (mk_bblock nil (RBreturn or)) rs m) - E0 (Returnstate s (regmap_optget or Vundef rs) m') + | exec_bblock: + forall s f sp pc rs rs' m m' t s' bb, + f.(fn_code)!pc = Some bb -> + step_instr_list sp (InstrState rs m) bb.(bb_body) (InstrState rs' m') -> + step_cf_instr ge (State s f sp pc rs' m') bb.(bb_exit) t s' -> + step (State s f sp pc rs m) t s' | exec_function_internal: - forall s f args m m' stk, + 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 @@ -165,12 +74,12 @@ Section RELSEM. (init_regs args f.(fn_params)) m') | exec_function_external: - forall s ef args res t m m', + 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') | exec_return: - forall res f sp pc rs s vres m, + forall res f sp pc rs s vres m, step (Returnstate (Stackframe res f sp pc rs :: s) vres m) E0 (State s f sp pc (rs#res <- vres) m). diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index a3f058a..2744527 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -18,9 +18,12 @@ Require Import compcert.backend.Registers. Require Import compcert.common.AST. +Require Import compcert.common.Events. Require Import compcert.common.Globalenvs. Require Import compcert.common.Memory. Require Import compcert.common.Values. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. Require Import compcert.verilog.Op. Require Import vericert.common.Vericertlib. @@ -45,14 +48,6 @@ Inductive cf_instr : Type := | RBreturn : option reg -> cf_instr | RBgoto : node -> cf_instr. -Record bblock (bblock_body : Type) : Type := mk_bblock { - bb_body: bblock_body; - bb_exit: cf_instr - }. -Arguments mk_bblock [bblock_body]. -Arguments bb_body [bblock_body]. -Arguments bb_exit [bblock_body]. - Definition successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -102,47 +97,157 @@ Inductive instr_state : Type := (m: mem), instr_state. +Section DEFINITION. + + Context {bblock_body: Type}. + + Record bblock : Type := mk_bblock { + bb_body: bblock_body; + bb_exit: cf_instr + }. + + Definition code: Type := PTree.t bblock. + + Record function: Type := mkfunction { + fn_sig: signature; + fn_params: list reg; + fn_stacksize: Z; + fn_code: code; + fn_entrypoint: node + }. + + Definition fundef := AST.fundef function. + + Definition program := AST.program fundef unit. + + Definition funsig (fd: fundef) := + match fd with + | Internal f => fn_sig f + | External ef => ef_sig ef + end. + + Inductive stackframe : Type := + | Stackframe: + forall (res: reg) (**r where to store the result *) + (f: function) (**r calling function *) + (sp: val) (**r stack pointer in calling function *) + (pc: node) (**r program point in calling function *) + (rs: regset), (**r register state in calling function *) + stackframe. + + Inductive state : Type := + | State: + forall (stack: list stackframe) (**r call stack *) + (f: function) (**r current function *) + (sp: val) (**r stack pointer *) + (pc: node) (**r current program point in [c] *) + (rs: regset) (**r 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. + +End DEFINITION. + Section RELSEM. - Context (fundef: Type). + Context {bblock_body : Type}. + + Definition genv := Genv.t (@fundef bblock_body) unit. - Definition genv := Genv.t fundef unit. + Context (ge: genv). - Context (ge: genv) (sp: val). + 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_instr: instr_state -> instr -> instr_state -> Prop := + Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: - forall rs m, - step_instr (InstrState rs m) RBnop (InstrState rs m) + forall rs m sp, + step_instr sp (InstrState rs m) RBnop (InstrState rs m) | exec_RBop: - forall op v res args rs m, + forall op v res args rs m sp, eval_operation ge sp op rs##args m = Some v -> - step_instr (InstrState rs m) + step_instr sp (InstrState rs m) (RBop op args res) (InstrState (rs#res <- v) m) | exec_RBload: - forall addr rs args a chunk m v dst, + forall addr rs args a chunk m v dst sp, eval_addressing ge sp addr rs##args = Some a -> Mem.loadv chunk m a = Some v -> - step_instr (InstrState rs m) + step_instr sp (InstrState rs m) (RBload chunk addr args dst) (InstrState (rs#dst <- v) m) | exec_RBstore: - forall addr rs args a chunk m src m', + forall addr rs args a chunk m src m' sp, eval_addressing ge sp addr rs##args = Some a -> Mem.storev chunk m a rs#src = Some m' -> - step_instr (InstrState rs m) + step_instr sp (InstrState rs m) (RBstore chunk addr args src) (InstrState rs m'). - Inductive step_instr_list: instr_state -> list instr -> instr_state -> Prop := + Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := | exec_RBcons: - forall state i state' state'' instrs, - step_instr state i state' -> - step_instr_list state' instrs state'' -> - step_instr_list state (i :: instrs) state'' + forall state i state' state'' instrs sp, + step_instr sp state i state' -> + step_instr_list sp state' instrs state'' -> + step_instr_list sp state (i :: instrs) state'' | exec_RBnil: - forall state, - step_instr_list state nil state. + forall state sp, + step_instr_list sp state nil state. + + 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', + find_function ros rs = Some fd -> + funsig fd = sig -> + step_cf_instr (State s f sp pc rs m) (RBcall sig ros args res pc') + E0 (Callstate (Stackframe res f sp pc' rs :: s) fd rs##args m) + | exec_RBtailcall: + forall s f stk rs m sig ros args fd m' pc, + 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 (Vptr stk Ptrofs.zero) pc rs m) (RBtailcall sig ros args) + E0 (Callstate s fd rs##args m') + | exec_RBbuiltin: + forall s f sp rs m ef args res pc' vargs t vres m' pc, + 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 sp pc rs m) (RBbuiltin ef args res pc') + t (State s f sp pc' (regmap_setres res vres rs) m') + | exec_RBcond: + forall s f sp rs m cond args ifso ifnot b pc pc', + eval_condition cond rs##args m = Some b -> + pc' = (if b then ifso else ifnot) -> + step_cf_instr (State s f sp pc rs m) (RBcond cond args ifso ifnot) + E0 (State s f sp pc' rs m) + | exec_RBjumptable: + forall s f sp rs m arg tbl n pc pc', + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + step_cf_instr (State s f sp pc rs m) (RBjumptable arg tbl) + E0 (State s f sp pc' rs m) + | exec_Ireturn: + forall s f stk rs m or pc m', + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs m) (RBreturn or) + E0 (Returnstate s (regmap_optget or Vundef rs) m'). End RELSEM. diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index 1e5dd43..3ad54f5 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -28,165 +28,44 @@ Require Import compcert.lib.Integers. Require Import compcert.lib.Maps. Require Import compcert.verilog.Op. -Require Import vericert.common.Vericertlib. Require Import vericert.hls.RTLBlockInstr. -Definition bblock_body := list (list instr). -Definition bblock := RTLBlockInstr.bblock bblock_body. +Definition bb := list (list instr). -Definition is_empty (b: bblock_body) : Prop := - Forall (fun n => Forall (fun m => m = RBnop) n) b. - -Definition code : Type := PTree.t bblock. - -Record function: Type := mkfunction { - fn_sig: signature; - fn_params: list reg; - fn_stacksize: Z; - fn_code: code; - fn_entrypoint: node -}. - -Definition fundef := AST.fundef function. - -Definition program := AST.program fundef unit. - -Definition funsig (fd: fundef) := - match fd with - | Internal f => fn_sig f - | External ef => ef_sig ef - end. - -Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) := - let max_body := fold_left (fun x l => fold_left max_reg_instr l x) bb.(bb_body) m in - max_reg_cfi max_body bb.(bb_exit). - -Definition max_reg_function (f: function) := - Pos.max - (PTree.fold max_reg_bblock f.(fn_code) 1%positive) - (fold_left Pos.max f.(fn_params) 1%positive). - -Definition max_pc_function (f: function) : positive := - PTree.fold (fun m pc i => (Pos.max m - (pc + match Zlength i.(bb_body) - with Z.pos p => p | _ => 1 end))%positive) - f.(fn_code) 1%positive. - -Definition genv := Genv.t fundef unit. - -Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := - match rl, vl with - | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs) - | _, _ => Regmap.init Vundef - end. - -Inductive stackframe : Type := - | Stackframe: - forall (res: reg) (**r where to store the result *) - (f: function) (**r calling function *) - (sp: val) (**r stack pointer in calling function *) - (pc: node) (**r program point in calling function *) - (rs: regset), (**r register state in calling function *) - stackframe. - -Inductive state : Type := - | State: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (sp: val) (**r stack pointer *) - (pc: node) (**r current program point in [c] *) - (rs: regset) (**r register state *) - (m: mem), (**r memory state *) - state - | Block: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (sp: val) (**r stack pointer *) - (bb: bblock) (**r current basic block *) - (rs: regset) (**r 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. +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe bb. +Definition state := @state bb. +Definition genv := @genv bb. Section RELSEM. Context (ge: genv). - 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_instr_block (sp : val) + : instr_state -> bb -> instr_state -> Prop := + | exec_instr_block_cons: + forall state i state' state'' instrs, + step_instr_list ge sp state i state' -> + step_instr_block sp state' instrs state'' -> + step_instr_block sp state (i :: instrs) state'' + | exec_instr_block_nil: + forall state, + step_instr_block sp state nil state. Inductive step: state -> trace -> state -> Prop := - | exec_bblock_start: - forall stack f sp pc rs m bb, - (fn_code f)!pc = Some bb -> - step (State stack f sp pc rs m) E0 (Block stack f sp bb rs m) - | exec_bblock_instr_cons: - forall sp rs m rs' m' stack f bb bbs cfi, - step_instr_list _ ge sp (InstrState rs m) bb (InstrState rs' m') -> - step (Block stack f sp (mk_bblock (bb :: bbs) cfi) rs m) E0 - (Block stack f sp (mk_bblock bbs cfi) rs' m') - | exec_RBcall: - forall s f sp rs m res fd ros sig args pc' el, - find_function ros rs = Some fd -> - funsig fd = sig -> - is_empty el -> - step (Block s f sp (mk_bblock el (RBcall sig ros args res pc')) rs m) - E0 (Callstate (Stackframe res f sp pc' rs :: s) fd rs##args m) - | exec_RBtailcall: - forall s f stk rs m sig ros args fd m' el, - find_function ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - is_empty el -> - step (Block s f (Vptr stk Ptrofs.zero) (mk_bblock el (RBtailcall sig ros args)) rs m) - E0 (Callstate s fd rs##args m') - | exec_RBbuiltin: - forall s f sp rs m ef args res pc' vargs t vres m' el, - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - is_empty el -> - step (Block s f sp (mk_bblock el (RBbuiltin ef args res pc')) rs m) - t (State s f sp pc' (regmap_setres res vres rs) m') - | exec_RBcond: - forall s f sp rs m cond args ifso ifnot b pc' el, - eval_condition cond rs##args m = Some b -> - pc' = (if b then ifso else ifnot) -> - is_empty el -> - step (Block s f sp (mk_bblock el (RBcond cond args ifso ifnot)) rs m) - E0 (State s f sp pc' rs m) - | exec_RBjumptable: - forall s f sp rs m arg tbl n pc' el, - rs#arg = Vint n -> - list_nth_z tbl (Int.unsigned n) = Some pc' -> - is_empty el -> - step (Block s f sp (mk_bblock el (RBjumptable arg tbl)) rs m) - E0 (State s f sp pc' rs m) - | exec_Ireturn: - forall s f stk rs m or m' el, - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - is_empty el -> - step (Block s f (Vptr stk Ptrofs.zero) (mk_bblock el (RBreturn or)) rs m) - E0 (Returnstate s (regmap_optget or Vundef rs) m') + | exec_bblock: + forall s f sp pc rs rs' m m' t s' bb, + f.(fn_code)!pc = Some bb -> + step_instr_block sp (InstrState rs m) bb.(bb_body) (InstrState rs' m') -> + step_cf_instr ge (State s f sp pc rs' m') bb.(bb_exit) t s' -> + step (State s f sp pc rs m) t s' | exec_function_internal: - forall s f args m m' stk, + 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 @@ -196,12 +75,12 @@ Section RELSEM. (init_regs args f.(fn_params)) m') | exec_function_external: - forall s ef args res t m m', + 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') | exec_return: - forall res f sp pc rs s vres m, + forall res f sp pc rs s vres m, step (Returnstate (Stackframe res f sp pc rs :: s) vres m) E0 (State s f sp pc (rs#res <- vres) m). @@ -222,3 +101,18 @@ Inductive final_state: state -> int -> Prop := Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). + +Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) := + let max_body := fold_left (fun x l => fold_left max_reg_instr l x) bb.(bb_body) m in + max_reg_cfi max_body bb.(bb_exit). + +Definition max_reg_function (f: function) := + Pos.max + (PTree.fold max_reg_bblock f.(fn_code) 1%positive) + (fold_left Pos.max f.(fn_params) 1%positive). + +Definition max_pc_function (f: function) : positive := + PTree.fold (fun m pc i => (Pos.max m + (pc + match Zlength i.(bb_body) + with Z.pos p => p | _ => 1 end))%positive) + f.(fn_code) 1%positive. -- cgit