From 9f252d9055ad16f9433caaf41f6490e45424e88a Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 19 May 2021 15:53:02 +0200 Subject: Adding a BTL record to help oracles --- scheduling/BTL.v | 143 ++++++++++++++++++++++++++++--------------------------- 1 file changed, 74 insertions(+), 69 deletions(-) (limited to 'scheduling/BTL.v') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 3daa40c4..44c25b75 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -23,24 +23,29 @@ Require Import RTL Op Registers OptionMonad. Definition exit := node. (* we may generalize this with register renamings at exit, like in "phi" nodes of SSA-form *) +(* node_info is a ghost record to provide and transfert information through oracles *) +Record node_info := { + _pos: node +}. + (** final instructions (that stops block execution) *) Inductive final: Type := | Bgoto (succ:exit) (** No operation -- just branch to [succ]. *) - | Breturn (res: option reg) + | Breturn (res: option reg) (ni: node_info) (** terminates the execution of the current function. It returns the value of the given register, or [Vundef] if none is given. *) - | Bcall (sig: signature) (fn: reg + ident) (args:list reg) (dest:reg) (succ:exit) + | Bcall (sig: signature) (fn: reg + ident) (args:list reg) (dest:reg) (succ:exit) (ni: node_info) (** invokes the function determined by [fn] (either a function pointer found in a register or a function name), giving it the values of registers [args] as arguments. It stores the return value in [dest] and branches to [succ]. *) - | Btailcall (sig:signature) (fn: reg + ident) (args: list reg) + | Btailcall (sig:signature) (fn: reg + ident) (args: list reg) (ni: node_info) (** performs a function invocation in tail-call position (the current function terminates after the call, returning the result of the callee) *) - | Bbuiltin (ef:external_function) (args:list (builtin_arg reg)) (dest:builtin_res reg) (succ:exit) + | Bbuiltin (ef:external_function) (args:list (builtin_arg reg)) (dest:builtin_res reg) (succ:exit) (ni: node_info) (** calls the built-in function identified by [ef], giving it the values of [args] as arguments. It stores the return value in [dest] and branches to [succ]. *) - | Bjumptable (arg:reg) (tbl:list exit) + | Bjumptable (arg:reg) (tbl:list exit) (ni: node_info) (** [Bjumptable arg tbl] transitions to the node that is the [n]-th element of the list [tbl], where [n] is the unsigned integer value of register [arg]. *) . @@ -50,14 +55,14 @@ Inductive iblock: Type := (* final instructions that stops block execution *) | BF (fi: final) (* basic instructions that continues block execution, except when aborting *) - | Bnop (* nop instruction *) - | Bop (op:operation) (args:list reg) (dest:reg) + | Bnop (ni: node_info) (* nop instruction *) + | Bop (op:operation) (args:list reg) (dest:reg) (ni: node_info) (** performs the arithmetic operation [op] over the values of registers [args], stores the result in [dest] *) - | Bload (trap:trapping_mode) (chunk:memory_chunk) (addr:addressing) (args:list reg) (dest:reg) + | Bload (trap:trapping_mode) (chunk:memory_chunk) (addr:addressing) (args:list reg) (dest:reg) (ni: node_info) (** loads a [chunk] quantity from the address determined by the addressing mode [addr] and the values of the [args] registers, stores the quantity just read into [dest]. If trap=NOTRAP, then failures lead to a default value written to [dest]. *) - | Bstore (chunk:memory_chunk) (addr:addressing) (args:list reg) (src:reg) + | Bstore (chunk:memory_chunk) (addr:addressing) (args:list reg) (src:reg) (ni: node_info) (** stores the value of register [src] in the [chunk] quantity at the the address determined by the addressing mode [addr] and the values of the [args] registers. *) @@ -65,7 +70,7 @@ Inductive iblock: Type := | Bseq (b1 b2: iblock) (** starts by running [b1] and stops here if execution of [b1] has reached a final instruction or aborted or continue with [b2] otherwise *) - | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (info:option bool) + | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (info:option bool) (ni: node_info) (** evaluates the boolean condition [cond] over the values of registers [args]. If the condition is true, it continues on [ifso]. If the condition is false, it continues on [ifnot]. @@ -195,23 +200,23 @@ Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop := (** internal big-step execution of one iblock *) Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option final -> Prop := | exec_final rs m fin: iblock_istep sp rs m (BF fin) rs m (Some fin) - | exec_nop rs m: iblock_istep sp rs m Bnop rs m None - | exec_op rs m op args res v + | exec_nop rs m ni: iblock_istep sp rs m (Bnop ni) rs m None + | exec_op rs m op args res v ni (EVAL: eval_operation ge sp op rs##args m = Some v) - : iblock_istep sp rs m (Bop op args res) (rs#res <- v) m None - | exec_load_TRAP rs m chunk addr args dst a v + : iblock_istep sp rs m (Bop op args res ni) (rs#res <- v) m None + | exec_load_TRAP rs m chunk addr args dst a v ni (EVAL: eval_addressing ge sp addr rs##args = Some a) (LOAD: Mem.loadv chunk m a = Some v) - : iblock_istep sp rs m (Bload TRAP chunk addr args dst) (rs#dst <- v) m None + : iblock_istep sp rs m (Bload TRAP chunk addr args dst ni) (rs#dst <- v) m None (* TODO: replace [exec_load_TRAP] above by this one. See "new scheme" for "NOTRAP" above | exec_load rs m trap chunk addr args dst v (LOAD: has_loaded sp rs m chunk addr args v trap) : iblock_istep sp rs m (Bload trap chunk addr args dst) (rs#dst <- v) m None *) - | exec_store rs m chunk addr args src a m' + | exec_store rs m chunk addr args src a m' ni (EVAL: eval_addressing ge sp addr rs##args = Some a) (STORE: Mem.storev chunk m a rs#src = Some m') - : iblock_istep sp rs m (Bstore chunk addr args src) rs m' None + : iblock_istep sp rs m (Bstore chunk addr args src ni) rs m' None | exec_seq_stop rs m b1 b2 rs' m' fin (EXEC: iblock_istep sp rs m b1 rs' m' (Some fin)) : iblock_istep sp rs m (Bseq b1 b2) rs' m' (Some fin) @@ -219,10 +224,10 @@ Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option fi (EXEC1: iblock_istep sp rs m b1 rs1 m1 None) (EXEC2: iblock_istep sp rs1 m1 b2 rs' m' ofin) : iblock_istep sp rs m (Bseq b1 b2) rs' m' ofin - | exec_cond rs m cond args ifso ifnot i b rs' m' ofin + | exec_cond rs m cond args ifso ifnot i b rs' m' ofin ni (EVAL: eval_condition cond rs##args m = Some b) (EXEC: iblock_istep sp rs m (if b then ifso else ifnot) rs' m' ofin) - : iblock_istep sp rs m (Bcond cond args ifso ifnot i) rs' m' ofin + : iblock_istep sp rs m (Bcond cond args ifso ifnot i ni) rs' m' ofin . Local Hint Constructors iblock_istep: core. @@ -234,18 +239,18 @@ Fixpoint iblock_istep_run sp ib rs m: option outcome := | BF fin => Some {| _rs := rs; _m := m; _fin := Some fin |} (* basic instructions *) - | Bnop => + | Bnop _ => Some {| _rs := rs; _m:= m; _fin := None |} - | Bop op args res => + | Bop op args res _ => SOME v <- eval_operation ge sp op rs##args m IN Some {| _rs := rs#res <- v; _m:= m; _fin := None |} - | Bload TRAP chunk addr args dst => + | Bload TRAP chunk addr args dst _ => SOME a <- eval_addressing ge sp addr rs##args IN SOME v <- Mem.loadv chunk m a IN Some {| _rs := rs#dst <- v; _m:= m; _fin := None |} - | Bload NOTRAP chunk addr args dst => + | Bload NOTRAP chunk addr args dst _ => None (* TODO *) - | Bstore chunk addr args src => + | Bstore chunk addr args src _ => SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN Some {| _rs := rs; _m:= m'; _fin := None |} @@ -256,7 +261,7 @@ Fixpoint iblock_istep_run sp ib rs m: option outcome := | None => iblock_istep_run sp b2 out1.(_rs) out1.(_m) | _ => Some out1 (* stop execution on the 1st final instruction *) end - | Bcond cond args ifso ifnot _ => + | Bcond cond args ifso ifnot _ _ => SOME b <- eval_condition cond rs##args m IN iblock_istep_run sp (if b then ifso else ifnot) rs m end. @@ -275,32 +280,32 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := | exec_Bgoto pc: final_step stack f sp rs m (Bgoto pc) E0 (State stack f sp pc rs m) - | exec_Breturn or stk m': + | exec_Breturn or stk m' ni: sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - final_step stack f sp rs m (Breturn or) + final_step stack f sp rs m (Breturn or ni) E0 (Returnstate stack (regmap_optget or Vundef rs) m') - | exec_Bcall sig ros args res pc' fd: + | exec_Bcall sig ros args res pc' fd ni: find_function ros rs = Some fd -> funsig fd = sig -> - final_step stack f sp rs m (Bcall sig ros args res pc') + final_step stack f sp rs m (Bcall sig ros args res pc' ni) E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) - | exec_Btailcall sig ros args stk m' fd: + | exec_Btailcall sig ros args stk m' fd ni: find_function ros rs = Some fd -> funsig fd = sig -> sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - final_step stack f sp rs m (Btailcall sig ros args) + final_step stack f sp rs m (Btailcall sig ros args ni) E0 (Callstate stack fd rs##args m') - | exec_Bbuiltin ef args res pc' vargs t vres m': + | exec_Bbuiltin ef args res pc' vargs t vres m' ni: eval_builtin_args ge (fun r => rs#r) sp m args vargs -> external_call ef ge vargs m t vres m' -> - final_step stack f sp rs m (Bbuiltin ef args res pc') + final_step stack f sp rs m (Bbuiltin ef args res pc' ni) t (State stack f sp pc' (regmap_setres res vres rs) m') - | exec_Bjumptable arg tbl n pc': + | exec_Bjumptable arg tbl n pc' ni: rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - final_step stack f sp rs m (Bjumptable arg tbl) + final_step stack f sp rs m (Bjumptable arg tbl ni) E0 (State stack f sp pc' rs m) . @@ -388,16 +393,16 @@ Inductive match_final_inst (dupmap: PTree.t node): final -> instruction -> Prop | mfi_goto pc pc': dupmap!pc = (Some pc') -> match_final_inst dupmap (Bgoto pc) (Inop pc') *) - | mfi_return or: match_final_inst dupmap (Breturn or) (Ireturn or) - | mfi_call pc pc' s ri lr r: - dupmap!pc = (Some pc') -> match_final_inst dupmap (Bcall s ri lr r pc) (Icall s ri lr r pc') - | mfi_tailcall s ri lr: - match_final_inst dupmap (Btailcall s ri lr) (Itailcall s ri lr) - | mfi_builtin pc pc' ef la br: - dupmap!pc = (Some pc') -> match_final_inst dupmap (Bbuiltin ef la br pc) (Ibuiltin ef la br pc') - | mfi_jumptable ln ln' r: + | mfi_return or ni: match_final_inst dupmap (Breturn or ni) (Ireturn or) + | mfi_call pc pc' s ri lr r ni: + dupmap!pc = (Some pc') -> match_final_inst dupmap (Bcall s ri lr r pc ni) (Icall s ri lr r pc') + | mfi_tailcall s ri lr ni: + match_final_inst dupmap (Btailcall s ri lr ni) (Itailcall s ri lr) + | mfi_builtin pc pc' ef la br ni: + dupmap!pc = (Some pc') -> match_final_inst dupmap (Bbuiltin ef la br pc ni) (Ibuiltin ef la br pc') + | mfi_jumptable ln ln' r ni: list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' -> - match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln') + match_final_inst dupmap (Bjumptable r ln ni) (Ijumptable r ln') . Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := @@ -416,18 +421,18 @@ Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> i cfg!pc = Some i -> match_final_inst dupmap fi i -> match_iblock dupmap cfg isfst pc (BF fi) None - | mib_nop isfst pc pc': + | mib_nop isfst pc pc' ni: cfg!pc = Some (Inop pc') -> - match_iblock dupmap cfg isfst pc Bnop (Some pc') - | mib_op isfst pc pc' op lr r: + match_iblock dupmap cfg isfst pc (Bnop ni) (Some pc') + | mib_op isfst pc pc' op lr r ni: cfg!pc = Some (Iop op lr r pc') -> - match_iblock dupmap cfg isfst pc (Bop op lr r) (Some pc') - | mib_load isfst pc pc' m a lr r: + match_iblock dupmap cfg isfst pc (Bop op lr r ni) (Some pc') + | mib_load isfst pc pc' m a lr r ni: cfg!pc = Some (Iload TRAP m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r) (Some pc') - | mib_store isfst pc pc' m a lr r: + match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r ni) (Some pc') + | mib_store isfst pc pc' m a lr r ni: cfg!pc = Some (Istore m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bstore m a lr r) (Some pc') + match_iblock dupmap cfg isfst pc (Bstore m a lr r ni) (Some pc') | mib_exit pc pc': dupmap!pc = (Some pc') -> match_iblock dupmap cfg false pc' (Bgoto pc) None @@ -445,12 +450,12 @@ Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> i And the verifier could fail if there is such dead code! *) *) - | mib_cond isfst c lr bso bnot pcso pcnot pc opc1 opc2 opc i i': + | mib_cond isfst c lr bso bnot pcso pcnot pc opc1 opc2 opc i i' ni: cfg!pc = Some (Icond c lr pcso pcnot i') -> match_iblock dupmap cfg false pcso bso opc1 -> match_iblock dupmap cfg false pcnot bnot opc2 -> is_join_opt opc1 opc2 opc -> - match_iblock dupmap cfg isfst pc (Bcond c lr bso bnot i) opc + match_iblock dupmap cfg isfst pc (Bcond c lr bso bnot i ni) opc . Definition match_cfg dupmap (cfg: code) (cfg':RTL.code): Prop := @@ -530,14 +535,14 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if negb isfst then OK None else Error (msg "BTL.verify_block: isfst is true Bgoto") - | Breturn or => + | Breturn or _ => match cfg!pc with | Some (Ireturn or') => if option_eq Pos.eq_dec or or' then OK None else Error (msg "BTL.verify_block: different opt reg in Breturn") | _ => Error (msg "BTL.verify_block: incorrect cfg Breturn") end - | Bcall s ri lr r pc1 => + | Bcall s ri lr r pc1 _ => match cfg!pc with | Some (Icall s' ri' lr' r' pc2) => do u <- verify_is_copy dupmap pc1 pc2; @@ -551,7 +556,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different signatures in Bcall") | _ => Error (msg "BTL.verify_block: incorrect cfg Bcall") end - | Btailcall s ri lr => + | Btailcall s ri lr _ => match cfg!pc with | Some (Itailcall s' ri' lr') => if (signature_eq s s') then @@ -562,7 +567,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different signatures in Btailcall") | _ => Error (msg "BTL.verify_block: incorrect cfg Btailcall") end - | Bbuiltin ef la br pc1 => + | Bbuiltin ef la br pc1 _ => match cfg!pc with | Some (Ibuiltin ef' la' br' pc2) => do u <- verify_is_copy dupmap pc1 pc2; @@ -574,7 +579,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different ef in Bbuiltin") | _ => Error (msg "BTL.verify_block: incorrect cfg Bbuiltin") end - | Bjumptable r ln => + | Bjumptable r ln _ => match cfg!pc with | Some (Ijumptable r' ln') => do u <- verify_is_copy_list dupmap ln ln'; @@ -583,12 +588,12 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) | _ => Error (msg "BTL.verify_block: incorrect cfg Bjumptable") end end - | Bnop => + | Bnop _ => match cfg!pc with | Some (Inop pc') => OK (Some pc') | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") end - | Bop op lr r => + | Bop op lr r _ => match cfg!pc with | Some (Iop op' lr' r' pc') => if (eq_operation op op') then @@ -600,7 +605,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different operations in Bop") | _ => Error (msg "BTL.verify_block: incorrect cfg Bop") end - | Bload tm m a lr r => + | Bload tm m a lr r _ => match cfg!pc with | Some (Iload tm' m' a' lr' r' pc') => if (trapping_mode_eq tm TRAP && trapping_mode_eq tm' TRAP) then @@ -616,7 +621,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: NOTRAP trapping_mode unsupported in Bload") | _ => Error (msg "BTL.verify_block: incorrect cfg Bload") end - | Bstore m a lr r => + | Bstore m a lr r _ => match cfg!pc with | Some (Istore m' a' lr' r' pc') => if (chunk_eq m m') then @@ -636,7 +641,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) verify_block dupmap cfg false pc' b2 | None => Error (msg "BTL.verify_block: None next pc in Bseq (deadcode)") end - | Bcond c lr bso bnot i => + | Bcond c lr bso bnot i _ => match cfg!pc with | Some (Icond c' lr' pcso pcnot i') => if (list_eq_dec Pos.eq_dec lr lr') then @@ -794,7 +799,7 @@ Definition is_goto (ib: iblock): bool := Definition is_atom (ib: iblock): bool := match ib with - | Bseq _ _ | Bcond _ _ _ _ _ => false + | Bseq _ _ | Bcond _ _ _ _ _ _ => false | _ => true end. @@ -804,10 +809,10 @@ Inductive is_expand: iblock -> Prop := is_atom ib1 = true -> is_expand ib2 -> is_expand (Bseq ib1 ib2) - | exp_Bcond cond args ib1 ib2 i: + | exp_Bcond cond args ib1 ib2 i ni: is_expand ib1 -> is_expand ib2 -> - is_expand (Bcond cond args ib1 ib2 i) + is_expand (Bcond cond args ib1 ib2 i ni) | exp_others ib: is_atom ib = true -> is_expand ib @@ -817,8 +822,8 @@ Local Hint Constructors is_expand: core. Fixpoint expand (ib: iblock) (k: option iblock): iblock := match ib with | Bseq ib1 ib2 => expand ib1 (Some (expand ib2 k)) - | Bcond cond args ib1 ib2 i => - Bcond cond args (expand ib1 k) (expand ib2 k) i + | Bcond cond args ib1 ib2 i ni => + Bcond cond args (expand ib1 k) (expand ib2 k) i ni | BF fin => fin | ib => match k with -- cgit