From af2c82cfd7f2318fcd81da2ea4cf3fd695db3b40 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 28 Apr 2021 18:14:53 +0200 Subject: start the new "BTL" IR. --- scheduling/BTL.v | 575 +++++++++++++++++++++++++++++++++++++++++++++ scheduling/BTLtoRTL.v | 23 ++ scheduling/BTLtoRTLproof.v | 127 ++++++++++ 3 files changed, 725 insertions(+) create mode 100644 scheduling/BTL.v create mode 100644 scheduling/BTLtoRTL.v create mode 100644 scheduling/BTLtoRTLproof.v (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v new file mode 100644 index 00000000..8bdafb89 --- /dev/null +++ b/scheduling/BTL.v @@ -0,0 +1,575 @@ +(** The BTL intermediate language: abstract syntax and semantics. + + BTL stands for "Block Transfer Language". + + Informally, a block is a piece of "loop-free" code, with a single entry-point, + hence, such that transformation preserving locally the semantics of each block, + preserve also globally the semantics of the function. + + a BTL function is a CFG where each node is such a block, represented by structured code. + + BTL gives a structured view of RTL code. + It is dedicated to optimizations validated by "symbolic simulation" over blocks. + + +*) + +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad. + +(** * Abstract syntax *) + +Definition exit := node. (* we may generalize this with register renamings at exit, + like in "phi" nodes of SSA-form *) + +(** final instructions (that stops block execution) *) +Inductive final: Type := + | Bgoto (succ:exit) (** No operation -- just branch to [succ]. *) + | Breturn (res: option reg) + (** 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) + (** 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) + (** 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) + (** 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 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]. *) + . + +(* instruction block *) +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) + (** 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) + (** 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) + (** 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. *) +(* composed instructions *) + | 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) + (** 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]. + [info] is a ghost field there to provide information relative to branch prediction. *) + . +Coercion BF: final >-> iblock. + + +(** NB: - a RTL [(Inop pc)] ending a branch of block is encoded by [(Bseq Bnop (Bgoto pc))]. + - a RTL [(Inop pc)] in the middle of a branch is simply encoded by [Bnop]. + - the same trick appears for all "basic" instructions and [Icond]. +*) + + + +Record iblock_info := { + entry: iblock; + input_regs: Regset.t (* extra liveness information ignored by BTL semantics *) +}. + +Definition code: Type := PTree.t iblock_info. + +Record function: Type := mkfunction { + fn_sig: signature; + fn_params: list reg; + fn_stacksize: Z; + fn_code: code; + fn_entrypoint: node +}. + +(** A function description comprises a control-flow graph (CFG) [fn_code] + (a partial finite mapping from nodes to instructions). As in Cminor, + [fn_sig] is the function signature and [fn_stacksize] the number of bytes + for its stack-allocated activation record. [fn_params] is the list + of registers that are bound to the values of arguments at call time. + [fn_entrypoint] is the node of the first instruction of the function + in the CFG. *) + +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. + +(** * Operational semantics *) + +Definition genv := Genv.t fundef unit. + +(** The dynamic semantics of BTL is similar to RTL, + except that the step of one instruction is generalized into the run of one [iblock]. +*) + +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 *) + (succ: exit) (**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. + +(** outcome of a block execution *) +Record outcome := { + _rs: regset; + _m: mem; + _fin: option final +}. + +Section RELSEM. + +Variable 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. + +Local Open Scope option_monad_scope. + +(** 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 + (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 + (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 + | exec_store rs m chunk addr args src a m' + (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 + | 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) + | exec_seq_continue rs m b1 b2 rs1 m1 rs' m' ofin + (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 + (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 + . +Local Hint Constructors iblock_istep: core. + +(** A functional variant of [iblock_istep_run] of [iblock_istep]. +Lemma [iblock_istep_run_equiv] below provides a proof that "relation" [iblock_istep] is a "partial function". +*) +Fixpoint iblock_istep_run sp ib rs m: option outcome := + match ib with + | BF fin => + Some {| _rs := rs; _m := m; _fin := Some fin |} + (* basic instructions *) + | Bnop => + Some {| _rs := rs; _m:= m; _fin := None |} + | 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 => + 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 => + None (* TODO *) + | 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 |} + (* composed instructions *) + | Bseq b1 b2 => + SOME out1 <- iblock_istep_run sp b1 rs m IN + match out1.(_fin) with + | 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 _ => + SOME b <- eval_condition cond rs##args m IN + iblock_istep_run sp (if b then ifso else ifnot) rs m + end. + +Lemma iblock_istep_run_equiv sp rs m ib rs' m' ofin: + iblock_istep sp rs m ib rs' m' ofin <-> iblock_istep_run sp ib rs m = Some {| _rs := rs'; _m:= m'; _fin := ofin |}. +Proof. + constructor. + - induction 1; simpl; try autodestruct; try_simplify_someHyps. + - generalize rs m rs' m' ofin; clear rs m rs' m' ofin. + induction ib; simpl; repeat (try autodestruct; try_simplify_someHyps). + destruct o; try_simplify_someHyps; subst; eauto. +Qed. + +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': + sp = (Vptr stk Ptrofs.zero) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + final_step stack f sp rs m (Breturn or) + E0 (Returnstate stack (regmap_optget or Vundef rs) m') + | exec_Bcall sig ros args res pc' fd: + find_function ros rs = Some fd -> + funsig fd = sig -> + final_step stack f sp rs m (Bcall sig ros args res pc') + E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) + | exec_Btailcall sig ros args stk m' fd: + 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) + E0 (Callstate stack fd rs##args m') + | exec_Bbuiltin 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' -> + final_step stack f sp rs m (Bbuiltin ef args res pc') + t (State stack f sp pc' (regmap_setres res vres rs) m') + | exec_Bjumptable arg tbl n pc': + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + final_step stack f sp rs m (Bjumptable arg tbl) + E0 (State stack f sp pc' rs m) +. + +(** big-step execution of one iblock *) +Definition iblock_step stack f sp rs m ib t s: Prop := + exists rs' m' fin, iblock_istep sp rs m ib rs' m' (Some fin) /\ final_step stack f sp rs' m' fin t s. + +(** The transitions are presented as an inductive predicate + [step ge st1 t st2], where [ge] is the global environment, + [st1] the initial state, [st2] the final state, and [t] the trace + of system calls performed during this transition. *) + +Inductive step: state -> trace -> state -> Prop := + | exec_iblock stack ib f sp pc rs m t s: + (fn_code f)!pc = Some ib -> + iblock_step stack f sp rs m ib.(entry) t s -> + step (State stack f sp pc rs m) t s + | exec_function_internal stack f args m m' stk: + Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> + step (Callstate stack (Internal f) args m) + E0 (State stack + f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + m') + | exec_function_external stack ef args res t m m': + external_call ef ge args m t res m' -> + step (Callstate stack (External ef) args m) + t (Returnstate stack res m') + | exec_return stack res f sp pc rs vres m: + step (Returnstate (Stackframe res f sp pc rs :: stack) vres m) + E0 (State stack f sp pc (rs#res <- vres) m) +. + +End RELSEM. + +(** Execution of whole programs are described as sequences of transitions + from an initial state to a final state. An initial state is a [Callstate] + corresponding to the invocation of the ``main'' function of the program + without arguments and with an empty call stack. *) + +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). + +(** A final state is a [Returnstate] with an empty call stack. *) + +Inductive final_state: state -> int -> Prop := + | final_state_intro: forall r m, + final_state (Returnstate nil (Vint r) m) r. + +(** The small-step semantics for a program. *) + +Definition semantics (p: program) := + Semantics step (initial_state p) final_state (Genv.globalenv p). + + +(** Matching BTL and RTL code + +We should be able to define a single verifier able to prove a bisimulation between BTL and RTL code. + +NB 1: the proof of BTL -> RTL (plus simulation) should be much easier than RTL -> BTL (star simulation). + +NB 2: our scheme allows the BTL to duplicate some RTL code. +- in other words: RTL -> BTL allows tail duplication, loop unrolling, etc. Exactly like "Duplicate" on RTL. +- BTL -> RTL allows to "undo" some duplications (e.g. remove duplications that have not enabled interesting optimizations) ! + +Hence, in a sense, our verifier imitates the approach of Duplicate, where [dupmap] maps the BTL nodes to the RTL nodes. + +The [match_*] definitions gives a "relational" specification of the verifier... +*) + +Require Import Errors. + +Inductive match_final_inst (dupmap: PTree.t node): final -> instruction -> Prop := +(* TODO: it may be simplify the oracle for BTL -> RTL, but may makes the verifier more complex ? + | 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: + list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' -> + match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln') +. + +Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := + | ijo_None_left o: is_join_opt None o o + | ijo_None_right o: is_join_opt o None o + | ijo_Some x: is_join_opt (Some x) (Some x) (Some x) + . + +(* [match_iblock dupmap cfg isfst pc ib opc] means that [ib] match a block in a RTL code starting at [pc], with: + - [isfst] (in "input") indicates that no step in the surrounding block has been executed before entering [pc] + - if [opc] (in "ouput") is [None], this means that all branches of the block ends on a final instruction + - if [opc] is [Some pc'], this means that all branches of the block that do not exit, join on [pc']. +*) +Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> iblock -> (option node) -> Prop := + | mib_BF isfst fi pc 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': + cfg!pc = Some (Inop pc') -> + match_iblock dupmap cfg isfst pc Bnop (Some pc') + | mib_op isfst pc pc' op lr r: + 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' tm m a lr r: + cfg!pc = Some (Iload tm m a lr r pc') -> + match_iblock dupmap cfg isfst pc (Bload tm m a lr r) (Some pc') + | mib_store isfst pc pc' m a lr r: + cfg!pc = Some (Istore m a lr r pc') -> + match_iblock dupmap cfg isfst pc (Bstore m a lr r) (Some pc') + | mib_exit pc pc': + dupmap!pc = (Some pc') -> + match_iblock dupmap cfg false pc' (Bgoto pc) None + (* NB: on RTL side, we exit the block by a "basic" instruction (or Icond). + Thus some step should have been executed before [pc'] in the RTL code *) + | mib_seq_Some isfst b1 b2 pc1 pc2 opc: + match_iblock dupmap cfg isfst pc1 b1 (Some pc2) -> + match_iblock dupmap cfg false pc2 b2 opc -> + match_iblock dupmap cfg isfst pc1 (Bseq b1 b2) opc +(* | mib_seq_None isfst b1 b2 pc: + match_iblock dupmap cfg isfst pc b1 None -> + match_iblock dupmap cfg isfst (Bseq b1 b2) pc None + (* TODO: here [b2] is dead code ! Is this case really useful ? + The oracle could remove this dead code. + 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': + 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 + . + +(* Partial copy of Duplicate and Duplicateproof. *) + +Definition match_cfg dupmap (cfg: code) (cfg':RTL.code): Prop := + forall pc pc', dupmap!pc = Some pc' -> + exists ib, cfg!pc = Some ib /\ match_iblock dupmap cfg' true pc' ib.(entry) None. + +Record match_function dupmap f f': Prop := { + dupmap_correct: match_cfg dupmap (fn_code f) (RTL.fn_code f'); + dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f'); + preserv_fnsig: fn_sig f = RTL.fn_sig f'; + preserv_fnparams: fn_params f = RTL.fn_params f'; + preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' +}. + +Inductive match_fundef: fundef -> RTL.fundef -> Prop := + | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc' = Some pc) + : match_stackframes (Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). + +Inductive match_states: state -> RTL.state -> Prop := + | match_states_intro + dupmap st f sp pc rs m st' f' pc' + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc' = Some pc) + : match_states (State st f sp pc rs m) (RTL.State st' f' sp pc' rs m) + | match_states_call + st st' f f' args m + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_fundef f f') + : match_states (Callstate st f args m) (RTL.Callstate st' f' args m) + | match_states_return + st st' v m + (STACKS: list_forall2 match_stackframes st st') + : match_states (Returnstate st v m) (RTL.Returnstate st' v m) + . + +(** Shared verifier between RTL -> BTL and BTL -> RTL *) + +Local Open Scope error_monad_scope. + +Definition verify_is_copy dupmap n n' := + match dupmap!n with + | None => Error(msg "verify_is_copy None") + | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "verify_is_copy invalid map") end + end. + +Fixpoint verify_is_copy_list dupmap ln ln' := + match ln with + | n::ln => match ln' with + | n'::ln' => do _ <- verify_is_copy dupmap n n'; + verify_is_copy_list dupmap ln ln' + | nil => Error (msg "verify_is_copy_list: ln' bigger than ln") end + | nil => match ln' with + | n :: ln' => Error (msg "verify_is_copy_list: ln bigger than ln'") + | nil => OK tt end + end. + +Lemma verify_is_copy_correct dupmap n n' tt: + verify_is_copy dupmap n n' = OK tt -> + dupmap ! n = Some n'. +Proof. + unfold verify_is_copy; repeat autodestruct. + intros NP H; destruct (_ ! n) eqn:REVM; [|inversion H]. + eapply Pos.compare_eq in NP. congruence. +Qed. +Local Hint Resolve verify_is_copy_correct: core. + +Lemma verify_is_copy_list_correct dupmap ln: forall ln' tt, + verify_is_copy_list dupmap ln ln' = OK tt -> + list_forall2 (fun n n' => dupmap ! n = Some n') ln ln'. +Proof. + induction ln. + - intros. destruct ln'; monadInv H. constructor. + - intros. destruct ln'; monadInv H. constructor; eauto. +Qed. + +(* TODO: deuxième étape (après BTLtoRTLproof.plus_simulation) *) +Parameter verify_block: PTree.t node -> RTL.code -> bool -> node -> iblock -> res (option node). + +(* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *) +Lemma verify_block_correct: forall dupmap cfg isfst pc ib fin, + verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin. +Admitted. +Local Hint Resolve verify_block_correct: core. + +Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit := + match l with + | nil => OK tt + | (pc, pc')::l => + match cfg!pc with + | Some ib => do o <- verify_block dupmap cfg' true pc' ib.(entry); + match o with + | None => verify_blocks dupmap cfg cfg' l + | _ => Error(msg "verify_blocks.end") + end + | _ => Error(msg "verify_blocks.entry") + end + end. + +Definition verify_cfg dupmap (cfg: code) (cfg':RTL.code): res unit := + verify_blocks dupmap cfg cfg' (PTree.elements dupmap). + +Lemma verify_cfg_correct dupmap cfg cfg' tt: + verify_cfg dupmap cfg cfg' = OK tt -> + match_cfg dupmap cfg cfg'. +Proof. + unfold verify_cfg. + intros X pc pc' H; generalize X; clear X. + exploit PTree.elements_correct; eauto. + generalize tt pc pc' H; clear tt pc pc' H. + generalize (PTree.elements dupmap). + induction l as [|[pc1 pc1']l]; simpl. + - tauto. + - intros pc pc' DUP u H. + unfold bind. + repeat autodestruct. + intros; subst. + destruct H as [H|H]; eauto. + inversion H; subst. + eexists; split; eauto. +Qed. + +Definition verify_function dupmap f f' : res unit := + do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); + verify_cfg dupmap (fn_code f) (RTL.fn_code f'). + +Lemma verify_function_correct dupmap f f' tt: + verify_function dupmap f f' = OK tt -> + fn_sig f = RTL.fn_sig f' -> + fn_params f = RTL.fn_params f' -> + fn_stacksize f = RTL.fn_stacksize f' -> + match_function dupmap f f'. +Proof. + unfold verify_function; intro VERIF. monadInv VERIF. + constructor; eauto. + eapply verify_cfg_correct; eauto. +Qed. diff --git a/scheduling/BTLtoRTL.v b/scheduling/BTLtoRTL.v new file mode 100644 index 00000000..b64fd87a --- /dev/null +++ b/scheduling/BTLtoRTL.v @@ -0,0 +1,23 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad BTL. + +Require Import Errors Linking. + +(** External oracle *) +Parameter btl2rtl: function -> RTL.code * node * (PTree.t node). + +Local Open Scope error_monad_scope. + +Definition transf_function (f: function) : res RTL.function := + let (tcte, dupmap) := btl2rtl f in + let (tc, te) := tcte in + let f' := RTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in + do u <- verify_function dupmap f f'; + OK f'. + +Definition transf_fundef (f: fundef) : res RTL.fundef := + transf_partial_fundef transf_function f. + +Definition transf_program (p: program) : res RTL.program := + transform_partial_program transf_fundef p. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v new file mode 100644 index 00000000..26b5bae6 --- /dev/null +++ b/scheduling/BTLtoRTLproof.v @@ -0,0 +1,127 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad BTL. + +Require Import Errors Linking BTLtoRTL. + + +(*****************************) +(* Put this in an other file *) + +Require Import Linking. + +Lemma transf_function_correct f f': + transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. +Proof. + unfold transf_function; unfold bind. repeat autodestruct. + intros H _ _ X. inversion X; subst; clear X. + eexists; eapply verify_function_correct; simpl; eauto. +Qed. + +Lemma transf_fundef_correct f f': + transf_fundef f = OK f' -> match_fundef f f'. +Proof. + intros TRANSF; destruct f; simpl; monadInv TRANSF. + + exploit transf_function_correct; eauto. + intros (dupmap & MATCH_F). + eapply match_Internal; eauto. + + eapply match_External. +Qed. + +Definition match_prog (p: program) (tp: RTL.program) := + match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp. + +Lemma transf_program_match: + forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog. +Proof. + intros. eapply match_transform_partial_program_contextual; eauto. +Qed. + +Section RTL_SIMULATES_BTL. + +Variable prog: program. +Variable tprog: RTL.program. + +Hypothesis TRANSL: match_prog prog tprog. + +Let ge := Genv.globalenv prog. +Let tge := Genv.globalenv tprog. + +Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. +Proof. + rewrite <- (Genv.find_symbol_match TRANSL). reflexivity. +Qed. + +Lemma senv_preserved: Senv.equiv ge tge. +Proof. + eapply (Genv.senv_match TRANSL). +Qed. + +Lemma functions_translated (v: val) (f: fundef): + Genv.find_funct ge v = Some f -> + exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tge v = Some tf /\ linkorder cunit prog. +Proof. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & A & B & C). + repeat eexists; intuition eauto. + + unfold incl; auto. + + eapply linkorder_refl. +Qed. + +Lemma function_ptr_translated v f: + Genv.find_funct_ptr ge v = Some f -> + exists tf, + Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = OK tf. +Proof. + intros. + exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. +Qed. + +Lemma function_sig_translated f tf: transf_fundef f = OK tf -> RTL.funsig tf = funsig f. +Proof. + intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto. + erewrite preserv_fnsig; eauto. +Qed. + +Lemma transf_initial_states s1: + initial_state prog s1 -> + exists s2, RTL.initial_state tprog s2 /\ match_states s1 s2. +Proof. + intros. inv H. + exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF). + eexists. split. + - econstructor; eauto. + + eapply (Genv.init_mem_transf_partial TRANSL); eauto. + + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. + symmetry. eapply match_program_main. eauto. + + erewrite function_sig_translated; eauto. + - constructor; eauto. + constructor. + apply transf_fundef_correct; auto. +Qed. + +Lemma transf_final_states s1 s2 r: + match_states s1 s2 -> final_state s1 r -> RTL.final_state s2 r. +Proof. + intros. inv H0. inv H. inv STACKS. constructor. +Qed. + +Theorem plus_simulation s1 t s1': + step ge s1 t s1' -> + forall s2 (MS: match_states s1 s2), + exists s2', + plus RTL.step tge s2 t s2' + /\ match_states s1' s2'. +Admitted. (* TODO: première étape *) + +Theorem transf_program_correct: + forward_simulation (semantics prog) (RTL.semantics tprog). +Proof. + eapply forward_simulation_plus with match_states. + - eapply senv_preserved. + - eapply transf_initial_states. + - eapply transf_final_states. + - eapply plus_simulation. +Qed. + +End RTL_SIMULATES_BTL. -- cgit From 989e85b700161867c2f80df55c2dfaaaa5fe82b4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 30 Apr 2021 16:52:21 +0200 Subject: BTLtoRTL: proof for internal/external/return states --- scheduling/BTL.v | 4 ++-- scheduling/BTLtoRTLproof.v | 25 +++++++++++++++++++++++++ 2 files changed, 27 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 8bdafb89..8529010c 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -451,7 +451,7 @@ Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := | match_stackframe_intro dupmap res f sp pc rs f' pc' (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc' = Some pc) + (DUPLIC: dupmap!pc = Some pc') : match_stackframes (Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). Inductive match_states: state -> RTL.state -> Prop := @@ -459,7 +459,7 @@ Inductive match_states: state -> RTL.state -> Prop := dupmap st f sp pc rs m st' f' pc' (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc' = Some pc) + (DUPLIC: dupmap!pc = Some pc') : match_states (State st f sp pc rs m) (RTL.State st' f' sp pc' rs m) | match_states_call st st' f f' args m diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 26b5bae6..a3be20b5 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -112,6 +112,31 @@ Theorem plus_simulation s1 t s1': exists s2', plus RTL.step tge s2 t s2' /\ match_states s1' s2'. +Proof. + induction 1; intros; inv MS. + - eapply dupmap_correct in DUPLIC; eauto. + destruct DUPLIC as (ib' & FNC & MIB). + inv H0. destruct H1 as (m' & fin & IBIS & FS). + admit. + (* Définir une relation inductive qui provient du MS, + et qui relie le iblock_istep avec le RTL. Il faut prendre en compte + le "isfst" pour savoir si on doit avoir du "stuttering" côté RTL. *) + (* exec_function_internal *) + - inversion TRANSF as [dupmap f0 f0' MATCHF|]; subst. eexists. split. + + eapply plus_one. apply RTL.exec_function_internal. + erewrite <- preserv_fnstacksize; eauto. + + erewrite <- preserv_fnparams; eauto. + eapply match_states_intro; eauto. + apply dupmap_entrypoint. assumption. + (* exec_function_external *) + - inversion TRANSF as [|]; subst. eexists. split. + + eapply plus_one. econstructor. + eapply external_call_symbols_preserved; eauto. apply senv_preserved. + + constructor. assumption. + (* exec_return *) + - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split. + + eapply plus_one. constructor. + + inv H1. econstructor; eauto. Admitted. (* TODO: première étape *) Theorem transf_program_correct: -- cgit From 3751a5570441faed6147f0d7e80dffbb2342d258 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 1 May 2021 17:11:48 +0200 Subject: debroussaillage et precisions... --- scheduling/BTL.v | 30 ++++++++++----------- scheduling/BTLtoRTLproof.v | 66 ++++++++++++++++++++++++++++++++++++++++------ 2 files changed, 73 insertions(+), 23 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 8529010c..28941e01 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -292,26 +292,26 @@ Definition iblock_step stack f sp rs m ib t s: Prop := of system calls performed during this transition. *) Inductive step: state -> trace -> state -> Prop := - | exec_iblock stack ib f sp pc rs m t s: - (fn_code f)!pc = Some ib -> - iblock_step stack f sp rs m ib.(entry) t s -> - step (State stack f sp pc rs m) t s - | exec_function_internal stack f args m m' stk: - Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> - step (Callstate stack (Internal f) args m) - E0 (State stack + | exec_iblock stack ib f sp pc rs m t s + (PC: (fn_code f)!pc = Some ib) + (STEP: iblock_step stack f sp rs m ib.(entry) t s) + :step (State stack f sp pc rs m) t s + | exec_function_internal stack f args m m' stk + (ALLOC: Mem.alloc m 0 f.(fn_stacksize) = (m', stk)) + :step (Callstate stack (Internal f) args m) + E0 (State stack f (Vptr stk Ptrofs.zero) f.(fn_entrypoint) (init_regs args f.(fn_params)) m') - | exec_function_external stack ef args res t m m': - external_call ef ge args m t res m' -> - step (Callstate stack (External ef) args m) - t (Returnstate stack res m') - | exec_return stack res f sp pc rs vres m: - step (Returnstate (Stackframe res f sp pc rs :: stack) vres m) - E0 (State stack f sp pc (rs#res <- vres) m) + | exec_function_external stack ef args res t m m' + (EXTCALL: external_call ef ge args m t res m') + :step (Callstate stack (External ef) args m) + t (Returnstate stack res m') + | exec_return stack res f sp pc rs vres m + :step (Returnstate (Stackframe res f sp pc rs :: stack) vres m) + E0 (State stack f sp pc (rs#res <- vres) m) . End RELSEM. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index a3be20b5..92e05d48 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,6 +106,57 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +Lemma iblock_step_simulation sp dupmap stack stack' f f' cfg' ib rs0 m0 rs1 m1 ofin + (STACKS: list_forall2 match_stackframes stack stack') + (TRANSF: match_function dupmap f f') + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin) + fin pc0 t s isfst + (OFIN: ofin = Some fin) + (MIB : match_iblock dupmap cfg' isfst pc0 ib None) + (FS : final_step ge stack f sp rs1 m1 fin t s) + :exists s', (if isfst then plus RTL.step else star RTL.step) tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. +Proof. + (* TODO: généraliser ce lemme pour pouvoir le prouver par induction sur IBIS: + => il faut en particulier généraliser l'hypothèse MIB qui relie le iblock_istep ib en cours d'exécution. + avec le code RTL à partir de pc0. + => ici, le "isfst" a déjà été généralisé: quand il vaut "false", ça veut dire qu'on a le droit de faire du "stuttering" côté RTL. + => il reste à comprendre comment généraliser le "None" en "opc" + ainsi que l'hypothese OFIN pour autoriser le cas "ofin=None" (nécessaire pour l'induction). + Idée: si "ofin = None" alors il y a un pc1 tq "opc = Some pc1" qui permet d'exécuter la suite du bloc... + + Au final, il faut sans doute introduire un "Inductive" pour capturer ces idées dans un prédicat "lisible"/"manipulable"... + *) + induction IBIS. + - (* exec_final *) + try_simplify_someHyps. (* TODO: introduire un lemme pour ce cas specifique ? *) + admit. + - (* exec_nop *) + try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + - (* exec_op *) + try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + - (* exec_load_TRAP *) + try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + - (* exec_load_store *) + try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + - (* exec_seq_stop *) + try_simplify_someHyps. + inv MIB. + eapply IHIBIS; eauto. + (* TODO: c'est ici qu'on voit que l'hypothèse MIB est trop restrictive actuellement + normalement, l'hypothèse d'induction IHIBIS devrait permettre de conclure quasi-directement ici. + *) + admit. + - (* exec_seq_continue *) + (* TODO: ici l'hypothèse d'induction IHIBIS1 n'est pas utilisable à cause de OFIN trop restrictive *) + try_simplify_someHyps. + inv MIB. + admit. + - (* exec_cond *) + try_simplify_someHyps. + inv MIB. + admit. +Admitted. + Theorem plus_simulation s1 t s1': step ge s1 t s1' -> forall s2 (MS: match_states s1 s2), @@ -113,14 +164,11 @@ Theorem plus_simulation s1 t s1': plus RTL.step tge s2 t s2' /\ match_states s1' s2'. Proof. - induction 1; intros; inv MS. + destruct 1; intros; inv MS. - eapply dupmap_correct in DUPLIC; eauto. destruct DUPLIC as (ib' & FNC & MIB). - inv H0. destruct H1 as (m' & fin & IBIS & FS). - admit. - (* Définir une relation inductive qui provient du MS, - et qui relie le iblock_istep avec le RTL. Il faut prendre en compte - le "isfst" pour savoir si on doit avoir du "stuttering" côté RTL. *) + try_simplify_someHyps. destruct STEP as (rs' & m' & fin & IBIS & FS). + intros; exploit iblock_step_simulation; eauto. (* exec_function_internal *) - inversion TRANSF as [dupmap f0 f0' MATCHF|]; subst. eexists. split. + eapply plus_one. apply RTL.exec_function_internal. @@ -134,10 +182,12 @@ Proof. eapply external_call_symbols_preserved; eauto. apply senv_preserved. + constructor. assumption. (* exec_return *) - - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split. + - inversion STACKS as [|a1 al b1 bl H1 HL]; subst. + destruct b1 as [res' f' sp' pc' rs']. + eexists. split. + eapply plus_one. constructor. + inv H1. econstructor; eauto. -Admitted. (* TODO: première étape *) +Qed. Theorem transf_program_correct: forward_simulation (semantics prog) (RTL.semantics tprog). -- cgit From 63d38609a1418ae99474923ecb6f512e23cee8f5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 2 May 2021 15:48:10 +0200 Subject: BTL roadmap --- scheduling/BTLroadmap.md | 79 ++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 79 insertions(+) create mode 100644 scheduling/BTLroadmap.md (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md new file mode 100644 index 00000000..6babf47c --- /dev/null +++ b/scheduling/BTLroadmap.md @@ -0,0 +1,79 @@ +# BTL Development Roadmap + +BTL aims to be an IR dedicated to defensive certification of middle-end optimizations (before register allocation). +It provides a CFG of "loop-free" blocks, where each block is run in one step emitting at most a single observational event. +The "local" optimizations (i.e. preserving "locally" the semantics of such blocks) would be checked by symbolic execution with rewriting rules. +The only info required from oracles: a "dupmap" mapping block entries. +In contrast, the "global" optimizations would require some invariants annotations at block entry (provided by oracles). + +Examples of optimizations that we aim to support: + + - instruction scheduling + - instruction rewritings (instruction selection) + - branch duplication, "if-lifting" (e.g. side-exit moved up in "traces"). + - strength-reduction + - SSA optimizations + +We sketch below the various kinds of supported optimizations in development order... + +## Block boundaries, branch duplication or factorization + +Two possibility of branch duplications (e.g tail-duplication, loop unrolling, etc): + +- during RTL -> BTL (while "building" BTL blocks) +- during BTL -> BTL. Typically, the "if-lifting" à la Justus could be performed/verified in a single pass here ! + +Branch factorization should also be possible in BTL -> RTL pass. Example: revert "loop-unrolling". + +**IMPLEM NOTE** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. + +**CURRENT STATUS** + +- verifier: specified but not yet implemented. +- BTL -> RTL: proof started. + +## "Basic" symbolic execution / symbolic simulation + +IMPLEM PLAN in two steps: + +1. a very basic version without rewriting rules, and without support for simulation modulo-liveness. + +2. Support a "SVundef" symbolic value, and implement symbolic simulation wrt "less_undef" relation. +This will allow more general rewriting rules and will generalize modulo-liveness (see next section). + +## Simulation modulo liveness + +Extends BTL with the possibility of destroying a set of registers at each exit (a destroyed register is simply set to Vundef). + +We can the provide a "MarkUnused" pass of BTL -> BTL that insert these destructions at each exit using liveness info ! +(NB: the liveness info could be provided by the RTL -> BTL oracle like on RTLpath). + +In the symbolic execution, destroying a register corresponds to associate it SVundef value. + +NB: this approach through "destroyed registers" and "SVundef" seems more modular (thus simpler) and more general than the current approach on RTLpath. + +The proof of "simulation modulo liveness" of a BTL B1 by a BTL B2 will simply corresponds to compose two forward simulations: B1 --> MarkUnused(B1) --> B2. + +## Port of RTLpath optimizations to BTL + +- Generalize superblock scheduling for a notion of "extended-blocks" such that each instruction of a block, except the block entry, has a single predecessor in the block. +- Port rewriting rules of RTLpath. +- Justus's "poor man SSA" + if-lifting. + +## Invariants at block entry + +Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. + +Case-study: support of strength-reduction. + +## Support of SSA-optimizations + +Extends BTL with "register renamings" at exits (in a sense, "register renaming" extends the feature of "register destroying"). + +This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. + +## Alias analysis in the symbolic simulation + +... + + -- cgit From 846fe8fcae71c964b635a56a5e7fdf20eb4b85e5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 2 May 2021 16:36:34 +0200 Subject: more implem notes in BTLroadmap --- scheduling/BTLroadmap.md | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 6babf47c..27e74de5 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -25,7 +25,7 @@ Two possibility of branch duplications (e.g tail-duplication, loop unrolling, et Branch factorization should also be possible in BTL -> RTL pass. Example: revert "loop-unrolling". -**IMPLEM NOTE** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. +**IMPLEM NOTE:** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. **CURRENT STATUS** @@ -41,6 +41,8 @@ IMPLEM PLAN in two steps: 2. Support a "SVundef" symbolic value, and implement symbolic simulation wrt "less_undef" relation. This will allow more general rewriting rules and will generalize modulo-liveness (see next section). +**IMPLEM NOTE:** use a symbolic execution in Continuation-Passing Style, in order to have of simple handling of "side-exits" within "branch unfolding". + ## Simulation modulo liveness Extends BTL with the possibility of destroying a set of registers at each exit (a destroyed register is simply set to Vundef). @@ -54,6 +56,11 @@ NB: this approach through "destroyed registers" and "SVundef" seems more modular The proof of "simulation modulo liveness" of a BTL B1 by a BTL B2 will simply corresponds to compose two forward simulations: B1 --> MarkUnused(B1) --> B2. +**IMPLEM NOTE:** "MarkUnused" checks the correctness of liveness informations while performing its transformation. +It should computes alive register from block entry, by forward proceeding of "internal instructions" (like in RTLPathLivegen). +When arriving on a final instruction, we marked as "destroyed" all alive registers that are not in the input registers of the exit +Semantically, the destruction of registers happens "in parallel" of the final step of the block. + ## Port of RTLpath optimizations to BTL - Generalize superblock scheduling for a notion of "extended-blocks" such that each instruction of a block, except the block entry, has a single predecessor in the block. -- cgit From dcaec659404e11fd927b5e8af3e44958d4c6c950 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 3 May 2021 11:44:42 +0200 Subject: Trying a inductive predicate for BTL-RTL proof --- scheduling/BTLtoRTLproof.v | 51 +++++++++++++++++++++++++++++++++++++--------- 1 file changed, 41 insertions(+), 10 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 92e05d48..7f6f21c7 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,14 +106,28 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -Lemma iblock_step_simulation sp dupmap stack stack' f f' cfg' ib rs0 m0 rs1 m1 ofin +Inductive iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg: state -> trace -> bool -> option final -> option node -> Prop := + | ibis_synced ib rs1 m1 opc pc1 + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 None) + (HOPC: opc = Some pc1) + (MIB : match_iblock dupmap cfg true pc0 ib opc) + : iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg (State stack f sp pc0 rs0 m0) E0 true None opc + | ibis_stutter ib rs1 m1 fin ofin s t + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) + (HFIN: ofin = Some fin) + (MIB : match_iblock dupmap cfg false pc0 ib None) + (FS : final_step ge stack f sp rs1 m1 fin t s) + : iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg s t false ofin None. + +Lemma iblock_step_simulation + sp dupmap stack stack' f f' rs0 m0 ofin pc0 opc t s isfst (STACKS: list_forall2 match_stackframes stack stack') (TRANSF: match_function dupmap f f') - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin) - fin pc0 t s isfst - (OFIN: ofin = Some fin) - (MIB : match_iblock dupmap cfg' isfst pc0 ib None) - (FS : final_step ge stack f sp rs1 m1 fin t s) + (*(IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin)*) + (*(MIB : match_iblock dupmap cfg' isfst pc0 ib opc)*) + (*(FS : final_step ge stack f sp rs1 m1 fin t s)*) +(*FNC : (fn_code f) ! pc = Some ib*) + (IBGEN: iblock_istep_gen sp dupmap stack f rs0 m0 pc0 (RTL.fn_code f') s t isfst ofin opc) :exists s', (if isfst then plus RTL.step else star RTL.step) tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. (* TODO: généraliser ce lemme pour pouvoir le prouver par induction sur IBIS: @@ -126,7 +140,22 @@ Proof. Au final, il faut sans doute introduire un "Inductive" pour capturer ces idées dans un prédicat "lisible"/"manipulable"... *) - induction IBIS. + (* XXX keep IBIS, MIB, and FS ? *) + induction IBGEN. + - inv IBIS. + + inv MIB. + eexists. split. apply plus_one. + eapply exec_Inop; eauto. + econstructor; eauto. + try_simplify_someHyps. + admit. + + admit. + + admit. + + admit. + + admit. + + admit. + - admit. + (* induction IBIS. - (* exec_final *) try_simplify_someHyps. (* TODO: introduire un lemme pour ce cas specifique ? *) admit. @@ -154,7 +183,7 @@ Proof. - (* exec_cond *) try_simplify_someHyps. inv MIB. - admit. + admit.*) Admitted. Theorem plus_simulation s1 t s1': @@ -168,7 +197,8 @@ Proof. - eapply dupmap_correct in DUPLIC; eauto. destruct DUPLIC as (ib' & FNC & MIB). try_simplify_someHyps. destruct STEP as (rs' & m' & fin & IBIS & FS). - intros; exploit iblock_step_simulation; eauto. + admit. + (*intros; exploit iblock_step_simulation; eauto.*) (* exec_function_internal *) - inversion TRANSF as [dupmap f0 f0' MATCHF|]; subst. eexists. split. + eapply plus_one. apply RTL.exec_function_internal. @@ -187,7 +217,8 @@ Proof. eexists. split. + eapply plus_one. constructor. + inv H1. econstructor; eauto. -Qed. +(*Qed.*) +Admitted. Theorem transf_program_correct: forward_simulation (semantics prog) (RTL.semantics tprog). -- cgit From c8b2438685f1a6c4a98cf58727d32a9e474c4f34 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 3 May 2021 13:53:43 +0200 Subject: A better structure for inductive prop --- scheduling/BTLtoRTLproof.v | 60 +++++++++++++++++++--------------------------- 1 file changed, 25 insertions(+), 35 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 7f6f21c7..904b3814 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,28 +106,29 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -Inductive iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg: state -> trace -> bool -> option final -> option node -> Prop := - | ibis_synced ib rs1 m1 opc pc1 - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 None) +Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := + | ibis_synced opc pc1 (HOPC: opc = Some pc1) (MIB : match_iblock dupmap cfg true pc0 ib opc) - : iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg (State stack f sp pc0 rs0 m0) E0 true None opc - | ibis_stutter ib rs1 m1 fin ofin s t - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) + : iblock_istep_gen sp dupmap stack f pc0 cfg ib E0 true None opc + | ibis_stutter rs1 m1 fin ofin s t (HFIN: ofin = Some fin) (MIB : match_iblock dupmap cfg false pc0 ib None) (FS : final_step ge stack f sp rs1 m1 fin t s) - : iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg s t false ofin None. + : iblock_istep_gen sp dupmap stack f pc0 cfg ib t false ofin None. + +(*Lemma iblock_istep_gen_correct *) +(*TODO *) Lemma iblock_step_simulation - sp dupmap stack stack' f f' rs0 m0 ofin pc0 opc t s isfst + sp dupmap stack stack' f f' rs0 m0 rs1 m1 ib ofin pc0 opc t s isfst (STACKS: list_forall2 match_stackframes stack stack') (TRANSF: match_function dupmap f f') - (*(IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin)*) + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin) (*(MIB : match_iblock dupmap cfg' isfst pc0 ib opc)*) (*(FS : final_step ge stack f sp rs1 m1 fin t s)*) (*FNC : (fn_code f) ! pc = Some ib*) - (IBGEN: iblock_istep_gen sp dupmap stack f rs0 m0 pc0 (RTL.fn_code f') s t isfst ofin opc) + (IBGEN: iblock_istep_gen sp dupmap stack f pc0 (RTL.fn_code f') ib t isfst ofin opc) :exists s', (if isfst then plus RTL.step else star RTL.step) tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. (* TODO: généraliser ce lemme pour pouvoir le prouver par induction sur IBIS: @@ -141,35 +142,24 @@ Proof. Au final, il faut sans doute introduire un "Inductive" pour capturer ces idées dans un prédicat "lisible"/"manipulable"... *) (* XXX keep IBIS, MIB, and FS ? *) - induction IBGEN. - - inv IBIS. - + inv MIB. - eexists. split. apply plus_one. - eapply exec_Inop; eauto. - econstructor; eauto. - try_simplify_someHyps. - admit. - + admit. - + admit. - + admit. - + admit. - + admit. - - admit. - (* induction IBIS. + induction IBIS. - (* exec_final *) try_simplify_someHyps. (* TODO: introduire un lemme pour ce cas specifique ? *) admit. - - (* exec_nop *) - try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + - inversion IBGEN; subst; try_simplify_someHyps. + inv MIB. + eexists. split. apply plus_one. + eapply exec_Inop; eauto. + try_simplify_someHyps. + admit. - (* exec_op *) - try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + admit. (* cas absurde car hypothese OFIN trop restrictive *) - (* exec_load_TRAP *) - try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + admit. (* cas absurde car hypothese OFIN trop restrictive *) - (* exec_load_store *) - try_simplify_someHyps. (* cas absurde car hypothese OFIN trop restrictive *) + admit. (* cas absurde car hypothese OFIN trop restrictive *) - (* exec_seq_stop *) - try_simplify_someHyps. - inv MIB. + (*inv MIB.*) eapply IHIBIS; eauto. (* TODO: c'est ici qu'on voit que l'hypothèse MIB est trop restrictive actuellement normalement, l'hypothèse d'induction IHIBIS devrait permettre de conclure quasi-directement ici. @@ -178,12 +168,12 @@ Proof. - (* exec_seq_continue *) (* TODO: ici l'hypothèse d'induction IHIBIS1 n'est pas utilisable à cause de OFIN trop restrictive *) try_simplify_someHyps. - inv MIB. + (*inv MIB.*) admit. - (* exec_cond *) try_simplify_someHyps. - inv MIB. - admit.*) + (*inv MIB.*) + admit. Admitted. Theorem plus_simulation s1 t s1': -- cgit From cc5f8c4e4d71d824a4d1818745d6bf1015b98f98 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 3 May 2021 13:55:15 +0200 Subject: une autre suggestion --- scheduling/BTLtoRTLproof.v | 14 ++++++++++++++ 1 file changed, 14 insertions(+) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 7f6f21c7..ab3cab69 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,6 +106,20 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +(* suggestion *) +Lemma iblock_stepi_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): + forall cfg' pc0 opc isfst + (MIB: match_iblock dupmap cfg' isfst pc0 ib opc) + , + match ofin with + | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) + | Some fin => True (* TODO: A CHANGER *) + end. +Proof. + induction IBIS; simpl; auto. +Admitted. + Inductive iblock_istep_gen sp dupmap stack f rs0 m0 pc0 cfg: state -> trace -> bool -> option final -> option node -> Prop := | ibis_synced ib rs1 m1 opc pc1 (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 None) -- cgit From 4b47467a486f50b9b3dc861e93e1399544ac6432 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 3 May 2021 15:58:24 +0200 Subject: some advance --- scheduling/BTLtoRTLproof.v | 47 ++++++++++++++++++++++++++++++++++++++-------- 1 file changed, 39 insertions(+), 8 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 97ac75fb..8f892020 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -107,17 +107,45 @@ Proof. Qed. (* suggestion *) -Lemma iblock_stepi_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin +Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): - forall cfg' pc0 opc isfst - (MIB: match_iblock dupmap cfg' isfst pc0 ib opc) + forall pc0 opc isfst + (MIB: match_iblock dupmap (RTL.fn_code f') isfst pc0 ib opc) , match ofin with | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => True (* TODO: A CHANGER *) end. Proof. - induction IBIS; simpl; auto. + induction IBIS; simpl; auto; + intros. + - inv MIB. exists pc'; split; auto. + apply plus_one. apply exec_Inop; trivial. + - inv MIB. exists pc'; split; auto. + apply plus_one. eapply exec_Iop; eauto. + erewrite <- eval_operation_preserved; eauto. + intros; rewrite symbols_preserved; trivial. + - inv MIB. exists pc'; split; auto. + apply plus_one. eapply exec_Iload; eauto. + erewrite <- eval_addressing_preserved; eauto. + intros; rewrite symbols_preserved; trivial. + - inv MIB. exists pc'; split; auto. + apply plus_one. eapply exec_Istore; eauto. + erewrite <- eval_addressing_preserved; eauto. + intros; rewrite symbols_preserved; trivial. + - + destruct ofin; simpl; auto. + inv MIB. + admit. + (*eexists; split; auto.*) + (*apply IHIBIS2.*) + (*inv MIB. exists pc'; split; auto.*) + (*apply plus_one. eapply exec_Istore; eauto.*) + (*erewrite <- eval_addressing_preserved; eauto.*) + (*intros; rewrite symbols_preserved; trivial.*) + - destruct ofin; simpl; auto. inv MIB. + destruct b. + eapply IHIBIS. Admitted. Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := @@ -154,10 +182,13 @@ Proof. *) (* XXX keep IBIS, MIB, and FS ? *) induction IBIS. - - (* exec_final *) - try_simplify_someHyps. (* TODO: introduire un lemme pour ce cas specifique ? *) - admit. + - admit. + - inv IBGEN; try_simplify_someHyps. + inv MIB. eexists; split. + apply plus_one. econstructor; eauto. + (* - inversion IBGEN; subst; try_simplify_someHyps. + exploit iblock_istep_simulation; eauto. inv MIB. eexists. split. apply plus_one. eapply exec_Inop; eauto. @@ -184,7 +215,7 @@ Proof. - (* exec_cond *) try_simplify_someHyps. (*inv MIB.*) - admit. + admit.*) Admitted. Theorem plus_simulation s1 t s1': -- cgit From 2c17196bd28d4936381a43fa1652848d275639ec Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 3 May 2021 18:22:37 +0200 Subject: essais --- scheduling/BTLtoRTLproof.v | 47 ++++++++++++++++++++++++++++++---------------- 1 file changed, 31 insertions(+), 16 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 8f892020..ef640d22 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,6 +106,12 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +Definition is_goto (i: final): option exit := + match i with + | Bgoto pc => Some pc + | _ => None + end. + (* suggestion *) Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): @@ -114,11 +120,14 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin , match ofin with | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) - | Some fin => True (* TODO: A CHANGER *) + | Some fin => + (* A CHANGER ? *) + exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin opc + /\ star RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) end. Proof. - induction IBIS; simpl; auto; - intros. + induction IBIS; simpl; intros. + - repeat econstructor; eauto. - inv MIB. exists pc'; split; auto. apply plus_one. apply exec_Inop; trivial. - inv MIB. exists pc'; split; auto. @@ -133,19 +142,25 @@ Proof. apply plus_one. eapply exec_Istore; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. - - - destruct ofin; simpl; auto. - inv MIB. - admit. - (*eexists; split; auto.*) - (*apply IHIBIS2.*) - (*inv MIB. exists pc'; split; auto.*) - (*apply plus_one. eapply exec_Istore; eauto.*) - (*erewrite <- eval_addressing_preserved; eauto.*) - (*intros; rewrite symbols_preserved; trivial.*) - - destruct ofin; simpl; auto. inv MIB. - destruct b. - eapply IHIBIS. + - inv MIB. + exploit IHIBIS; eauto. + intros (isfst' & pc1 & M1 & STAR). + inv M1. + - inv MIB. + exploit IHIBIS1; eauto. + intros (pc1 & EQpc1 & STEP1); inv EQpc1. + exploit IHIBIS2; eauto. + destruct ofin; simpl. + + intros (ifst2 & pc2 & M2 & STEP2). + repeat eexists; eauto. + eapply star_trans; eauto. + eapply plus_star; eauto. + eauto. + + intros (pc2 & EQpc2 & STEP2); inv EQpc2. + eexists; split; auto. + eapply plus_trans; eauto. + - inv MIB. + destruct b; exploit IHIBIS; eauto. Admitted. Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := -- cgit From 0ed3ee91bad8dc17a77d742ed9c43742c0a833ba Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 4 May 2021 07:15:14 +0200 Subject: essai avec le cond_star_step --- scheduling/BTLtoRTLproof.v | 57 ++++++++++++++++++++++++++++++++++++---------- 1 file changed, 45 insertions(+), 12 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index ef640d22..31652697 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,11 +106,30 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +(* FIXME: unused ? Definition is_goto (i: final): option exit := match i with | Bgoto pc => Some pc | _ => None end. +*) + +(* variant of [star RTL.step] but requiring proposition [P] on the [refl] (stutttering) case. *) +Inductive cond_star_step (P: Prop): RTL.state -> trace -> RTL.state -> Prop := + | css_refl s: P -> cond_star_step P s E0 s + | css_plus s1 t s2: plus RTL.step tge s1 t s2 -> cond_star_step P s1 t s2 + . + +Lemma css_plus_trans P Q s0 s1 s2 t: + plus RTL.step tge s0 E0 s1 -> + cond_star_step P s1 t s2 -> + cond_star_step Q s0 t s2. +Proof. + intros PLUS STAR. + eapply css_plus. + inv STAR; auto. + eapply plus_trans; eauto. +Qed. (* suggestion *) Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin @@ -122,12 +141,15 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => (* A CHANGER ? *) - exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin opc - /\ star RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) + exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None + /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) end. Proof. induction IBIS; simpl; intros. - - repeat econstructor; eauto. + - assert (X: opc = None). { inv MIB; auto. } + subst. + repeat eexists; eauto. + eapply css_refl; auto. - inv MIB. exists pc'; split; auto. apply plus_one. apply exec_Inop; trivial. - inv MIB. exists pc'; split; auto. @@ -144,8 +166,6 @@ Proof. intros; rewrite symbols_preserved; trivial. - inv MIB. exploit IHIBIS; eauto. - intros (isfst' & pc1 & M1 & STAR). - inv M1. - inv MIB. exploit IHIBIS1; eauto. intros (pc1 & EQpc1 & STEP1); inv EQpc1. @@ -153,9 +173,7 @@ Proof. destruct ofin; simpl. + intros (ifst2 & pc2 & M2 & STEP2). repeat eexists; eauto. - eapply star_trans; eauto. - eapply plus_star; eauto. - eauto. + eapply css_plus_trans; eauto. + intros (pc2 & EQpc2 & STEP2); inv EQpc2. eexists; split; auto. eapply plus_trans; eauto. @@ -163,6 +181,20 @@ Proof. destruct b; exploit IHIBIS; eauto. Admitted. +Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s + (STACKS: list_forall2 match_stackframes stack stack') + (TRANSF: match_function dupmap f f') + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) + (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) + (FS : final_step ge stack f sp rs1 m1 fin t s) + :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. +Proof. + intros; exploit iblock_istep_simulation; eauto. + simpl. +Admitted. + +(* BROUILLON + Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := | ibis_synced opc pc1 (HOPC: opc = Some pc1) @@ -233,6 +265,8 @@ Proof. admit.*) Admitted. +*) + Theorem plus_simulation s1 t s1': step ge s1 t s1' -> forall s2 (MS: match_states s1 s2), @@ -244,8 +278,7 @@ Proof. - eapply dupmap_correct in DUPLIC; eauto. destruct DUPLIC as (ib' & FNC & MIB). try_simplify_someHyps. destruct STEP as (rs' & m' & fin & IBIS & FS). - admit. - (*intros; exploit iblock_step_simulation; eauto.*) + intros; exploit iblock_step_simulation; eauto. (* exec_function_internal *) - inversion TRANSF as [dupmap f0 f0' MATCHF|]; subst. eexists. split. + eapply plus_one. apply RTL.exec_function_internal. @@ -264,8 +297,8 @@ Proof. eexists. split. + eapply plus_one. constructor. + inv H1. econstructor; eauto. -(*Qed.*) -Admitted. +Qed. +(* Admitted. *) Theorem transf_program_correct: forward_simulation (semantics prog) (RTL.semantics tprog). -- cgit From d44d44c8fffbd85d501e30b7dfdf93683ed9bb4d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 10:42:08 +0200 Subject: proof for Icond in iblock_istep --- scheduling/BTLtoRTLproof.v | 32 +++++++++++++++++++++++++++++++- 1 file changed, 31 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 31652697..8df8ef1e 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -179,7 +179,37 @@ Proof. eapply plus_trans; eauto. - inv MIB. destruct b; exploit IHIBIS; eauto. -Admitted. + + (* taking ifso *) + destruct ofin; simpl. + * (* ofin is Some final *) + intros (isfst' & pc1 & MI & STAR). + repeat eexists; eauto. + eapply css_plus_trans. + eapply plus_one. eapply exec_Icond; eauto. + eauto. + * (* ofin is None *) + intros (pc1 & OPC & PLUS). inv OPC. + inv H10; eexists; split; eauto. + all: + eapply plus_trans; + [ eapply plus_one; eapply exec_Icond; eauto + | eauto | eauto ]. + + (* taking ifnot *) + destruct ofin; simpl. + * (* ofin is Some final *) + intros (isfst' & pc1 & MI & STAR). + repeat eexists; eauto. + eapply css_plus_trans. + eapply plus_one. eapply exec_Icond; eauto. + eauto. + * (* ofin is None *) + intros (pc1 & OPC & PLUS). inv OPC. + inv H10; eexists; split; eauto. + all: + eapply plus_trans; + [ eapply plus_one; eapply exec_Icond; eauto + | eauto | eauto ]. +Qed. Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s (STACKS: list_forall2 match_stackframes stack stack') -- cgit From 6e7822b5313eaa01478530a5c47cc0d8f3776681 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 11:55:23 +0200 Subject: Some try in step_simulation --- scheduling/BTLtoRTLproof.v | 23 ++++++++++++++++++++--- 1 file changed, 20 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 8df8ef1e..bfc09d65 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -131,7 +131,6 @@ Proof. eapply plus_trans; eauto. Qed. -(* suggestion *) Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): forall pc0 opc isfst @@ -220,7 +219,26 @@ Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fi :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. - simpl. + simpl. intros (isfst' & pc1 & MI & STAR). inv MI. + - (* Final instruction except Bgoto *) + inv FS. + + (* absurd case *) inv H3. + + inv H3. inv STAR. + * + eexists; split. eapply plus_one. + eapply exec_Ireturn; eauto. + erewrite <- preserv_fnstacksize; eauto. + econstructor; eauto. + * inv H. + admit. + + admit. + + admit. + + admit. + + admit. + - (* Bgoto *) + inv STAR; try discriminate. inv FS. + eexists; split. eauto. + econstructor; eauto. Admitted. (* BROUILLON @@ -328,7 +346,6 @@ Proof. + eapply plus_one. constructor. + inv H1. econstructor; eauto. Qed. -(* Admitted. *) Theorem transf_program_correct: forward_simulation (semantics prog) (RTL.semantics tprog). -- cgit From f592b3b46908a97e2181f270cf631e9f8d60d726 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 12:37:54 +0200 Subject: one example case on main proof --- scheduling/BTLtoRTLproof.v | 38 ++++++++++++++++++++------------------ 1 file changed, 20 insertions(+), 18 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index bfc09d65..94ec84bf 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -219,26 +219,28 @@ Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fi :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. - simpl. intros (isfst' & pc1 & MI & STAR). inv MI. - - (* Final instruction except Bgoto *) - inv FS. - + (* absurd case *) inv H3. - + inv H3. inv STAR. - * - eexists; split. eapply plus_one. - eapply exec_Ireturn; eauto. + simpl. intros (isfst' & pc1 & MI & STAR). + induction FS. + - inv MI. + + inv STAR; inv H3. + + inv STAR; try discriminate. + eexists; split. eauto. + econstructor; eauto. + - inv MI. + inv STAR. + + inv H5. eexists; split. + eapply plus_one. eapply exec_Ireturn; eauto. erewrite <- preserv_fnstacksize; eauto. econstructor; eauto. - * inv H. - admit. - + admit. - + admit. - + admit. - + admit. - - (* Bgoto *) - inv STAR; try discriminate. inv FS. - eexists; split. eauto. - econstructor; eauto. + + inv H5. eexists; split. + eapply plus_trans. eauto. + eapply plus_one. eapply exec_Ireturn; eauto. + erewrite <- preserv_fnstacksize; eauto. eauto. + econstructor; eauto. + - admit. + - admit. + - admit. + - admit. Admitted. (* BROUILLON -- cgit From a345b9ff0ff90a3b90b20c13c472d52361ddcea6 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 13:06:52 +0200 Subject: builtin case OK, call and tailcall started but not finished --- scheduling/BTLtoRTLproof.v | 30 +++++++++++++++++++++++++++++- 1 file changed, 29 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 94ec84bf..4595eb83 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -226,6 +226,7 @@ Proof. + inv STAR; try discriminate. eexists; split. eauto. econstructor; eauto. + - inv MI. inv STAR. + inv H5. eexists; split. @@ -237,8 +238,35 @@ Proof. eapply plus_one. eapply exec_Ireturn; eauto. erewrite <- preserv_fnstacksize; eauto. eauto. econstructor; eauto. + - inv MI. + inv STAR. + + inv H5. + pose symbols_preserved as SYMPRES. + destruct ros. + * simpl in H. apply functions_translated in H. + destruct H as (tf & cunit & TFUN & GFIND & LO). + eexists; split. + eapply plus_one. eapply exec_Icall; eauto. + apply function_sig_translated. assumption. + admit. + (* TODO ?? *) + * admit. + + admit. - admit. - - admit. + - inv MI. pose symbols_preserved as SYMPRES. + inv STAR. + + inv H5. eexists; split. + eapply plus_one. eapply exec_Ibuiltin; eauto. + eapply eval_builtin_args_preserved; eauto. + eapply external_call_symbols_preserved; eauto. eapply senv_preserved. + econstructor; eauto. + + inv H5. eexists; split. + eapply plus_trans. eauto. + eapply plus_one. eapply exec_Ibuiltin; eauto. + eapply eval_builtin_args_preserved; eauto. + eapply external_call_symbols_preserved; eauto. eapply senv_preserved. + eauto. econstructor; eauto. +(* Icond *) - admit. - admit. Admitted. -- cgit From e02c100fac830f48a6b5c1a23b26b08d01c54c12 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 4 May 2021 14:17:06 +0200 Subject: suggestions... --- scheduling/BTLtoRTLproof.v | 90 +++++++++++++++++++++++++++++++++++----------- 1 file changed, 70 insertions(+), 20 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 4595eb83..8dad43c7 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -131,6 +131,10 @@ Proof. eapply plus_trans; eauto. Qed. + +Local Hint Constructors RTL.step match_states: core. +Local Hint Resolve css_refl plus_one: core. + Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): forall pc0 opc isfst @@ -148,9 +152,9 @@ Proof. - assert (X: opc = None). { inv MIB; auto. } subst. repeat eexists; eauto. - eapply css_refl; auto. - - inv MIB. exists pc'; split; auto. - apply plus_one. apply exec_Inop; trivial. + (* eapply css_refl; auto. *) + - inv MIB. exists pc'; split; eauto. + (* apply plus_one; eauto. apply exec_Inop; trivial. *) - inv MIB. exists pc'; split; auto. apply plus_one. eapply exec_Iop; eauto. erewrite <- eval_operation_preserved; eauto. @@ -163,8 +167,8 @@ Proof. apply plus_one. eapply exec_Istore; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. - - inv MIB. - exploit IHIBIS; eauto. + - inv MIB; eauto. + (* exploit IHIBIS; eauto. *) - inv MIB. exploit IHIBIS1; eauto. intros (pc1 & EQpc1 & STEP1); inv EQpc1. @@ -177,39 +181,66 @@ Proof. eexists; split; auto. eapply plus_trans; eauto. - inv MIB. + assert (JOIN: is_join_opt opc1 opc2 opc). { exact H10. } clear H10. destruct b; exploit IHIBIS; eauto. + (* taking ifso *) destruct ofin; simpl. * (* ofin is Some final *) intros (isfst' & pc1 & MI & STAR). repeat eexists; eauto. - eapply css_plus_trans. - eapply plus_one. eapply exec_Icond; eauto. - eauto. + eapply css_plus_trans; eauto. + (* eapply plus_one. eapply exec_Icond; eauto. + eauto. *) * (* ofin is None *) intros (pc1 & OPC & PLUS). inv OPC. - inv H10; eexists; split; eauto. + inv JOIN; eexists; split; eauto. all: - eapply plus_trans; + eapply plus_trans; eauto. + (* [ eapply plus_one; eapply exec_Icond; eauto - | eauto | eauto ]. + | eauto | eauto ].*) + (* taking ifnot *) destruct ofin; simpl. * (* ofin is Some final *) intros (isfst' & pc1 & MI & STAR). repeat eexists; eauto. - eapply css_plus_trans. - eapply plus_one. eapply exec_Icond; eauto. - eauto. + eapply css_plus_trans; eauto. + (* eapply plus_one. eapply exec_Icond; eauto. + eauto. *) * (* ofin is None *) - intros (pc1 & OPC & PLUS). inv OPC. - inv H10; eexists; split; eauto. + intros (pc1 & OPC & PLUS). subst. + inv JOIN; eexists; split; eauto. all: - eapply plus_trans; + eapply plus_trans; eauto. +(* [ eapply plus_one; eapply exec_Icond; eauto - | eauto | eauto ]. + | eauto | eauto ]. *) Qed. +Lemma css_star P s0 s1 t: + cond_star_step P s0 t s1 -> + star RTL.step tge s0 t s1. +Proof. + destruct 1. + - eapply star_refl; eauto. + - eapply plus_star; eauto. +Qed. + +Lemma final_simu_except_goto sp dupmap stack stack' f f' rs0 m0 rs1 m1 pc0 fin t s + (STACKS : list_forall2 match_stackframes stack stack') + (TRANSF : match_function dupmap f f') + (FS : final_step ge stack f sp rs1 m1 fin t s) + (pc1 : node) + (STAR : star RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 + (RTL.State stack' f' sp pc1 rs1 m1)) + (i : instruction) + (ATpc1 : (RTL.fn_code f') ! pc1 = Some i) + (MF : match_final_inst dupmap fin i) + : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. +Proof. + inv MF; inv FS. +Admitted. + Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s (STACKS: list_forall2 match_stackframes stack stack') (TRANSF: match_function dupmap f f') @@ -217,6 +248,27 @@ Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fi (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) (FS : final_step ge stack f sp rs1 m1 fin t s) :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. +Proof. + intros; exploit iblock_istep_simulation; eauto. + simpl. intros (isfst' & pc1 & MI & STAR). clear IBIS MIB. + inv MI. + - (* final inst except goto *) + exploit final_simu_except_goto; eauto. + eapply css_star; eauto. + - (* goto *) + inv FS. + inv STAR; try congruence. + eexists; split. eauto. + econstructor; eauto. +Qed. + +Lemma iblock_step_simulation_draft sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s + (STACKS: list_forall2 match_stackframes stack stack') + (TRANSF: match_function dupmap f f') + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) + (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) + (FS : final_step ge stack f sp rs1 m1 fin t s) + :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. simpl. intros (isfst' & pc1 & MI & STAR). @@ -266,8 +318,6 @@ Proof. eapply eval_builtin_args_preserved; eauto. eapply external_call_symbols_preserved; eauto. eapply senv_preserved. eauto. econstructor; eauto. -(* Icond *) - - admit. - admit. Admitted. -- cgit From fa16ab84c204bdc9470c979c3a27fc45cfb012ba Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 17:31:54 +0200 Subject: some advance and simplifications --- scheduling/BTLtoRTLproof.v | 138 +++++++++++++++++++++++---------------------- 1 file changed, 71 insertions(+), 67 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 8dad43c7..f23bfcf8 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,6 +106,25 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +(* TODO copied from Duplicateproof.v *) +(*Lemma list_nth_z_dupmap: + forall dupmap ln ln' (pc pc': node) val, + list_nth_z ln val = Some pc' -> + list_forall2 (fun n n' => dupmap!n' = Some n) ln ln' -> + exists pc', + list_nth_z ln' val = Some pc' + /\ dupmap!pc' = Some pc. +Proof. +Admitted. + induction ln; intros until val; intros LNZ LFA. + - inv LNZ. + - inv LNZ. destruct (zeq val 0) eqn:ZEQ. + + inv H0. destruct ln'; inv LFA. + simpl. exists p. split; auto. + + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto. + intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. + Qed.*) + (* FIXME: unused ? Definition is_goto (i: final): option exit := match i with @@ -131,9 +150,8 @@ Proof. eapply plus_trans; eauto. Qed. - Local Hint Constructors RTL.step match_states: core. -Local Hint Resolve css_refl plus_one: core. +Local Hint Resolve css_refl plus_one transf_fundef_correct: core. Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): @@ -143,33 +161,37 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin match ofin with | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => - (* A CHANGER ? *) + (* XXX A CHANGER ? *) exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) end. Proof. induction IBIS; simpl; intros. - - assert (X: opc = None). { inv MIB; auto. } + - (* exec_final *) + assert (X: opc = None). { inv MIB; auto. } subst. repeat eexists; eauto. - (* eapply css_refl; auto. *) - - inv MIB. exists pc'; split; eauto. - (* apply plus_one; eauto. apply exec_Inop; trivial. *) - - inv MIB. exists pc'; split; auto. + - (* exec_nop *) + inv MIB. exists pc'; split; eauto. + - (* exec_op *) + inv MIB. exists pc'; split; auto. apply plus_one. eapply exec_Iop; eauto. erewrite <- eval_operation_preserved; eauto. intros; rewrite symbols_preserved; trivial. - - inv MIB. exists pc'; split; auto. + - (* exec_load *) + inv MIB. exists pc'; split; auto. apply plus_one. eapply exec_Iload; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. - - inv MIB. exists pc'; split; auto. + - (* exec_store *) + inv MIB. exists pc'; split; auto. apply plus_one. eapply exec_Istore; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. - - inv MIB; eauto. - (* exploit IHIBIS; eauto. *) - - inv MIB. + - (* exec_seq_stop *) + inv MIB; eauto. + - (* exec_seq_continue *) + inv MIB. exploit IHIBIS1; eauto. intros (pc1 & EQpc1 & STEP1); inv EQpc1. exploit IHIBIS2; eauto. @@ -180,7 +202,8 @@ Proof. + intros (pc2 & EQpc2 & STEP2); inv EQpc2. eexists; split; auto. eapply plus_trans; eauto. - - inv MIB. + - (* exec_cond *) + inv MIB. assert (JOIN: is_join_opt opc1 opc2 opc). { exact H10. } clear H10. destruct b; exploit IHIBIS; eauto. + (* taking ifso *) @@ -189,34 +212,25 @@ Proof. intros (isfst' & pc1 & MI & STAR). repeat eexists; eauto. eapply css_plus_trans; eauto. - (* eapply plus_one. eapply exec_Icond; eauto. - eauto. *) * (* ofin is None *) intros (pc1 & OPC & PLUS). inv OPC. inv JOIN; eexists; split; eauto. all: eapply plus_trans; eauto. - (* - [ eapply plus_one; eapply exec_Icond; eauto - | eauto | eauto ].*) + (* taking ifnot *) destruct ofin; simpl. * (* ofin is Some final *) intros (isfst' & pc1 & MI & STAR). repeat eexists; eauto. eapply css_plus_trans; eauto. - (* eapply plus_one. eapply exec_Icond; eauto. - eauto. *) * (* ofin is None *) intros (pc1 & OPC & PLUS). subst. inv JOIN; eexists; split; eauto. all: eapply plus_trans; eauto. -(* - [ eapply plus_one; eapply exec_Icond; eauto - | eauto | eauto ]. *) Qed. +(* TODO move above *) Lemma css_star P s0 s1 t: cond_star_step P s0 t s1 -> star RTL.step tge s0 t s1. @@ -226,19 +240,35 @@ Proof. - eapply plus_star; eauto. Qed. -Lemma final_simu_except_goto sp dupmap stack stack' f f' rs0 m0 rs1 m1 pc0 fin t s +Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 pc1 fin t s (STACKS : list_forall2 match_stackframes stack stack') (TRANSF : match_function dupmap f f') (FS : final_step ge stack f sp rs1 m1 fin t s) - (pc1 : node) - (STAR : star RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 - (RTL.State stack' f' sp pc1 rs1 m1)) (i : instruction) (ATpc1 : (RTL.fn_code f') ! pc1 = Some i) (MF : match_final_inst dupmap fin i) - : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. + : exists s', RTL.step tge (RTL.State stack' f' sp pc1 rs1 m1) t s' /\ match_states s s'. Proof. inv MF; inv FS. + - (* return *) + eexists; split. + eapply exec_Ireturn; eauto. + erewrite <- preserv_fnstacksize; eauto. + econstructor; eauto. + - (* call *) + pose symbols_preserved as SYMPRES. + rename H7 into FIND. + destruct ri. + + simpl in FIND; apply functions_translated in FIND. + destruct FIND as (tf & cunit & TFUN & GFIND & LO). + eexists; split. eapply exec_Icall; eauto. + apply function_sig_translated. assumption. + repeat (econstructor; eauto). + + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). + eexists; split. eapply exec_Icall; eauto. simpl. rewrite symbols_preserved. + rewrite GFS. eassumption. apply function_sig_translated. assumption. + repeat (econstructor; eauto). Admitted. Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s @@ -254,7 +284,11 @@ Proof. inv MI. - (* final inst except goto *) exploit final_simu_except_goto; eauto. - eapply css_star; eauto. + intros (s' & STEP & MS). eexists; split. + + eapply plus_right. + eapply css_star; eauto. + eapply STEP. econstructor. + + eapply MS. - (* goto *) inv FS. inv STAR; try congruence. @@ -262,6 +296,8 @@ Proof. econstructor; eauto. Qed. +(* BROUILLON + Lemma iblock_step_simulation_draft sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s (STACKS: list_forall2 match_stackframes stack stack') (TRANSF: match_function dupmap f f') @@ -279,49 +315,17 @@ Proof. eexists; split. eauto. econstructor; eauto. - - inv MI. - inv STAR. - + inv H5. eexists; split. - eapply plus_one. eapply exec_Ireturn; eauto. - erewrite <- preserv_fnstacksize; eauto. - econstructor; eauto. - + inv H5. eexists; split. - eapply plus_trans. eauto. - eapply plus_one. eapply exec_Ireturn; eauto. - erewrite <- preserv_fnstacksize; eauto. eauto. - econstructor; eauto. - - inv MI. - inv STAR. - + inv H5. - pose symbols_preserved as SYMPRES. - destruct ros. - * simpl in H. apply functions_translated in H. - destruct H as (tf & cunit & TFUN & GFIND & LO). - eexists; split. - eapply plus_one. eapply exec_Icall; eauto. - apply function_sig_translated. assumption. - admit. - (* TODO ?? *) - * admit. - + admit. - - admit. - inv MI. pose symbols_preserved as SYMPRES. inv STAR. + inv H5. eexists; split. - eapply plus_one. eapply exec_Ibuiltin; eauto. - eapply eval_builtin_args_preserved; eauto. - eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - econstructor; eauto. - + inv H5. eexists; split. - eapply plus_trans. eauto. - eapply plus_one. eapply exec_Ibuiltin; eauto. - eapply eval_builtin_args_preserved; eauto. - eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - eauto. econstructor; eauto. + eapply plus_one. eapply exec_Ijumptable; eauto. + exploit list_nth_z_dupmap; eauto. intros (pc'1 & LNZ & REVM). + eexists. split. + + eapply exec_Ijumptable; eauto. + + econstructor; eauto. - admit. Admitted. -(* BROUILLON Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := | ibis_synced opc pc1 -- cgit From 85dc46da493448a1d207ed5003ea476cd397b0d9 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 22:24:33 +0200 Subject: finishing proofs and cleanup --- scheduling/BTLtoRTLproof.v | 182 ++++++++++++--------------------------------- 1 file changed, 48 insertions(+), 134 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index f23bfcf8..da65d6ac 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -5,8 +5,8 @@ Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking BTLtoRTL. -(*****************************) -(* Put this in an other file *) +(**********************************) +(* TODO Put this in an other file *) Require Import Linking. @@ -106,32 +106,22 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -(* TODO copied from Duplicateproof.v *) -(*Lemma list_nth_z_dupmap: +(* Inspired from Duplicateproof.v *) +Lemma list_nth_z_dupmap: forall dupmap ln ln' (pc pc': node) val, - list_nth_z ln val = Some pc' -> - list_forall2 (fun n n' => dupmap!n' = Some n) ln ln' -> - exists pc', + list_nth_z ln val = Some pc -> + list_forall2 (fun n n' => dupmap!n = Some n') ln ln' -> + exists (pc': node), list_nth_z ln' val = Some pc' - /\ dupmap!pc' = Some pc. + /\ dupmap!pc = Some pc'. Proof. -Admitted. induction ln; intros until val; intros LNZ LFA. - inv LNZ. - inv LNZ. destruct (zeq val 0) eqn:ZEQ. + inv H0. destruct ln'; inv LFA. - simpl. exists p. split; auto. + simpl. exists n. split; auto. + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto. - intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. - Qed.*) - -(* FIXME: unused ? -Definition is_goto (i: final): option exit := - match i with - | Bgoto pc => Some pc - | _ => None - end. -*) +Qed. (* variant of [star RTL.step] but requiring proposition [P] on the [refl] (stutttering) case. *) Inductive cond_star_step (P: Prop): RTL.state -> trace -> RTL.state -> Prop := @@ -150,6 +140,15 @@ Proof. eapply plus_trans; eauto. Qed. +Lemma css_star P s0 s1 t: + cond_star_step P s0 t s1 -> + star RTL.step tge s0 t s1. +Proof. + destruct 1. + - eapply star_refl; eauto. + - eapply plus_star; eauto. +Qed. + Local Hint Constructors RTL.step match_states: core. Local Hint Resolve css_refl plus_one transf_fundef_correct: core. @@ -230,16 +229,6 @@ Proof. eapply plus_trans; eauto. Qed. -(* TODO move above *) -Lemma css_star P s0 s1 t: - cond_star_step P s0 t s1 -> - star RTL.step tge s0 t s1. -Proof. - destruct 1. - - eapply star_refl; eauto. - - eapply plus_star; eauto. -Qed. - Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 pc1 fin t s (STACKS : list_forall2 match_stackframes stack stack') (TRANSF : match_function dupmap f f') @@ -269,7 +258,35 @@ Proof. eexists; split. eapply exec_Icall; eauto. simpl. rewrite symbols_preserved. rewrite GFS. eassumption. apply function_sig_translated. assumption. repeat (econstructor; eauto). -Admitted. + - (* tailcall *) + pose symbols_preserved as SYMPRES. + rename H2 into FIND. + destruct ri. + + simpl in FIND; apply functions_translated in FIND. + destruct FIND as (tf & cunit & TFUN & GFIND & LO). + eexists; split. eapply exec_Itailcall; eauto. + apply function_sig_translated. assumption. + erewrite <- preserv_fnstacksize; eauto. + repeat (econstructor; eauto). + + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). + eexists; split. eapply exec_Itailcall; eauto. simpl. rewrite symbols_preserved. + rewrite GFS. eassumption. apply function_sig_translated. assumption. + erewrite <- preserv_fnstacksize; eauto. + repeat (econstructor; eauto). + - (* builtin *) + pose symbols_preserved as SYMPRES. + eexists. split. + eapply exec_Ibuiltin; eauto. eapply eval_builtin_args_preserved; eauto. + eapply external_call_symbols_preserved; eauto. eapply senv_preserved. + econstructor; eauto. + - (* jumptable *) + pose symbols_preserved as SYMPRES. + exploit list_nth_z_dupmap; eauto. intros (pc'1 & LNZ & REVM). + eexists. split. + eapply exec_Ijumptable; eauto. + econstructor; eauto. +Qed. Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s (STACKS: list_forall2 match_stackframes stack stack') @@ -296,109 +313,6 @@ Proof. econstructor; eauto. Qed. -(* BROUILLON - -Lemma iblock_step_simulation_draft sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s - (STACKS: list_forall2 match_stackframes stack stack') - (TRANSF: match_function dupmap f f') - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) - (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) - (FS : final_step ge stack f sp rs1 m1 fin t s) - :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. -Proof. - intros; exploit iblock_istep_simulation; eauto. - simpl. intros (isfst' & pc1 & MI & STAR). - induction FS. - - inv MI. - + inv STAR; inv H3. - + inv STAR; try discriminate. - eexists; split. eauto. - econstructor; eauto. - - - inv MI. pose symbols_preserved as SYMPRES. - inv STAR. - + inv H5. eexists; split. - eapply plus_one. eapply exec_Ijumptable; eauto. - exploit list_nth_z_dupmap; eauto. intros (pc'1 & LNZ & REVM). - eexists. split. - + eapply exec_Ijumptable; eauto. - + econstructor; eauto. - - admit. -Admitted. - - -Inductive iblock_istep_gen sp dupmap stack f pc0 cfg ib: trace -> bool -> option final -> option node -> Prop := - | ibis_synced opc pc1 - (HOPC: opc = Some pc1) - (MIB : match_iblock dupmap cfg true pc0 ib opc) - : iblock_istep_gen sp dupmap stack f pc0 cfg ib E0 true None opc - | ibis_stutter rs1 m1 fin ofin s t - (HFIN: ofin = Some fin) - (MIB : match_iblock dupmap cfg false pc0 ib None) - (FS : final_step ge stack f sp rs1 m1 fin t s) - : iblock_istep_gen sp dupmap stack f pc0 cfg ib t false ofin None. - -Lemma iblock_step_simulation - sp dupmap stack stack' f f' rs0 m0 rs1 m1 ib ofin pc0 opc t s isfst - (STACKS: list_forall2 match_stackframes stack stack') - (TRANSF: match_function dupmap f f') - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin) - (*(MIB : match_iblock dupmap cfg' isfst pc0 ib opc)*) - (*(FS : final_step ge stack f sp rs1 m1 fin t s)*) -(*FNC : (fn_code f) ! pc = Some ib*) - (IBGEN: iblock_istep_gen sp dupmap stack f pc0 (RTL.fn_code f') ib t isfst ofin opc) - :exists s', (if isfst then plus RTL.step else star RTL.step) tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. -Proof. - (* TODO: généraliser ce lemme pour pouvoir le prouver par induction sur IBIS: - => il faut en particulier généraliser l'hypothèse MIB qui relie le iblock_istep ib en cours d'exécution. - avec le code RTL à partir de pc0. - => ici, le "isfst" a déjà été généralisé: quand il vaut "false", ça veut dire qu'on a le droit de faire du "stuttering" côté RTL. - => il reste à comprendre comment généraliser le "None" en "opc" - ainsi que l'hypothese OFIN pour autoriser le cas "ofin=None" (nécessaire pour l'induction). - Idée: si "ofin = None" alors il y a un pc1 tq "opc = Some pc1" qui permet d'exécuter la suite du bloc... - - Au final, il faut sans doute introduire un "Inductive" pour capturer ces idées dans un prédicat "lisible"/"manipulable"... - *) - (* XXX keep IBIS, MIB, and FS ? *) - induction IBIS. - - admit. - - inv IBGEN; try_simplify_someHyps. - inv MIB. eexists; split. - apply plus_one. econstructor; eauto. - (* - - inversion IBGEN; subst; try_simplify_someHyps. - exploit iblock_istep_simulation; eauto. - inv MIB. - eexists. split. apply plus_one. - eapply exec_Inop; eauto. - try_simplify_someHyps. - admit. - - (* exec_op *) - admit. (* cas absurde car hypothese OFIN trop restrictive *) - - (* exec_load_TRAP *) - admit. (* cas absurde car hypothese OFIN trop restrictive *) - - (* exec_load_store *) - admit. (* cas absurde car hypothese OFIN trop restrictive *) - - (* exec_seq_stop *) - (*inv MIB.*) - eapply IHIBIS; eauto. - (* TODO: c'est ici qu'on voit que l'hypothèse MIB est trop restrictive actuellement - normalement, l'hypothèse d'induction IHIBIS devrait permettre de conclure quasi-directement ici. - *) - admit. - - (* exec_seq_continue *) - (* TODO: ici l'hypothèse d'induction IHIBIS1 n'est pas utilisable à cause de OFIN trop restrictive *) - try_simplify_someHyps. - (*inv MIB.*) - admit. - - (* exec_cond *) - try_simplify_someHyps. - (*inv MIB.*) - admit.*) -Admitted. - -*) - Theorem plus_simulation s1 t s1': step ge s1 t s1' -> forall s2 (MS: match_states s1 s2), -- cgit From bf2ed66aa8c87b95c1d57cdd7f22dc131c7728cf Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 4 May 2021 23:08:53 +0200 Subject: Factorize as suggested for call/tailcall --- scheduling/BTLtoRTLproof.v | 86 +++++++++++++++++++++++----------------------- 1 file changed, 43 insertions(+), 43 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index da65d6ac..67a01c20 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -106,6 +106,22 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. +Lemma find_function_preserved ri rs0 fd + (FIND : find_function ge ri rs0 = Some fd) + : exists fd', RTL.find_function tge ri rs0 = Some fd' + /\ transf_fundef fd = OK fd'. +Proof. + pose symbols_preserved as SYMPRES. + destruct ri. + + simpl in FIND; apply functions_translated in FIND. + destruct FIND as (tf & cunit & TFUN & GFIND & LO). + eexists; split. eauto. assumption. + + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). + eexists; split. simpl. rewrite symbols_preserved. + rewrite GFS. eassumption. assumption. +Qed. + (* Inspired from Duplicateproof.v *) Lemma list_nth_z_dupmap: forall dupmap ln ln' (pc pc': node) val, @@ -153,17 +169,16 @@ Local Hint Constructors RTL.step match_states: core. Local Hint Resolve css_refl plus_one transf_fundef_correct: core. Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): - forall pc0 opc isfst - (MIB: match_iblock dupmap (RTL.fn_code f') isfst pc0 ib opc) - , - match ofin with - | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) - | Some fin => - (* XXX A CHANGER ? *) - exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None - /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) - end. + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin): + forall pc0 opc isfst + (MIB: match_iblock dupmap (RTL.fn_code f') isfst pc0 ib opc), + match ofin with + | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) + | Some fin => + (* XXX A CHANGER ? *) + exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None + /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) + end. Proof. induction IBIS; simpl; intros. - (* exec_final *) @@ -245,35 +260,20 @@ Proof. erewrite <- preserv_fnstacksize; eauto. econstructor; eauto. - (* call *) - pose symbols_preserved as SYMPRES. rename H7 into FIND. - destruct ri. - + simpl in FIND; apply functions_translated in FIND. - destruct FIND as (tf & cunit & TFUN & GFIND & LO). - eexists; split. eapply exec_Icall; eauto. - apply function_sig_translated. assumption. - repeat (econstructor; eauto). - + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. - apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). - eexists; split. eapply exec_Icall; eauto. simpl. rewrite symbols_preserved. - rewrite GFS. eassumption. apply function_sig_translated. assumption. - repeat (econstructor; eauto). + exploit find_function_preserved; eauto. + intros (fd' & FIND' & TRANSFU). + eexists; split. eapply exec_Icall; eauto. + apply function_sig_translated. assumption. + repeat (econstructor; eauto). - (* tailcall *) - pose symbols_preserved as SYMPRES. rename H2 into FIND. - destruct ri. - + simpl in FIND; apply functions_translated in FIND. - destruct FIND as (tf & cunit & TFUN & GFIND & LO). - eexists; split. eapply exec_Itailcall; eauto. - apply function_sig_translated. assumption. - erewrite <- preserv_fnstacksize; eauto. - repeat (econstructor; eauto). - + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. - apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). - eexists; split. eapply exec_Itailcall; eauto. simpl. rewrite symbols_preserved. - rewrite GFS. eassumption. apply function_sig_translated. assumption. - erewrite <- preserv_fnstacksize; eauto. - repeat (econstructor; eauto). + exploit find_function_preserved; eauto. + intros (fd' & FIND' & TRANSFU). + eexists; split. eapply exec_Itailcall; eauto. + apply function_sig_translated. assumption. + erewrite <- preserv_fnstacksize; eauto. + repeat (econstructor; eauto). - (* builtin *) pose symbols_preserved as SYMPRES. eexists. split. @@ -289,12 +289,12 @@ Proof. Qed. Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s - (STACKS: list_forall2 match_stackframes stack stack') - (TRANSF: match_function dupmap f f') - (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) - (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) - (FS : final_step ge stack f sp rs1 m1 fin t s) - :exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. + (STACKS: list_forall2 match_stackframes stack stack') + (TRANSF: match_function dupmap f f') + (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) + (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) + (FS : final_step ge stack f sp rs1 m1 fin t s) + : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. simpl. intros (isfst' & pc1 & MI & STAR). clear IBIS MIB. -- cgit From f427f9b257baee8a962a4b3e6c1e3c1804fd65da Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 5 May 2021 08:51:51 +0200 Subject: clean deprecated comments --- scheduling/BTLtoRTLproof.v | 7 +------ 1 file changed, 1 insertion(+), 6 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 67a01c20..3944e756 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -4,10 +4,6 @@ Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking BTLtoRTL. - -(**********************************) -(* TODO Put this in an other file *) - Require Import Linking. Lemma transf_function_correct f f': @@ -175,7 +171,6 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin match ofin with | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => - (* XXX A CHANGER ? *) exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) end. @@ -218,7 +213,7 @@ Proof. eapply plus_trans; eauto. - (* exec_cond *) inv MIB. - assert (JOIN: is_join_opt opc1 opc2 opc). { exact H10. } clear H10. + rename H10 into JOIN. (* is_join_opt opc1 opc2 opc *) destruct b; exploit IHIBIS; eauto. + (* taking ifso *) destruct ofin; simpl. -- cgit From e9f8fa5f0635f3c1af489bcf084d38f89ea58b13 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 5 May 2021 14:03:04 +0200 Subject: advance in cfg checker --- scheduling/BTL.v | 139 ++++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 138 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 28941e01..82a54061 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -513,11 +513,148 @@ Proof. Qed. (* TODO: deuxième étape (après BTLtoRTLproof.plus_simulation) *) -Parameter verify_block: PTree.t node -> RTL.code -> bool -> node -> iblock -> res (option node). +Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := + match ib with + | BF fi => + match fi with + | Bgoto pc => + match dupmap!pc with + | Some _ => OK None + | None => Error (msg "verify_block: incorrect dupmap Bgoto") + end + | _ => + match cfg!pc with + | Some i => Error (msg "TODO") (* TODO *) + | None => Error (msg "verify_block: incorrect cfg BF") + end + end + | Bnop => + match cfg!pc with + | Some (Inop pc') => OK (Some pc') + | _ => Error (msg "verify_block: incorrect cfg Bnop") + end + | Bop op lr r => + match cfg!pc with + | Some (Iop op' lr' r' pc') => + if (eq_operation op op') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then + OK (Some pc') + else Error (msg "verify_block: different r in Bop") + else Error (msg "verify_block: different lr in Bop") + else Error (msg "verify_block: different operations in Bop") + | _ => Error (msg "verify_block: incorrect cfg Bop") + end + | Bload tm m a lr r => + match cfg!pc with + | Some (Iload tm' m' a' lr' r' pc') => + if (trapping_mode_eq tm tm') then + if (chunk_eq m m') then + if (eq_addressing a a') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then + OK (Some pc') + else Error (msg "verify_block: different r in Bload") + else Error (msg "verify_block: different lr in Bload") + else Error (msg "verify_block: different addressing in Bload") + else Error (msg "verify_block: different mchunk in Bload") + else Error (msg "verify_block: different trapping_mode in Bload") + | _ => Error (msg "verify_block: incorrect cfg Bload") + end + | Bstore m a lr r => + match cfg!pc with + | Some (Istore m' a' lr' r' pc') => + if (chunk_eq m m') then + if (eq_addressing a a') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then OK (Some pc') + else Error (msg "verify_block: different r in Bstore") + else Error (msg "verify_block: different lr in Bstore") + else Error (msg "verify_block: different addressing in Bstore") + else Error (msg "verify_block: different mchunk in Bstore") + | _ => Error (msg "verify_block: incorrect cfg Bstore") + end + | Bseq b1 b2 => + do opc <- verify_block dupmap cfg isfst pc b1; + match opc with + | Some pc' => + verify_block dupmap cfg false pc' b2 + | None => Error (msg "verify_block: None next pc in Bseq (deadcode)") + end + | 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 + if (eq_condition c c') then + do opc1 <- verify_block dupmap cfg false pcso bso; + do opc2 <- verify_block dupmap cfg false pcnot bnot; + match opc1, opc2 with + | None, o => OK o + | o, None => OK o + | Some x, Some x' => + if Pos.eq_dec x x' then OK (Some x) + else Error (msg "verify_block: is_join_opt incorrect for Bcond") + end + else Error (msg "verify_block: incompatible conditions in Bcond") + else Error (msg "verify_block: different lr in Bcond") + | _ => Error (msg "verify_block: incorrect cfg Bcond") + end + end. (* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *) Lemma verify_block_correct: forall dupmap cfg isfst pc ib fin, verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin. +Proof. + induction ib; intros; + try (unfold verify_block in H; destruct (cfg!pc) eqn:EQCFG; [ idtac | discriminate; fail ]). + - admit. + - (* Bnop *) + destruct i; inv H. + constructor; assumption. + - (* Bop *) + destruct i; try discriminate. + destruct (eq_operation _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bload *) + destruct i; try discriminate. + destruct (trapping_mode_eq _ _); try discriminate. + destruct (chunk_eq _ _); try discriminate. + destruct (eq_addressing _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bstore *) + destruct i; try discriminate. + destruct (chunk_eq _ _); try discriminate. + destruct (eq_addressing _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bseq *) + monadInv H. + destruct x; try discriminate. + eapply mib_seq_Some. + eapply IHib1; eauto. admit. + (*eapply IHib2; eauto.*) + (*destruct (cfg!pc) eqn:ECFG; try discriminate.*) + (*destruct i; try discriminate.*) + - (* Bcond *) + destruct i; try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (eq_condition _ _); try discriminate. + fold (verify_block dupmap cfg false n ib1) in H. + fold (verify_block dupmap cfg false n0 ib2) in H. + monadInv H. + destruct x, x0; try destruct (Pos.eq_dec _ _); try discriminate. + + inv EQ2. eapply mib_cond; eauto. + admit. + admit. + admit. + + admit. + + admit. + + admit. Admitted. Local Hint Resolve verify_block_correct: core. -- cgit From 7473fed7c8e1b2fdef276b7aa754fb00792d47ca Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 5 May 2021 19:59:47 +0200 Subject: finish verify_block and proof --- scheduling/BTL.v | 144 +++++++++++++++++++++++++++++++++++++++++++++++-------- 1 file changed, 123 insertions(+), 21 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 82a54061..7bc091ef 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -512,20 +512,91 @@ Proof. - intros. destruct ln'; monadInv H. constructor; eauto. Qed. +(* TODO Copied from duplicate, should we import ? *) +Lemma product_eq {A B: Type} : + (forall (a b: A), {a=b} + {a<>b}) -> + (forall (c d: B), {c=d} + {c<>d}) -> + forall (x y: A+B), {x=y} + {x<>y}. +Proof. + intros H H'. intros. decide equality. +Qed. + +(* TODO Copied from duplicate, should we import ? *) +(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application" + * error when doing so *) +Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}. +Proof. + intros. + apply (builtin_arg_eq Pos.eq_dec). +Defined. +Global Opaque builtin_arg_eq_pos. + +(* TODO Copied from duplicate, should we import ? *) +Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}. +Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed. +Global Opaque builtin_res_eq_pos. + (* TODO: deuxième étape (après BTLtoRTLproof.plus_simulation) *) Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := match ib with | BF fi => match fi with - | Bgoto pc => - match dupmap!pc with - | Some _ => OK None - | None => Error (msg "verify_block: incorrect dupmap Bgoto") + | Bgoto pc1 => + do u <- verify_is_copy dupmap pc1 pc; + if negb isfst then + OK None + else Error (msg "verify_block: isfst is true Bgoto") + | Breturn or => + match cfg!pc with + | Some (Ireturn or') => + if option_eq Pos.eq_dec or or' then OK None + else Error (msg "verify_block: different opt reg in Breturn") + | _ => Error (msg "verify_block: incorrect cfg Breturn") end - | _ => + | Bcall s ri lr r pc1 => match cfg!pc with - | Some i => Error (msg "TODO") (* TODO *) - | None => Error (msg "verify_block: incorrect cfg BF") + | Some (Icall s' ri' lr' r' pc2) => + do u <- verify_is_copy dupmap pc1 pc2; + if (signature_eq s s') then + if (product_eq Pos.eq_dec ident_eq ri ri') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then OK None + else Error (msg "verify_block: different r r' in Bcall") + else Error (msg "verify_block: different lr in Bcall") + else Error (msg "verify_block: different ri in Bcall") + else Error (msg "verify_block: different signatures in Bcall") + | _ => Error (msg "verify_block: incorrect cfg Bcall") + end + | Btailcall s ri lr => + match cfg!pc with + | Some (Itailcall s' ri' lr') => + if (signature_eq s s') then + if (product_eq Pos.eq_dec ident_eq ri ri') then + if (list_eq_dec Pos.eq_dec lr lr') then OK None + else Error (msg "verify_block: different lr in Btailcall") + else Error (msg "verify_block: different ri in Btailcall") + else Error (msg "verify_block: different signatures in Btailcall") + | _ => Error (msg "verify_block: incorrect cfg Btailcall") + end + | Bbuiltin ef la br pc1 => + match cfg!pc with + | Some (Ibuiltin ef' la' br' pc2) => + do u <- verify_is_copy dupmap pc1 pc2; + if (external_function_eq ef ef') then + if (list_eq_dec builtin_arg_eq_pos la la') then + if (builtin_res_eq_pos br br') then OK None + else Error (msg "verify_block: different brr in Bbuiltin") + else Error (msg "verify_block: different lbar in Bbuiltin") + else Error (msg "verify_block: different ef in Bbuiltin") + | _ => Error (msg "verify_block: incorrect cfg Bbuiltin") + end + | Bjumptable r ln => + match cfg!pc with + | Some (Ijumptable r' ln') => + do u <- verify_is_copy_list dupmap ln ln'; + if (Pos.eq_dec r r') then OK None + else Error (msg "verify_block: different r in Bjumptable") + | _ => Error (msg "verify_block: incorrect cfg Bjumptable") end end | Bnop => @@ -602,12 +673,51 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) end. (* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *) -Lemma verify_block_correct: forall dupmap cfg isfst pc ib fin, +Lemma verify_block_correct dupmap cfg ib: forall pc isfst fin, verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin. Proof. induction ib; intros; try (unfold verify_block in H; destruct (cfg!pc) eqn:EQCFG; [ idtac | discriminate; fail ]). - - admit. + - (* BF *) + destruct fi; unfold verify_block in H. + + (* Bgoto *) + monadInv H. + destruct (isfst); simpl in EQ0; inv EQ0. + eapply verify_is_copy_correct in EQ. + constructor; assumption. + + (* Breturn *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + destruct (option_eq _ _ _); try discriminate. inv H. + eapply mib_BF; eauto. constructor. + + (* Bcall *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_correct in EQ. + destruct (signature_eq _ _); try discriminate. + destruct (product_eq _ _ _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. + + (* Btailcall *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + destruct (signature_eq _ _); try discriminate. + destruct (product_eq _ _ _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. subst. + inv H. eapply mib_BF; eauto. constructor. + + (* Bbuiltin *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_correct in EQ. + destruct (external_function_eq _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (builtin_res_eq_pos _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. + + (* Bjumptable *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_list_correct in EQ. + destruct (Pos.eq_dec _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. - (* Bnop *) destruct i; inv H. constructor; assumption. @@ -636,10 +746,8 @@ Proof. monadInv H. destruct x; try discriminate. eapply mib_seq_Some. - eapply IHib1; eauto. admit. - (*eapply IHib2; eauto.*) - (*destruct (cfg!pc) eqn:ECFG; try discriminate.*) - (*destruct i; try discriminate.*) + eapply IHib1; eauto. + eapply IHib2; eauto. - (* Bcond *) destruct i; try discriminate. destruct (list_eq_dec _ _ _); try discriminate. @@ -648,14 +756,8 @@ Proof. fold (verify_block dupmap cfg false n0 ib2) in H. monadInv H. destruct x, x0; try destruct (Pos.eq_dec _ _); try discriminate. - + inv EQ2. eapply mib_cond; eauto. - admit. - admit. - admit. - + admit. - + admit. - + admit. -Admitted. + all: inv EQ2; eapply mib_cond; eauto; econstructor. +Qed. Local Hint Resolve verify_block_correct: core. Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit := -- cgit From c7bc74b1860f30df678bf384a32cd9c6eb113beb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 6 May 2021 09:07:07 +0200 Subject: start RTL -> BTL --- scheduling/BTL.v | 119 ++++++++++++---------- scheduling/BTLtoRTLproof.v | 50 ++++++++++ scheduling/RTLtoBTL.v | 23 +++++ scheduling/RTLtoBTLproof.v | 239 +++++++++++++++++++++++++++++++++++++++++++++ 4 files changed, 382 insertions(+), 49 deletions(-) create mode 100644 scheduling/RTLtoBTL.v create mode 100644 scheduling/RTLtoBTLproof.v (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 7bc091ef..fb67d150 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -342,6 +342,8 @@ Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). +(** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) + (** Matching BTL and RTL code We should be able to define a single verifier able to prove a bisimulation between BTL and RTL code. @@ -429,49 +431,10 @@ Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> i match_iblock dupmap cfg isfst pc (Bcond c lr bso bnot i) opc . -(* Partial copy of Duplicate and Duplicateproof. *) - Definition match_cfg dupmap (cfg: code) (cfg':RTL.code): Prop := forall pc pc', dupmap!pc = Some pc' -> exists ib, cfg!pc = Some ib /\ match_iblock dupmap cfg' true pc' ib.(entry) None. -Record match_function dupmap f f': Prop := { - dupmap_correct: match_cfg dupmap (fn_code f) (RTL.fn_code f'); - dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f'); - preserv_fnsig: fn_sig f = RTL.fn_sig f'; - preserv_fnparams: fn_params f = RTL.fn_params f'; - preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' -}. - -Inductive match_fundef: fundef -> RTL.fundef -> Prop := - | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') - | match_External ef: match_fundef (External ef) (External ef). - -Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := - | match_stackframe_intro - dupmap res f sp pc rs f' pc' - (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc = Some pc') - : match_stackframes (Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). - -Inductive match_states: state -> RTL.state -> Prop := - | match_states_intro - dupmap st f sp pc rs m st' f' pc' - (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc = Some pc') - : match_states (State st f sp pc rs m) (RTL.State st' f' sp pc' rs m) - | match_states_call - st st' f f' args m - (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_fundef f f') - : match_states (Callstate st f args m) (RTL.Callstate st' f' args m) - | match_states_return - st st' v m - (STACKS: list_forall2 match_stackframes st st') - : match_states (Returnstate st v m) (RTL.Returnstate st' v m) - . - (** Shared verifier between RTL -> BTL and BTL -> RTL *) Local Open Scope error_monad_scope. @@ -536,7 +499,6 @@ Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}. Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed. Global Opaque builtin_res_eq_pos. -(* TODO: deuxième étape (après BTLtoRTLproof.plus_simulation) *) Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := match ib with | BF fi => @@ -801,14 +763,73 @@ Definition verify_function dupmap f f' : res unit := do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); verify_cfg dupmap (fn_code f) (RTL.fn_code f'). -Lemma verify_function_correct dupmap f f' tt: - verify_function dupmap f f' = OK tt -> - fn_sig f = RTL.fn_sig f' -> - fn_params f = RTL.fn_params f' -> - fn_stacksize f = RTL.fn_stacksize f' -> - match_function dupmap f f'. +Definition is_seq (ib: iblock): bool := + match ib with + | Bseq _ _ => true + | _ => false + end. + +Definition is_atom (ib: iblock): bool := + match ib with + | Bseq _ _ | Bcond _ _ _ _ _ => false + | _ => true + end. + + +Inductive is_right_assoc: iblock -> Prop := + | ira_Bseq ib1 ib2: + is_seq ib1 = false -> + is_right_assoc ib1 -> + is_right_assoc ib2 -> + is_right_assoc (Bseq ib1 ib2) + | ira_Bcond cond args ib1 ib2 i: + is_right_assoc ib1 -> + is_right_assoc ib2 -> + is_right_assoc (Bcond cond args ib1 ib2 i) + | ira_others ib: + is_atom ib = true -> + is_right_assoc ib + . +Local Hint Constructors is_right_assoc: core. + +Fixpoint right_assoc (ib: iblock): iblock := + match ib with + | Bseq ib1 ib2 => rec_rotate ib1 (right_assoc ib2) + | Bcond cond args ib1 ib2 i => + Bcond cond args (right_assoc ib1) (right_assoc ib2) i + | ib => ib + end +with rec_rotate (ib: iblock) rem: iblock := + match ib with + | Bseq ib1 ib2 => + rec_rotate ib1 (rec_rotate ib2 rem) + | Bcond cond args ib1 ib2 i => + Bseq (Bcond cond args (right_assoc ib1) (right_assoc ib2) i) rem + | ib => Bseq ib rem +end. + +Lemma right_assoc_rec_rotate_correct ib: + is_right_assoc (right_assoc ib) + /\ (forall rem, is_right_assoc rem -> is_right_assoc (rec_rotate ib rem)). Proof. - unfold verify_function; intro VERIF. monadInv VERIF. - constructor; eauto. - eapply verify_cfg_correct; eauto. + induction ib; simpl; intuition. Qed. + +Lemma right_assoc_correct ib: + is_right_assoc (right_assoc ib). +Proof. + destruct (right_assoc_rec_rotate_correct ib); auto. +Qed. + +Definition right_assoc_code (cfg: code): code := + (PTree.map (fun _ ib => {| entry:=right_assoc ib.(entry); input_regs := ib.(input_regs) |}) cfg). + +Lemma right_assoc_code_correct cfg pc ib : + (right_assoc_code cfg)!pc=Some ib -> is_right_assoc (ib.(entry)). +Proof. + unfold right_assoc_code. + rewrite PTree.gmap. + destruct (cfg!pc); simpl; intros; try_simplify_someHyps. + apply right_assoc_correct; auto. +Qed. + diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 3944e756..08a77ae4 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -6,6 +6,56 @@ Require Import Errors Linking BTLtoRTL. Require Import Linking. +Record match_function dupmap f f': Prop := { + dupmap_correct: match_cfg dupmap (fn_code f) (RTL.fn_code f'); + dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f'); + preserv_fnsig: fn_sig f = RTL.fn_sig f'; + preserv_fnparams: fn_params f = RTL.fn_params f'; + preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' +}. + +Inductive match_fundef: fundef -> RTL.fundef -> Prop := + | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc = Some pc') + : match_stackframes (Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). + +Inductive match_states: state -> RTL.state -> Prop := + | match_states_intro + dupmap st f sp pc rs m st' f' pc' + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc = Some pc') + : match_states (State st f sp pc rs m) (RTL.State st' f' sp pc' rs m) + | match_states_call + st st' f f' args m + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_fundef f f') + : match_states (Callstate st f args m) (RTL.Callstate st' f' args m) + | match_states_return + st st' v m + (STACKS: list_forall2 match_stackframes st st') + : match_states (Returnstate st v m) (RTL.Returnstate st' v m) + . + +Lemma verify_function_correct dupmap f f' tt: + verify_function dupmap f f' = OK tt -> + fn_sig f = RTL.fn_sig f' -> + fn_params f = RTL.fn_params f' -> + fn_stacksize f = RTL.fn_stacksize f' -> + match_function dupmap f f'. +Proof. + unfold verify_function; intro VERIF. monadInv VERIF. + constructor; eauto. + - eapply verify_cfg_correct; eauto. + - eapply verify_is_copy_correct; eauto. +Qed. + Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. Proof. diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v new file mode 100644 index 00000000..14aecb21 --- /dev/null +++ b/scheduling/RTLtoBTL.v @@ -0,0 +1,23 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad BTL. + +Require Import Errors Linking. + +(** External oracle *) +Parameter rtl2btl: RTL.function -> BTL.code * node * (PTree.t node). + +Local Open Scope error_monad_scope. + +Definition transf_function (f: RTL.function) : res BTL.function := + let (tcte, dupmap) := rtl2btl f in + let (tc, te) := tcte in + let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) (right_assoc_code tc) te in + do u <- verify_function dupmap f' f; + OK f'. + +Definition transf_fundef (f: RTL.fundef) : res fundef := + transf_partial_fundef transf_function f. + +Definition transf_program (p: RTL.program) : res program := + transform_partial_program transf_fundef p. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v new file mode 100644 index 00000000..4f76d0b7 --- /dev/null +++ b/scheduling/RTLtoBTLproof.v @@ -0,0 +1,239 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad BTL. + +Require Import Errors Linking RTLtoBTL. + +Require Import Linking. + +Record match_function dupmap f tf: Prop := { + dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); + dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); + code_right_assoc: forall pc ib, (fn_code tf)!pc=Some ib -> is_right_assoc ib.(entry); + preserv_fnsig: fn_sig tf = RTL.fn_sig f; + preserv_fnparams: fn_params tf = RTL.fn_params f; + preserv_fnstacksize: fn_stacksize tf = RTL.fn_stacksize f +}. + +Inductive match_fundef: RTL.fundef -> fundef -> Prop := + | match_Internal dupmap f tf: match_function dupmap f tf -> match_fundef (Internal f) (Internal tf) + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: RTL.stackframe -> stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc' = Some pc) + : match_stackframes (RTL.Stackframe res f sp pc rs) (Stackframe res f' sp pc' rs). + +Lemma verify_function_correct dupmap f f' tt: + verify_function dupmap f' f = OK tt -> + fn_sig f' = RTL.fn_sig f -> + fn_params f' = RTL.fn_params f -> + fn_stacksize f' = RTL.fn_stacksize f -> + (forall pc ib, (fn_code f')!pc=Some ib -> is_right_assoc ib.(entry)) -> + match_function dupmap f f'. +Proof. + unfold verify_function; intro VERIF. monadInv VERIF. + constructor; eauto. + - eapply verify_cfg_correct; eauto. + - eapply verify_is_copy_correct; eauto. +Qed. + +Lemma transf_function_correct f f': + transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. +Proof. + unfold transf_function; unfold bind. repeat autodestruct. + intros H _ _ X. inversion X; subst; clear X. + eexists; eapply verify_function_correct; simpl; eauto. + eapply right_assoc_code_correct; eauto. +Qed. + +Lemma transf_fundef_correct f f': + transf_fundef f = OK f' -> match_fundef f f'. +Proof. + intros TRANSF; destruct f; simpl; monadInv TRANSF. + + exploit transf_function_correct; eauto. + intros (dupmap & MATCH_F). + eapply match_Internal; eauto. + + eapply match_External. +Qed. + +Definition match_prog (p: RTL.program) (tp: program) := + match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp. + +Lemma transf_program_match: + forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog. +Proof. + intros. eapply match_transform_partial_program_contextual; eauto. +Qed. + +Section BTL_SIMULATES_RTL. + +Variable prog: RTL.program. +Variable tprog: program. + +Hypothesis TRANSL: match_prog prog tprog. + +Let ge := Genv.globalenv prog. +Let tge := Genv.globalenv tprog. + +Local Open Scope nat_scope. + +(* TODO: + faire un dessin ASCII art de la simulation avec match_states_intro + +Le iblock en paramètre représente l'état de l'exécution côté BTL +depuis le début de l'exécution du block +*) + +Inductive match_states: iblock -> RTL.state -> state -> Prop := + | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) + ib dupmap st f sp pc rs m st' f' pc' pc0' pc0 rs0 m0 isfst ib0 + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function dupmap f f') + (ATpc0: (fn_code f')!pc0' = Some ib0) + (DUPLIC: dupmap!pc0' = Some pc0) + (MIB: match_iblock dupmap (RTL.fn_code f) isfst pc' ib None) + (RIGHTA: is_right_assoc ib) + (RTL_RUN: star RTL.step ge (RTL.State st f sp pc0 rs0 m0) E0 (RTL.State st f sp pc rs m)) + (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) + : match_states ib (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m) + | match_states_call + ib st st' f f' args m + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_fundef f f') + : match_states ib (RTL.Callstate st f args m) (Callstate st' f' args m) + | match_states_return + ib st st' v m + (STACKS: list_forall2 match_stackframes st st') + : match_states ib (RTL.Returnstate st v m) (Returnstate st' v m) + . + +Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. +Proof. + rewrite <- (Genv.find_symbol_match TRANSL). reflexivity. +Qed. + +Lemma senv_preserved: Senv.equiv ge tge. +Proof. + eapply (Genv.senv_match TRANSL). +Qed. + +Lemma functions_translated (v: val) (f: RTL.fundef): + Genv.find_funct ge v = Some f -> + exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tge v = Some tf /\ linkorder cunit prog. +Proof. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & A & B & C). + repeat eexists; intuition eauto. + + unfold incl; auto. + + eapply linkorder_refl. +Qed. + +Lemma function_ptr_translated v f: + Genv.find_funct_ptr ge v = Some f -> + exists tf, + Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = OK tf. +Proof. + intros. + exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. +Qed. + +Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = RTL.funsig f. +Proof. + intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto. + erewrite preserv_fnsig; eauto. +Qed. + +Lemma transf_initial_states s1: + RTL.initial_state prog s1 -> + exists ib s2, initial_state tprog s2 /\ match_states ib s1 s2. +Proof. + intros. inv H. + exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF). + eexists. eexists. split. + - econstructor; eauto. + + eapply (Genv.init_mem_transf_partial TRANSL); eauto. + + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. + symmetry. eapply match_program_main. eauto. + + erewrite function_sig_translated; eauto. + - constructor; eauto. + constructor. + apply transf_fundef_correct; auto. +Admitted. (* TODO *) + +Lemma transf_final_states ib s1 s2 r: + match_states ib s1 s2 -> RTL.final_state s1 r -> final_state s2 r. +Proof. + intros. inv H0. inv H. inv STACKS. constructor. +Qed. + +Lemma find_function_preserved ri rs0 fd + (FIND : RTL.find_function ge ri rs0 = Some fd) + : exists fd', find_function tge ri rs0 = Some fd' + /\ transf_fundef fd = OK fd'. +Proof. + pose symbols_preserved as SYMPRES. + destruct ri. + + simpl in FIND; apply functions_translated in FIND. + destruct FIND as (tf & cunit & TFUN & GFIND & LO). + eexists; split. eauto. assumption. + + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF). + eexists; split. simpl. rewrite symbols_preserved. + rewrite GFS. eassumption. assumption. +Qed. + +(* Inspired from Duplicateproof.v *) +Lemma list_nth_z_dupmap: + forall dupmap ln ln' (pc pc': node) val, + list_nth_z ln val = Some pc -> + list_forall2 (fun n n' => dupmap!n = Some n') ln ln' -> + exists (pc': node), + list_nth_z ln' val = Some pc' + /\ dupmap!pc = Some pc'. +Proof. + induction ln; intros until val; intros LNZ LFA. + - inv LNZ. + - inv LNZ. destruct (zeq val 0) eqn:ZEQ. + + inv H0. destruct ln'; inv LFA. + simpl. exists n. split; auto. + + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto. +Qed. + +(* TODO: definir une measure sur les iblocks. +Cette mesure décroit strictement quand on exécute un "petit-pas" de iblock. +Par exemple, le nombre de noeuds (ou d'instructions "RTL") dans le iblock devrait convenir. +La hauteur de l'arbre aussi. +*) +Parameter measure: iblock -> nat. + +Theorem opt_simu s1 t s1': + RTL.step ge s1 t s1' -> + forall ib s2, + match_states ib s1 s2 -> + exists (ib' : iblock), + (exists s2', step tge s2 t s2' /\ match_states ib' s1' s2') + \/ (measure ib' < measure ib /\ t=E0 /\ match_states ib' s1' s2) + . +Admitted. + +Local Hint Resolve plus_one star_refl: core. + +Theorem transf_program_correct: + forward_simulation (RTL.semantics prog) (BTL.semantics tprog). +Proof. + eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=semantics tprog) (ltof _ measure) match_states). + constructor 1; simpl. + - apply well_founded_ltof. + - eapply transf_initial_states. + - eapply transf_final_states. + - intros s1 t s1' STEP i s2 MATCH. exploit opt_simu; eauto. clear MATCH STEP. + destruct 1 as (ib' & [ (s2' & STEP & MATCH) | (MEASURE & TRACE & MATCH) ]). + + repeat eexists; eauto. + + subst. repeat eexists; eauto. + - eapply senv_preserved. +Qed. + +End BTL_SIMULATES_RTL. -- cgit From fddab2ea68a4770f14188d120a83d45894acfbda Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 6 May 2021 10:01:02 +0200 Subject: cleaner match_states --- scheduling/RTLtoBTLproof.v | 37 +++++++++++++++++++++---------------- 1 file changed, 21 insertions(+), 16 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 4f76d0b7..2116ac70 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -83,11 +83,11 @@ Local Open Scope nat_scope. (* TODO: faire un dessin ASCII art de la simulation avec match_states_intro -Le iblock en paramètre représente l'état de l'exécution côté BTL +Le (option iblock) en paramètre représente l'état de l'exécution côté BTL depuis le début de l'exécution du block *) -Inductive match_states: iblock -> RTL.state -> state -> Prop := +Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) ib dupmap st f sp pc rs m st' f' pc' pc0' pc0 rs0 m0 isfst ib0 (STACKS: list_forall2 match_stackframes st st') @@ -98,16 +98,16 @@ Inductive match_states: iblock -> RTL.state -> state -> Prop := (RIGHTA: is_right_assoc ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pc0 rs0 m0) E0 (RTL.State st f sp pc rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) - : match_states ib (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m) + : match_states (Some ib) (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m) | match_states_call - ib st st' f f' args m + st st' f f' args m (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_fundef f f') - : match_states ib (RTL.Callstate st f args m) (Callstate st' f' args m) + : match_states None (RTL.Callstate st f args m) (Callstate st' f' args m) | match_states_return - ib st st' v m + st st' v m (STACKS: list_forall2 match_stackframes st st') - : match_states ib (RTL.Returnstate st v m) (Returnstate st' v m) + : match_states None (RTL.Returnstate st v m) (Returnstate st' v m) . Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. @@ -161,7 +161,7 @@ Proof. - constructor; eauto. constructor. apply transf_fundef_correct; auto. -Admitted. (* TODO *) +Qed. Lemma transf_final_states ib s1 s2 r: match_states ib s1 s2 -> RTL.final_state s1 r -> final_state s2 r. @@ -209,13 +209,18 @@ La hauteur de l'arbre aussi. *) Parameter measure: iblock -> nat. -Theorem opt_simu s1 t s1': +Definition omeasure (oib: option iblock): nat := + match oib with + | None => 0 + | Some ib => measure ib + end. + +Theorem opt_simu s1 t s1' oib s2: RTL.step ge s1 t s1' -> - forall ib s2, - match_states ib s1 s2 -> - exists (ib' : iblock), - (exists s2', step tge s2 t s2' /\ match_states ib' s1' s2') - \/ (measure ib' < measure ib /\ t=E0 /\ match_states ib' s1' s2) + match_states oib s1 s2 -> + exists (oib' : option iblock), + (exists s2', step tge s2 t s2' /\ match_states oib' s1' s2') + \/ (omeasure oib' < omeasure oib /\ t=E0 /\ match_states oib' s1' s2) . Admitted. @@ -224,13 +229,13 @@ Local Hint Resolve plus_one star_refl: core. Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (BTL.semantics tprog). Proof. - eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=semantics tprog) (ltof _ measure) match_states). + eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=semantics tprog) (ltof _ omeasure) match_states). constructor 1; simpl. - apply well_founded_ltof. - eapply transf_initial_states. - eapply transf_final_states. - intros s1 t s1' STEP i s2 MATCH. exploit opt_simu; eauto. clear MATCH STEP. - destruct 1 as (ib' & [ (s2' & STEP & MATCH) | (MEASURE & TRACE & MATCH) ]). + destruct 1 as (oib' & [ (s2' & STEP & MATCH) | (MEASURE & TRACE & MATCH) ]). + repeat eexists; eauto. + subst. repeat eexists; eauto. - eapply senv_preserved. -- cgit From c78cfdce8c3af0a923d62ae26a182757204a6031 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 6 May 2021 11:02:51 +0200 Subject: fix match_states --- scheduling/RTLtoBTLproof.v | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 2116ac70..ed661280 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -6,7 +6,7 @@ Require Import Errors Linking RTLtoBTL. Require Import Linking. -Record match_function dupmap f tf: Prop := { +Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); code_right_assoc: forall pc ib, (fn_code tf)!pc=Some ib -> is_right_assoc ib.(entry); @@ -80,25 +80,27 @@ Let tge := Genv.globalenv tprog. Local Open Scope nat_scope. -(* TODO: - faire un dessin ASCII art de la simulation avec match_states_intro +(* TODO: Le (option iblock) en paramètre représente l'état de l'exécution côté BTL depuis le début de l'exécution du block + + faire un dessin ASCII art de la simulation avec match_states_intro + *) Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) - ib dupmap st f sp pc rs m st' f' pc' pc0' pc0 rs0 m0 isfst ib0 + ib dupmap st f sp pc rs m st' f' pc0' pc0 rs0 m0 isfst ib0 (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') (ATpc0: (fn_code f')!pc0' = Some ib0) (DUPLIC: dupmap!pc0' = Some pc0) - (MIB: match_iblock dupmap (RTL.fn_code f) isfst pc' ib None) + (MIB: match_iblock dupmap (RTL.fn_code f) isfst pc ib None) (RIGHTA: is_right_assoc ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pc0 rs0 m0) E0 (RTL.State st f sp pc rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) - : match_states (Some ib) (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m) + : match_states (Some ib) (RTL.State st f sp pc rs m) (State st' f' sp pc0' rs0 m0) | match_states_call st st' f f' args m (STACKS: list_forall2 match_stackframes st st') -- cgit From 5757c5a377b54464b37bdce6a6f9630caefef826 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 6 May 2021 17:28:07 +0200 Subject: init BTL_SEtheory (by copy/paste from RTLpathSE_theory) --- scheduling/BTL_SEtheory.v | 1831 +++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 1831 insertions(+) create mode 100644 scheduling/BTL_SEtheory.v (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v new file mode 100644 index 00000000..53c00c8b --- /dev/null +++ b/scheduling/BTL_SEtheory.v @@ -0,0 +1,1831 @@ +(* A theory of symbolic execution on BTL + +NB: an efficient implementation with hash-consing will be defined in another file (some day) + +*) + +Require Import Coqlib Maps Floats. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL BTL OptionMonad. + +(* TODO remove this, when copy-paste of RTLpathSE_theory is clearly over... *) +Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) +Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) + +(** * Syntax and semantics of symbolic values *) + +(* symbolic value *) +Inductive sval := + | Sinput (r: reg) + | Sop (op:operation) (lsv: list_sval) (sm: smem) + | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) +with list_sval := + | Snil + | Scons (sv: sval) (lsv: list_sval) +(* symbolic memory *) +with smem := + | Sinit + | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval). + +Scheme sval_mut := Induction for sval Sort Prop +with list_sval_mut := Induction for list_sval Sort Prop +with smem_mut := Induction for smem Sort Prop. + +Fixpoint list_sval_inj (l: list sval): list_sval := + match l with + | nil => Snil + | v::l => Scons v (list_sval_inj l) + end. + +Local Open Scope option_monad_scope. + +Fixpoint eval_sval (ge: BTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := + match sv with + | Sinput r => Some (rs0#r) + | Sop op l sm => + SOME args <- eval_list_sval ge sp l rs0 m0 IN + SOME m <- eval_smem ge sp sm rs0 m0 IN + eval_operation ge sp op args m + | Sload sm trap chunk addr lsv => + match trap with + | TRAP => + SOME args <- eval_list_sval ge sp lsv rs0 m0 IN + SOME a <- eval_addressing ge sp addr args IN + SOME m <- eval_smem ge sp sm rs0 m0 IN + Mem.loadv chunk m a + | NOTRAP => + SOME args <- eval_list_sval ge sp lsv rs0 m0 IN + match (eval_addressing ge sp addr args) with + | None => Some (default_notrap_load_value chunk) + | Some a => + SOME m <- eval_smem ge sp sm rs0 m0 IN + match (Mem.loadv chunk m a) with + | None => Some (default_notrap_load_value chunk) + | Some val => Some val + end + end + end + end +with eval_list_sval (ge: BTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := + match lsv with + | Snil => Some nil + | Scons sv lsv' => + SOME v <- eval_sval ge sp sv rs0 m0 IN + SOME lv <- eval_list_sval ge sp lsv' rs0 m0 IN + Some (v::lv) + end +with eval_smem (ge: BTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem := + match sm with + | Sinit => Some m0 + | Sstore sm chunk addr lsv srce => + SOME args <- eval_list_sval ge sp lsv rs0 m0 IN + SOME a <- eval_addressing ge sp addr args IN + SOME m <- eval_smem ge sp sm rs0 m0 IN + SOME sv <- eval_sval ge sp srce rs0 m0 IN + Mem.storev chunk m a sv + end. + +Lemma eval_list_sval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: + (forall r : reg, eval_sval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> + eval_list_sval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Proof. + intros H; induction l as [|r l]; simpl; repeat autodestruct; auto. +Qed. + +Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := + SOME args <- eval_list_sval ge sp lsv rs0 m0 IN + SOME m <- eval_smem ge sp sm rs0 m0 IN + eval_condition cond args m. + + +(** * Auxiliary definitions on Builtins *) +(* TODO: clean this. Some generic stuffs could be put in [AST.v] *) + +Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *) + +Variable ge: BTL.genv. +Variable sp: val. +Variable m: mem. +Variable rs0: regset. +Variable m0: mem. + +Inductive seval_builtin_arg: builtin_arg sval -> val -> Prop := + | seval_BA: forall x v, + eval_sval ge sp x rs0 m0 = Some v -> + seval_builtin_arg (BA x) v + | seval_BA_int: forall n, + seval_builtin_arg (BA_int n) (Vint n) + | seval_BA_long: forall n, + seval_builtin_arg (BA_long n) (Vlong n) + | seval_BA_float: forall n, + seval_builtin_arg (BA_float n) (Vfloat n) + | seval_BA_single: forall n, + seval_builtin_arg (BA_single n) (Vsingle n) + | seval_BA_loadstack: forall chunk ofs v, + Mem.loadv chunk m (Val.offset_ptr sp ofs) = Some v -> + seval_builtin_arg (BA_loadstack chunk ofs) v + | seval_BA_addrstack: forall ofs, + seval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr sp ofs) + | seval_BA_loadglobal: forall chunk id ofs v, + Mem.loadv chunk m (Senv.symbol_address ge id ofs) = Some v -> + seval_builtin_arg (BA_loadglobal chunk id ofs) v + | seval_BA_addrglobal: forall id ofs, + seval_builtin_arg (BA_addrglobal id ofs) (Senv.symbol_address ge id ofs) + | seval_BA_splitlong: forall hi lo vhi vlo, + seval_builtin_arg hi vhi -> seval_builtin_arg lo vlo -> + seval_builtin_arg (BA_splitlong hi lo) (Val.longofwords vhi vlo) + | seval_BA_addptr: forall a1 a2 v1 v2, + seval_builtin_arg a1 v1 -> seval_builtin_arg a2 v2 -> + seval_builtin_arg (BA_addptr a1 a2) + (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2). + +Definition seval_builtin_args (al: list (builtin_arg sval)) (vl: list val) : Prop := + list_forall2 seval_builtin_arg al vl. + +Lemma seval_builtin_arg_determ: + forall a v, seval_builtin_arg a v -> forall v', seval_builtin_arg a v' -> v' = v. +Proof. + induction 1; intros v' EV; inv EV; try congruence. + f_equal; eauto. + apply IHseval_builtin_arg1 in H3. apply IHseval_builtin_arg2 in H5. subst; auto. +Qed. + +Lemma eval_builtin_args_determ: + forall al vl, seval_builtin_args al vl -> forall vl', seval_builtin_args al vl' -> vl' = vl. +Proof. + induction 1; intros v' EV; inv EV; f_equal; eauto using seval_builtin_arg_determ. +Qed. + +End SEVAL_BUILTIN_ARG. + +(* NB: generic function that could be put into [AST] file *) +Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B := + match arg with + | BA x => BA (f x) + | BA_int n => BA_int n + | BA_long n => BA_long n + | BA_float f => BA_float f + | BA_single s => BA_single s + | BA_loadstack chunk ptr => BA_loadstack chunk ptr + | BA_addrstack ptr => BA_addrstack ptr + | BA_loadglobal chunk id ptr => BA_loadglobal chunk id ptr + | BA_addrglobal id ptr => BA_addrglobal id ptr + | BA_splitlong ba1 ba2 => BA_splitlong (builtin_arg_map f ba1) (builtin_arg_map f ba2) + | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2) + end. + +Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg: forall arg varg, + (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + eval_builtin_arg ge (fun r => rs # r) sp m arg varg -> + seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg. +Proof. + induction arg. + all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence). + - intros varg SEVAL BARG. inv BARG. simpl. constructor. + eapply IHarg1; eauto. eapply IHarg2; eauto. + - intros varg SEVAL BARG. inv BARG. simpl. constructor. + eapply IHarg1; eauto. eapply IHarg2; eauto. +Qed. + +Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs: + (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + eval_builtin_args ge (fun r => rs # r) sp m args vargs -> + seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs. +Proof. + induction 2. + - constructor. + - simpl. constructor; [| assumption]. + eapply seval_builtin_arg_correct; eauto. +Qed. + +Lemma seval_builtin_arg_complete ge sp rs m rs0 m0 sreg: forall arg varg, + (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg -> + eval_builtin_arg ge (fun r => rs # r) sp m arg varg. +Proof. + induction arg. + all: intros varg SEVAL BARG; try (inv BARG; constructor; congruence). + - inv BARG. rewrite SEVAL in H0. inv H0. constructor. + - inv BARG. simpl. constructor. + eapply IHarg1; eauto. eapply IHarg2; eauto. + - inv BARG. simpl. constructor. + eapply IHarg1; eauto. eapply IHarg2; eauto. +Qed. + +Lemma seval_builtin_args_complete ge sp rs m rs0 m0 sreg: forall args vargs, + (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs -> + eval_builtin_args ge (fun r => rs # r) sp m args vargs. +Proof. + induction args. + - simpl. intros. inv H0. constructor. + - intros vargs SEVAL BARG. simpl in BARG. inv BARG. + constructor; [| eapply IHargs; eauto]. + eapply seval_builtin_arg_complete; eauto. +Qed. + +Fixpoint seval_builtin_sval ge sp bsv rs0 m0 := + match bsv with + | BA sv => SOME v <- eval_sval ge sp sv rs0 m0 IN Some (BA v) + | BA_splitlong sv1 sv2 => + SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN + SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN + Some (BA_splitlong v1 v2) + | BA_addptr sv1 sv2 => + SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN + SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN + Some (BA_addptr v1 v2) + | BA_int i => Some (BA_int i) + | BA_long l => Some (BA_long l) + | BA_float f => Some (BA_float f) + | BA_single s => Some (BA_single s) + | BA_loadstack chk ptr => Some (BA_loadstack chk ptr) + | BA_addrstack ptr => Some (BA_addrstack ptr) + | BA_loadglobal chk id ptr => Some (BA_loadglobal chk id ptr) + | BA_addrglobal id ptr => Some (BA_addrglobal id ptr) + end. + +Fixpoint eval_list_builtin_sval ge sp lbsv rs0 m0 := + match lbsv with + | nil => Some nil + | bsv::lbsv => SOME v <- seval_builtin_sval ge sp bsv rs0 m0 IN + SOME lv <- eval_list_builtin_sval ge sp lbsv rs0 m0 IN + Some (v::lv) + end. + +Lemma eval_list_builtin_sval_nil ge sp rs0 m0 lbs2: + eval_list_builtin_sval ge sp lbs2 rs0 m0 = Some nil -> + lbs2 = nil. +Proof. + destruct lbs2; simpl; auto. + intros. destruct (seval_builtin_sval _ _ _ _ _); + try destruct (eval_list_builtin_sval _ _ _ _ _); discriminate. +Qed. + +Lemma seval_builtin_sval_arg ge sp rs0 m0 bs: + forall ba m v, + seval_builtin_sval ge sp bs rs0 m0 = Some ba -> + eval_builtin_arg ge (fun id => id) sp m ba v -> + seval_builtin_arg ge sp m rs0 m0 bs v. +Proof. + induction bs; simpl; + try (intros ba m v H; inversion H; subst; clear H; + intros H; inversion H; subst; + econstructor; auto; fail). + - intros ba m v; destruct (eval_sval _ _ _ _ _) eqn: SV; + intros H; inversion H; subst; clear H. + intros H; inversion H; subst. + econstructor; auto. + - intros ba m v. + destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence. + destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence. + intros H; inversion H; subst; clear H. + intros H; inversion H; subst. + econstructor; eauto. + - intros ba m v. + destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence. + destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence. + intros H; inversion H; subst; clear H. + intros H; inversion H; subst. + econstructor; eauto. +Qed. + +Lemma seval_builtin_arg_sval ge sp m rs0 m0 v: forall bs, + seval_builtin_arg ge sp m rs0 m0 bs v -> + exists ba, + seval_builtin_sval ge sp bs rs0 m0 = Some ba + /\ eval_builtin_arg ge (fun id => id) sp m ba v. +Proof. + induction 1. + all: try (eexists; constructor; [simpl; reflexivity | constructor]). + 2-3: try assumption. + - eexists. constructor. + + simpl. rewrite H. reflexivity. + + constructor. + - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1). + destruct IHseval_builtin_arg2 as (ba2 & A2 & B2). + eexists. constructor. + + simpl. rewrite A1. rewrite A2. reflexivity. + + constructor; assumption. + - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1). + destruct IHseval_builtin_arg2 as (ba2 & A2 & B2). + eexists. constructor. + + simpl. rewrite A1. rewrite A2. reflexivity. + + constructor; assumption. +Qed. + +Lemma seval_builtin_sval_args ge sp rs0 m0 lbs: + forall lba m v, + eval_list_builtin_sval ge sp lbs rs0 m0 = Some lba -> + list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba v -> + seval_builtin_args ge sp m rs0 m0 lbs v. +Proof. + unfold seval_builtin_args; induction lbs; simpl; intros lba m v. + - intros H; inversion H; subst; clear H. + intros H; inversion H. econstructor. + - destruct (seval_builtin_sval _ _ _ _ _) eqn:SV; try congruence. + destruct (eval_list_builtin_sval _ _ _ _ _) eqn: SVL; try congruence. + intros H; inversion H; subst; clear H. + intros H; inversion H; subst; clear H. + econstructor; eauto. + eapply seval_builtin_sval_arg; eauto. +Qed. + +Lemma seval_builtin_args_sval ge sp m rs0 m0 lv: forall lbs, + seval_builtin_args ge sp m rs0 m0 lbs lv -> + exists lba, + eval_list_builtin_sval ge sp lbs rs0 m0 = Some lba + /\ list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba lv. +Proof. + induction 1. + - eexists. constructor. + + simpl. reflexivity. + + constructor. + - destruct IHlist_forall2 as (lba & A & B). + apply seval_builtin_arg_sval in H. destruct H as (ba & A' & B'). + eexists. constructor. + + simpl. rewrite A'. rewrite A. reflexivity. + + constructor; assumption. +Qed. + +Lemma seval_builtin_sval_correct ge sp m rs0 m0: forall bs1 v bs2, + seval_builtin_arg ge sp m rs0 m0 bs1 v -> + (seval_builtin_sval ge sp bs1 rs0 m0) = (seval_builtin_sval ge sp bs2 rs0 m0) -> + seval_builtin_arg ge sp m rs0 m0 bs2 v. +Proof. + intros. exploit seval_builtin_arg_sval; eauto. + intros (ba & X1 & X2). + eapply seval_builtin_sval_arg; eauto. + congruence. +Qed. + +Lemma eval_list_builtin_sval_correct ge sp m rs0 m0 vargs: forall lbs1, + seval_builtin_args ge sp m rs0 m0 lbs1 vargs -> + forall lbs2, (eval_list_builtin_sval ge sp lbs1 rs0 m0) = (eval_list_builtin_sval ge sp lbs2 rs0 m0) -> + seval_builtin_args ge sp m rs0 m0 lbs2 vargs. +Proof. + intros. exploit seval_builtin_args_sval; eauto. + intros (ba & X1 & X2). + eapply seval_builtin_sval_args; eauto. + congruence. +Qed. + +(** * Symbolic (final) value of a block *) + +Inductive sfval := + | Sgoto (pc: exit) + | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:exit) + (* NB: [res] the return register is hard-wired ! Is it restrictive ? *) + | Stailcall: signature -> sval + ident -> list_sval -> sfval + | Sbuiltin (ef:external_function) (sargs: list (builtin_arg sval)) (res: builtin_res reg) (pc:exit) + | Sjumptable (sv: sval) (tbl: list exit) + | Sreturn: option sval -> sfval +. + +Definition sfind_function (ge: BTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := + match svos with + | inl sv => SOME v <- eval_sval ge sp sv rs0 m0 IN Genv.find_funct ge v + | inr symb => SOME b <- Genv.find_symbol ge symb IN Genv.find_funct_ptr ge b + end +. + +Inductive sem_sfval (ge: BTL.genv) (sp:val) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := + | exec_Sgoto pc rs m: + sem_sfval ge sp stack f rs0 m0 (Sgoto pc) rs m E0 (State stack f sp pc rs m) + | exec_Sreturn stk osv rs m m' v: + sp = (Vptr stk Ptrofs.zero) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + match osv with Some sv => eval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> + sem_sfval ge sp stack f rs0 m0 (Sreturn osv) rs m + E0 (Returnstate stack v m') + | exec_Scall rs m sig svos lsv args res pc fd: + sfind_function ge sp svos rs0 m0 = Some fd -> + funsig fd = sig -> + eval_list_sval ge sp lsv rs0 m0 = Some args -> + sem_sfval ge sp stack f rs0 m0 (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) + | exec_Stailcall stk rs m sig svos args fd m' lsv: + sfind_function ge sp svos rs0 m0 = Some fd -> + funsig fd = sig -> + sp = Vptr stk Ptrofs.zero -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + eval_list_sval ge sp lsv rs0 m0 = Some args -> + sem_sfval ge sp stack f rs0 m0 (Stailcall sig svos lsv) rs m + E0 (Callstate stack fd args m') + | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: + seval_builtin_args ge sp m rs0 m0 sargs vargs -> + external_call ef ge vargs m t vres m' -> + sem_sfval ge sp stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m + t (State stack f sp pc (regmap_setres res vres rs) m') + | exec_Sjumptable sv tbl pc' n rs m: + eval_sval ge sp sv rs0 m0 = Some (Vint n) -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + sem_sfval ge sp stack f rs0 m0 (Sjumptable sv tbl) rs m + E0 (State stack f sp pc' rs m) +. + + +(** * Preservation properties *) + +Section SymbValPreserved. + +Variable ge ge': BTL.genv. + +Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. + +Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol ge id = Senv.find_symbol ge' id. +Proof. + destruct senv_preserved_BTL as (A & B & C). congruence. +Qed. + +Lemma senv_symbol_address_preserved id ofs: + Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. +Proof. + unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. + reflexivity. +Qed. + +Lemma eval_sval_preserved sp sv rs0 m0: + eval_sval ge sp sv rs0 m0 = eval_sval ge' sp sv rs0 m0. +Proof. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval ge sp lsv rs0 m0 = eval_list_sval ge' sp lsv rs0 m0) + (P1 := fun sm => eval_smem ge sp sm rs0 m0 = eval_smem ge' sp sm rs0 m0); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _ _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _ _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (eval_sval _ _ _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (eval_smem _ _ _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma seval_builtin_arg_preserved sp m rs0 m0: + forall bs varg, + seval_builtin_arg ge sp m rs0 m0 bs varg -> + seval_builtin_arg ge' sp m rs0 m0 bs varg. +Proof. + induction 1. + all: try (constructor; auto). + - rewrite <- eval_sval_preserved. assumption. + - rewrite <- senv_symbol_address_preserved. assumption. + - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. +Qed. + +Lemma seval_builtin_args_preserved sp m rs0 m0 lbs vargs: + seval_builtin_args ge sp m rs0 m0 lbs vargs -> + seval_builtin_args ge' sp m rs0 m0 lbs vargs. +Proof. + induction 1; constructor; eauto. + eapply seval_builtin_arg_preserved; auto. +Qed. + +Lemma list_sval_eval_preserved sp lsv rs0 m0: + eval_list_sval ge sp lsv rs0 m0 = eval_list_sval ge' sp lsv rs0 m0. +Proof. + induction lsv; simpl; auto. + rewrite eval_sval_preserved. destruct (eval_sval _ _ _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma smem_eval_preserved sp sm rs0 m0: + eval_smem ge sp sm rs0 m0 = eval_smem ge' sp sm rs0 m0. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsm; clear IHsm. destruct (eval_smem _ _ _ _); auto. + rewrite eval_sval_preserved; auto. +Qed. + +Lemma seval_condition_preserved sp cond lsv sm rs0 m0: + seval_condition ge sp cond lsv sm rs0 m0 = seval_condition ge' sp cond lsv sm rs0 m0. +Proof. + unfold seval_condition. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _ _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + +End SymbValPreserved. + + +(* Syntax and Semantics of symbolic internal states *) +(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) +Record sistate := { si_pre: BTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }. + +(* Predicate on which (rs, m) is a possible final state after evaluating [st] on (rs0, m0) *) +Definition sem_sistate (ge: BTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := + st.(si_pre) ge sp rs0 m0 + /\ eval_smem ge sp st.(si_smem) rs0 m0 = Some m + /\ forall (r:reg), eval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r). + +Definition abort_sistate (ge: BTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := + ~(st.(si_pre) ge sp rs0 m0) + \/ eval_smem ge sp st.(si_smem) rs0 m0 = None + \/ exists (r: reg), eval_sval ge sp (st.(si_sreg) r) rs0 m0 = None. + +Record sstate_exit := { se_sis:> sistate; se_sfv: sfval }. + +Definition sem_sexit ge (sp:val) (st: sstate_exit) stack f (rs0: regset) (m0: mem) rs m (t: trace) (s:BTL.state) := + sem_sistate ge sp st rs0 m0 rs m /\ + sem_sfval ge sp stack f rs0 m0 st.(se_sfv) rs m t s + . + +(** * Symbolic execution of final step *) +Definition sexec_final_sfv (i: final) (sis: sistate): sfval := + match i with + | Bgoto pc => Sgoto pc + | Bcall sig ros args res pc => + let svos := sum_left_map sis.(si_sreg) ros in + let sargs := list_sval_inj (List.map sis.(si_sreg) args) in + Scall sig svos sargs res pc + | Btailcall sig ros args => + let svos := sum_left_map sis.(si_sreg) ros in + let sargs := list_sval_inj (List.map sis.(si_sreg) args) in + Stailcall sig svos sargs + | Bbuiltin ef args res pc => + let sargs := List.map (builtin_arg_map sis.(si_sreg)) args in + Sbuiltin ef sargs res pc + | Breturn or => + let sor := SOME r <- or IN Some (sis.(si_sreg) r) in + Sreturn sor + | Bjumptable reg tbl => + let sv := sis.(si_sreg) reg in + Sjumptable sv tbl + end. + +Local Hint Constructors sem_sfval: core. + +Lemma sexec_final_svf_correct ge sp i f sis stack rs0 m0 t rs m s: + sem_sistate ge sp sis rs0 m0 rs m -> + final_step ge stack f sp rs m i t s -> + sem_sfval ge sp stack f rs0 m0 (sexec_final_sfv i sis) rs m t s. +Proof. + intros (PRE&MEM®). + destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. + + (* Bcall *) intros; eapply exec_Scall; auto. + - destruct ros; simpl in * |- *; auto. + rewrite REG; auto. + - erewrite eval_list_sval_inj; simpl; auto. + + (* Btailcall *) intros. eapply exec_Stailcall; auto. + - destruct ros; simpl in * |- *; auto. + rewrite REG; auto. + - erewrite eval_list_sval_inj; simpl; auto. + + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto. + eapply seval_builtin_args_correct; eauto. + + (* Bjumptable *) intros. eapply exec_Sjumptable; eauto. congruence. +Qed. + +Local Hint Constructors final_step: core. +Local Hint Resolve seval_builtin_args_complete: core. + +Lemma sexec_final_svf_complete ge sp i f sis stack rs0 m0 t rs m s: + sem_sistate ge sp sis rs0 m0 rs m -> + sem_sfval ge sp stack f rs0 m0 (sexec_final_sfv i sis) rs m t s + -> final_step ge stack f sp rs m i t s. +Proof. + intros (PRE&MEM®). + destruct i; simpl; intros LAST; inv LAST; eauto. + + (* Breturn *) + enough (v=regmap_optget res Vundef rs) as ->; eauto. + destruct res; simpl in *; congruence. + + (* Bcall *) + erewrite eval_list_sval_inj in *; try_simplify_someHyps. + intros; eapply exec_Bcall; eauto. + destruct fn; simpl in * |- *; auto. + rewrite REG in * |- ; auto. + + (* Btailcall *) + erewrite eval_list_sval_inj in *; try_simplify_someHyps. + intros; eapply exec_Btailcall; eauto. + destruct fn; simpl in * |- *; auto. + rewrite REG in * |- ; auto. + + (* Bjumptable *) + eapply exec_Bjumptable; eauto. + congruence. +Qed. + + +(* TODO: clean/recover this COPY-PASTE OF RTLpathSE_theory.v + + +(* Syntax and semantics of symbolic exit states *) + +Definition all_fallthrough ge sp (lx: list sistate_exit) rs0 m0: Prop := + forall ext, List.In ext lx -> + seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some false. + +Lemma all_fallthrough_revcons ge sp ext rs m lx: + all_fallthrough ge sp (ext::lx) rs m -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false + /\ all_fallthrough ge sp lx rs m. +Proof. + intros ALLFU. constructor. + - assert (In ext (ext::lx)) by (constructor; auto). apply ALLFU in H. assumption. + - intros ext' INEXT. assert (In ext' (ext::lx)) by (apply in_cons; auto). + apply ALLFU in H. assumption. +Qed. + +(** Semantic of an exit in pseudo code: + if si_cond (si_condargs) + si_elocal; goto if_so + else () +*) + +Definition ssem_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true + /\ ssem_local ge sp (si_elocal ext) rs m rs' m' + /\ (si_ifso ext) = pc'. + +(* Either an abort on the condition evaluation OR an abort on the sistate_local IF the condition was true *) +Definition sabort_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) : Prop := + let sev_cond := seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m in + sev_cond = None + \/ (sev_cond = Some true /\ sabort_local ge sp ext.(si_elocal) rs m). + +(** * Syntax and Semantics of symbolic internal state *) +Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. + +Definition all_fallthrough_upto_exit ge sp ext lx' lx rs m : Prop := + is_tail (ext::lx') lx /\ all_fallthrough ge sp lx' rs m. + +(** Semantic of a sistate in pseudo code: + si_exit1; si_exit2; ...; si_exitn; + si_local; goto si_pc *) + +(* Note: in RTLpath, is.(icontinue) = false iff we took an early exit *) + +Definition ssem_internal (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop := + if (is.(icontinue)) + then + ssem_local ge sp st.(si_local) rs m is.(irs) is.(imem) + /\ st.(si_pc) = is.(ipc) + /\ all_fallthrough ge sp st.(si_exits) rs m + else exists ext lx, + ssem_exit ge sp ext rs m is.(irs) is.(imem) is.(ipc) + /\ all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m. + +Definition sabort (ge: RTL.genv) (sp: val) (st: sistate) (rs: regset) (m: mem): Prop := + (* No early exit was met but we aborted on the si_local *) + (all_fallthrough ge sp st.(si_exits) rs m /\ sabort_local ge sp st.(si_local) rs m) + (* OR we aborted on an evaluation of one of the early exits *) + \/ (exists ext lx, all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m /\ sabort_exit ge sp ext rs m). + +Definition ssem_internal_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := + match ois with + | Some is => ssem_internal ge sp st rs0 m0 is + | None => sabort ge sp st rs0 m0 + end. + +Definition ssem_internal_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := + match ost with + | Some st => ssem_internal_opt ge sp st rs0 m0 ois + | None => ois=None + end. + +(** * An internal state represents a parallel program ! + + We prove below that the semantics [ssem_internal_opt] is deterministic. + + *) + +Definition istate_eq ist1 ist2 := + ist1.(icontinue) = ist2.(icontinue) /\ + ist1.(ipc) = ist2.(ipc) /\ + (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\ + ist1.(imem) = ist2.(imem). + +Lemma all_fallthrough_noexit ge sp ext lx rs0 m0 rs m pc: + ssem_exit ge sp ext rs0 m0 rs m pc -> + In ext lx -> + all_fallthrough ge sp lx rs0 m0 -> + False. +Proof. + Local Hint Resolve is_tail_in: core. + intros SSEM INE ALLF. + destruct SSEM as (SSEM & SSEM'). + unfold all_fallthrough in ALLF. rewrite ALLF in SSEM; eauto. + discriminate. +Qed. + +Lemma ssem_internal_exclude_incompatible_continue ge sp st rs m is1 is2: + is1.(icontinue) = true -> + is2.(icontinue) = false -> + ssem_internal ge sp st rs m is1 -> + ssem_internal ge sp st rs m is2 -> + False. +Proof. + Local Hint Resolve all_fallthrough_noexit: core. + unfold ssem_internal. + intros CONT1 CONT2. + rewrite CONT1, CONT2; simpl. + intuition eauto. + destruct H0 as (ext & lx & SSEME & ALLFU). + destruct ALLFU as (ALLFU & ALLFU'). + eapply all_fallthrough_noexit; eauto. +Qed. + +Lemma ssem_internal_determ_continue ge sp st rs m is1 is2: + ssem_internal ge sp st rs m is1 -> + ssem_internal ge sp st rs m is2 -> + is1.(icontinue) = is2.(icontinue). +Proof. + Local Hint Resolve ssem_internal_exclude_incompatible_continue: core. + destruct (Bool.bool_dec is1.(icontinue) is2.(icontinue)) as [|H]; auto. + intros H1 H2. assert (absurd: False); intuition. + destruct (icontinue is1) eqn: His1, (icontinue is2) eqn: His2; eauto. +Qed. + +Lemma ssem_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: + ssem_local ge sp st rs0 m0 rs1 m1 -> + ssem_local ge sp st rs0 m0 rs2 m2 -> + (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + unfold ssem_local. intuition try congruence. + generalize (H5 r); rewrite H4; congruence. +Qed. + +(* TODO: lemma to move in Coqlib *) +Lemma is_tail_bounded_total {A} (l1 l2 l3: list A): is_tail l1 l3 -> is_tail l2 l3 + -> is_tail l1 l2 \/ is_tail l2 l1. +Proof. + Local Hint Resolve is_tail_cons: core. + induction 1 as [|i l1 l3 T1 IND]; simpl; auto. + intros T2; inversion T2; subst; auto. +Qed. + +Lemma exit_cond_determ ge sp rs0 m0 l1 l2: + is_tail l1 l2 -> forall ext1 lx1 ext2 lx2, + l1=(ext1 :: lx1) -> + l2=(ext2 :: lx2) -> + all_fallthrough ge sp lx1 rs0 m0 -> + seval_condition ge sp (si_cond ext1) (si_scondargs ext1) (si_smem (si_elocal ext1)) rs0 m0 = Some true -> + all_fallthrough ge sp lx2 rs0 m0 -> + ext1=ext2. +Proof. + destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst; + inversion EQ2; subst; auto. + intros D1 EVAL NYE. + Local Hint Resolve is_tail_in: core. + unfold all_fallthrough in NYE. + rewrite NYE in EVAL; eauto. + try congruence. +Qed. + +Lemma ssem_exit_determ ge sp ext rs0 m0 rs1 m1 pc1 rs2 m2 pc2: + ssem_exit ge sp ext rs0 m0 rs1 m1 pc1 -> + ssem_exit ge sp ext rs0 m0 rs2 m2 pc2 -> + pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + Local Hint Resolve exit_cond_determ eq_sym: core. + intros SSEM1 SSEM2. destruct SSEM1 as (SEVAL1 & SLOC1 & PCEQ1). destruct SSEM2 as (SEVAL2 & SLOC2 & PCEQ2). subst. + destruct (ssem_local_determ ge sp (si_elocal ext) rs0 m0 rs1 m1 rs2 m2); auto. +Qed. + +Remark is_tail_inv_left {A: Type} (a a': A) l l': + is_tail (a::l) (a'::l') -> + (a = a' /\ l = l') \/ (In a l' /\ is_tail l (a'::l')). +Proof. + intros. inv H. + - left. eauto. + - right. econstructor. + + eapply is_tail_in; eauto. + + eapply is_tail_cons_left; eauto. +Qed. + +Lemma ssem_internal_determ ge sp st rs m is1 is2: + ssem_internal ge sp st rs m is1 -> + ssem_internal ge sp st rs m is2 -> + istate_eq is1 is2. +Proof. + unfold istate_eq. + intros SEM1 SEM2. + exploit (ssem_internal_determ_continue ge sp st rs m is1 is2); eauto. + intros CONTEQ. unfold ssem_internal in * |-. rewrite CONTEQ in * |- *. + destruct (icontinue is2). + - destruct (ssem_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); + intuition (try congruence). + - destruct SEM1 as (ext1 & lx1 & SSEME1 & ALLFU1). destruct SEM2 as (ext2 & lx2 & SSEME2 & ALLFU2). + destruct ALLFU1 as (ALLFU1 & ALLFU1'). destruct ALLFU2 as (ALLFU2 & ALLFU2'). + destruct SSEME1 as (SSEME1 & SSEME1' & SSEME1''). destruct SSEME2 as (SSEME2 & SSEME2' & SSEME2''). + assert (X:ext1=ext2). + { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2 :: lx2) (si_exits st)) as [TAIL|TAIL]; eauto. } + subst. destruct (ssem_local_determ ge sp (si_elocal ext2) rs m (irs is1) (imem is1) (irs is2) (imem is2)); auto. + intuition. congruence. +Qed. + +Lemma ssem_local_exclude_sabort_local ge sp loc rs m rs' m': + ssem_local ge sp loc rs m rs' m' -> + sabort_local ge sp loc rs m -> + False. +Proof. + intros SIML ABORT. inv SIML. destruct H0 as (H0 & H0'). + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. +Qed. + +Lemma ssem_local_exclude_sabort ge sp st rs m rs' m': + ssem_local ge sp (si_local st) rs m rs' m' -> + all_fallthrough ge sp (si_exits st) rs m -> + sabort ge sp st rs m -> + False. +Proof. + intros SIML ALLF ABORT. + inv ABORT. + - intuition; eapply ssem_local_exclude_sabort_local; eauto. + - destruct H as (ext & lx & ALLFU & SABORT). + destruct ALLFU as (TAIL & _). eapply is_tail_in in TAIL. + eapply ALLF in TAIL. + destruct SABORT as [CONDFAIL | (CONDTRUE & ABORTL)]; congruence. +Qed. + +Lemma ssem_exit_fallthrough_upto_exit ge sp ext ext' lx lx' exits rs m rs' m' pc': + ssem_exit ge sp ext rs m rs' m' pc' -> + all_fallthrough_upto_exit ge sp ext lx exits rs m -> + all_fallthrough_upto_exit ge sp ext' lx' exits rs m -> + is_tail (ext'::lx') (ext::lx). +Proof. + intros SSEME ALLFU ALLFU'. + destruct ALLFU as (ISTAIL & ALLFU). destruct ALLFU' as (ISTAIL' & ALLFU'). + destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto. + inv H. + - econstructor; eauto. + - eapply is_tail_in in H2. eapply ALLFU' in H2. + destruct SSEME as (SEVAL & _). congruence. +Qed. + +Lemma ssem_exit_exclude_sabort_exit ge sp ext rs m rs' m' pc': + ssem_exit ge sp ext rs m rs' m' pc' -> + sabort_exit ge sp ext rs m -> + False. +Proof. + intros A B. destruct A as (A & A' & A''). inv B. + - congruence. + - destruct H as (_ & H). eapply ssem_local_exclude_sabort_local; eauto. +Qed. + +Lemma ssem_exit_exclude_sabort ge sp ext st lx rs m rs' m' pc': + ssem_exit ge sp ext rs m rs' m' pc' -> + all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m -> + sabort ge sp st rs m -> + False. +Proof. + intros SSEM ALLFU ABORT. + inv ABORT. + - destruct H as (ALLF & _). destruct ALLFU as (TAIL & _). + eapply is_tail_in in TAIL. + destruct SSEM as (SEVAL & _ & _). + eapply ALLF in TAIL. congruence. + - destruct H as (ext' & lx' & ALLFU' & ABORT). + exploit ssem_exit_fallthrough_upto_exit; eauto. intros ITAIL. + destruct ALLFU as (ALLFU1 & ALLFU2). destruct ALLFU' as (ALLFU1' & ALLFU2'). + exploit (is_tail_inv_left ext' ext lx' lx); eauto. intro. inv H. + + inv H0. eapply ssem_exit_exclude_sabort_exit; eauto. + + destruct H0 as (INE & TAIL). eapply ALLFU2 in INE. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence. +Qed. + +Lemma ssem_internal_exclude_sabort ge sp st rs m is: + sabort ge sp st rs m -> + ssem_internal ge sp st rs m is -> False. +Proof. + intros ABORT SEM. + unfold ssem_internal in SEM. destruct icontinue. + - destruct SEM as (SEM1 & SEM2 & SEM3). + eapply ssem_local_exclude_sabort; eauto. + - destruct SEM as (ext & lx & SEM1 & SEM2). eapply ssem_exit_exclude_sabort; eauto. +Qed. + +Definition istate_eq_opt ist1 oist := + exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. + +Lemma ssem_internal_opt_determ ge sp st rs m ois is: + ssem_internal_opt ge sp st rs m ois -> + ssem_internal ge sp st rs m is -> + istate_eq_opt is ois. +Proof. + destruct ois as [is1|]; simpl; eauto. + - intros; eexists; intuition; eapply ssem_internal_determ; eauto. + - intros; exploit ssem_internal_exclude_sabort; eauto. destruct 1. +Qed. + +(** * Symbolic execution of one internal step *) + +Definition slocal_set_sreg (st:sistate_local) (r:reg) (sv:sval) := + {| si_pre:=(fun ge sp rs m => eval_sval ge sp (st.(si_sreg) r) rs m <> None /\ (st.(si_pre) ge sp rs m)); + si_sreg:=fun y => if Pos.eq_dec r y then sv else st.(si_sreg) y; + si_smem:= st.(si_smem)|}. + +Definition slocal_set_smem (st:sistate_local) (sm:smem) := + {| si_pre:=(fun ge sp rs m => eval_smem ge sp st.(si_smem) rs m <> None /\ (st.(si_pre) ge sp rs m)); + si_sreg:= st.(si_sreg); + si_smem:= sm |}. + +Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): sistate := + {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. + +Definition slocal_store st chunk addr args src : sistate_local := + let args := list_sval_inj (List.map (si_sreg st) args) in + let src := si_sreg st src in + let sm := Sstore (si_smem st) chunk addr args src + in slocal_set_smem st sm. + +Definition siexec_inst (i: instruction) (st: sistate): option sistate := + match i with + | Inop pc' => + Some (sist_set_local st pc' st.(si_local)) + | Iop op args dst pc' => + let prev := st.(si_local) in + let vargs := list_sval_inj (List.map prev.(si_sreg) args) in + let next := slocal_set_sreg prev dst (Sop op vargs prev.(si_smem)) in + Some (sist_set_local st pc' next) + | Iload trap chunk addr args dst pc' => + let prev := st.(si_local) in + let vargs := list_sval_inj (List.map prev.(si_sreg) args) in + let next := slocal_set_sreg prev dst (Sload prev.(si_smem) trap chunk addr vargs) in + Some (sist_set_local st pc' next) + | Istore chunk addr args src pc' => + let next := slocal_store st.(si_local) chunk addr args src in + Some (sist_set_local st pc' next) + | Icond cond args ifso ifnot _ => + let prev := st.(si_local) in + let vargs := list_sval_inj (List.map prev.(si_sreg) args) in + let ex := {| si_cond:=cond; si_scondargs:=vargs; si_elocal := prev; si_ifso := ifso |} in + Some {| si_pc := ifnot; si_exits := ex::st.(si_exits); si_local := prev |} + | _ => None + end. + +Lemma slocal_set_sreg_preserves_sabort_local ge sp st rs0 m0 r sv: + sabort_local ge sp st rs0 m0 -> + sabort_local ge sp (slocal_set_sreg st r sv) rs0 m0. +Proof. + unfold sabort_local. simpl; intuition. + destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST. + - subst; rewrite H; intuition. + - right. right. exists r1. rewrite HTEST. auto. +Qed. + +Lemma slocal_set_smem_preserves_sabort_local ge sp st rs0 m0 m: + sabort_local ge sp st rs0 m0 -> + sabort_local ge sp (slocal_set_smem st m) rs0 m0. +Proof. + unfold sabort_local. simpl; intuition. +Qed. + +Lemma all_fallthrough_upto_exit_cons ge sp ext lx ext' exits rs m: + all_fallthrough_upto_exit ge sp ext lx exits rs m -> + all_fallthrough_upto_exit ge sp ext lx (ext'::exits) rs m. +Proof. + intros. inv H. econstructor; eauto. +Qed. + +Lemma all_fallthrough_cons ge sp exits rs m ext: + all_fallthrough ge sp exits rs m -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false -> + all_fallthrough ge sp (ext::exits) rs m. +Proof. + intros. unfold all_fallthrough in *. intros. + inv H1; eauto. +Qed. + +Lemma siexec_inst_preserves_sabort i ge sp rs m st st': + siexec_inst i st = Some st' -> + sabort ge sp st rs m -> sabort ge sp st' rs m. +Proof. + intros SISTEP ABORT. + destruct i; simpl in SISTEP; try discriminate; inv SISTEP; unfold sabort; simpl. + (* NOP *) + * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - left. constructor; eauto. + - right. exists ext0, lx0. constructor; eauto. + (* OP *) + * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto. + - right. exists ext0, lx0. constructor; eauto. + (* LOAD *) + * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto. + - right. exists ext0, lx0. constructor; eauto. + (* STORE *) + * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - left. constructor; eauto. eapply slocal_set_smem_preserves_sabort_local; eauto. + - right. exists ext0, lx0. constructor; eauto. + (* COND *) + * remember ({| si_cond := _; si_scondargs := _; si_elocal := _; si_ifso := _ |}) as ext. + destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - destruct (seval_condition ge sp (si_cond ext) (si_scondargs ext) + (si_smem (si_elocal ext)) rs m) eqn:SEVAL; [destruct b|]. + (* case true *) + + right. exists ext, (si_exits st). + constructor. + ++ constructor. econstructor; eauto. eauto. + ++ unfold sabort_exit. right. constructor; eauto. + subst. simpl. eauto. + (* case false *) + + left. constructor; eauto. eapply all_fallthrough_cons; eauto. + (* case None *) + + right. exists ext, (si_exits st). constructor. + ++ constructor. econstructor; eauto. eauto. + ++ unfold sabort_exit. left. eauto. + - right. exists ext0, lx0. constructor; eauto. eapply all_fallthrough_upto_exit_cons; eauto. +Qed. + +Lemma siexec_inst_WF i st: + siexec_inst i st = None -> default_succ i = None. +Proof. + destruct i; simpl; unfold sist_set_local; simpl; congruence. +Qed. + +Lemma siexec_inst_default_succ i st st': + siexec_inst i st = Some st' -> default_succ i = Some (st'.(si_pc)). +Proof. + destruct i; simpl; unfold sist_set_local; simpl; try congruence; + intro H; inversion_clear H; simpl; auto. +Qed. + + +Lemma eval_list_sval_inj_not_none ge sp st rs0 m0: forall l, + (forall r, List.In r l -> eval_sval ge sp (si_sreg st r) rs0 m0 = None -> False) -> + eval_list_sval ge sp (list_sval_inj (map (si_sreg st) l)) rs0 m0 = None -> False. +Proof. + induction l. + - intuition discriminate. + - intros ALLR. simpl. + inversion_SOME v. + + intro SVAL. inversion_SOME lv; [discriminate|]. + assert (forall r : reg, In r l -> eval_sval ge sp (si_sreg st r) rs0 m0 = None -> False). + { intros r INR. eapply ALLR. right. assumption. } + intro SVALLIST. intro. eapply IHl; eauto. + + intros. exploit (ALLR a); simpl; eauto. +Qed. + +Lemma siexec_inst_correct ge sp i st rs0 m0 rs m: + ssem_local ge sp st.(si_local) rs0 m0 rs m -> + all_fallthrough ge sp st.(si_exits) rs0 m0 -> + ssem_internal_opt2 ge sp (siexec_inst i st) rs0 m0 (istep ge i sp rs m). +Proof. + intros (PRE & MEM & REG) NYE. + destruct i; simpl; auto. + + (* Nop *) + constructor; [|constructor]; simpl; auto. + constructor; auto. + + (* Op *) + inversion_SOME v; intros OP; simpl. + - constructor; [|constructor]; simpl; auto. + constructor; simpl; auto. + * constructor; auto. congruence. + * constructor; auto. + intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst. rewrite Regmap.gss; simpl; auto. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + - left. constructor; simpl; auto. + unfold sabort_local. right. right. + simpl. exists r. destruct (Pos.eq_dec r r); try congruence. + simpl. erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + + (* LOAD *) + inversion_SOME a0; intro ADD. + { inversion_SOME v; intros LOAD; simpl. + - explore_destruct; unfold ssem_internal, ssem_local; simpl; intuition. + * unfold ssem_internal. simpl. constructor; [|constructor]; auto. + constructor; constructor; simpl; auto. congruence. intro r0. + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + * unfold ssem_internal. simpl. constructor; [|constructor]; auto. + constructor; constructor; simpl; auto. congruence. intro r0. + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl. + inversion_SOME args; intros ARGS. + 2: { exploit eval_list_sval_inj_not_none; eauto; intuition congruence. } + exploit eval_list_sval_inj; eauto. intro ARGS'. erewrite ARGS in ARGS'. inv ARGS'. rewrite ADD. + inversion_SOME m2. intro SMEM. + assert (m = m2) by congruence. subst. rewrite LOAD. reflexivity. + - explore_destruct; unfold sabort, sabort_local; simpl. + * unfold sabort. simpl. left. constructor; auto. + right. right. exists r. simpl. destruct (Pos.eq_dec r r); try congruence. + simpl. erewrite eval_list_sval_inj; simpl; auto. + rewrite ADD; simpl; auto. try_simplify_eval_svalsomeHyps. + * unfold ssem_internal. simpl. constructor; [|constructor]; auto. + constructor; constructor; simpl; auto. congruence. intro r0. + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + } { rewrite ADD. destruct t. + - simpl. left; eauto. simpl. econstructor; eauto. + right. right. simpl. exists r. destruct (Pos.eq_dec r r); [|contradiction]. + simpl. inversion_SOME args. intro SLS. + eapply eval_list_sval_inj in REG. rewrite REG in SLS. inv SLS. + rewrite ADD. reflexivity. + - simpl. constructor; [|constructor]; simpl; auto. + constructor; simpl; constructor; auto; [congruence|]. + intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst. simpl. rewrite Regmap.gss. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + } + + (* STORE *) + inversion_SOME a0; intros ADD. + { inversion_SOME m'; intros STORE; simpl. + - unfold ssem_internal, ssem_local; simpl; intuition. + * congruence. + * erewrite eval_list_sval_inj; simpl; auto. + erewrite REG. + try_simplify_someHyps. + - unfold sabort, sabort_local; simpl. + left. constructor; auto. right. left. + erewrite eval_list_sval_inj; simpl; auto. + erewrite REG. + try_simplify_someHyps. } + { unfold sabort, sabort_local; simpl. + left. constructor; auto. right. left.eval_sval + erewrite eval_list_sval_inj; simpl; auto. + erewrite ADD; simpl; auto. } + + (* COND *) + Local Hint Resolve is_tail_refl: core. + Local Hint Unfold ssem_local: core. + inversion_SOME b; intros COND. + { destruct b; simpl; unfold ssem_internal, ssem_local; simpl. + - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st). + constructor; constructor; subst; simpl; auto. + unfold seval_condition. subst; simpl. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + - intuition. unfold all_fallthrough in * |- *. simpl. + intuition. subst. simpl. + unfold seval_condition. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. } + { unfold sabort. simpl. right. + remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st). + constructor; [constructor; subst; simpl; auto|]. + left. subst; simpl; auto. + unfold seval_condition. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. } +Qed. + + +Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m: + ssem_local ge sp (st.(si_local)) rs0 m0 rs m -> + siexec_inst i st = None -> + istep ge i sp rs m = None. +Proof. + intros (PRE & MEM & REG). + destruct i; simpl; unfold sist_set_local, ssem_internal, ssem_local; simpl; try_simplify_someHyps. +Qed. + +(** * Symbolic execution of the internal steps of a path *) +Fixpoint siexec_path (path:nat) (f: function) (st: sistate): option sistate := + match path with + | O => Some st + | S p => + SOME i <- (fn_code f)!(st.(si_pc)) IN + SOME st1 <- siexec_inst i st IN + siexec_path p f st1 + end. + +Lemma siexec_inst_add_exits i st st': + siexec_inst i st = Some st' -> + ( si_exits st' = si_exits st \/ exists ext, si_exits st' = ext :: si_exits st ). +Proof. + destruct i; simpl; intro SISTEP; inversion_clear SISTEP; unfold siexec_inst; simpl; (discriminate || eauto). +Qed. + +Lemma siexec_inst_preserves_allfu ge sp ext lx rs0 m0 st st' i: + all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs0 m0 -> + siexec_inst i st = Some st' -> + all_fallthrough_upto_exit ge sp ext lx (si_exits st') rs0 m0. +Proof. + intros ALLFU SISTEP. destruct ALLFU as (ISTAIL & ALLF). + constructor; eauto. + destruct i; simpl in SISTEP; inversion_clear SISTEP; simpl; (discriminate || eauto). +Qed. + +Lemma siexec_path_correct_false ge sp f rs0 m0 st' is: + forall path, + is.(icontinue)=false -> + forall st, ssem_internal ge sp st rs0 m0 is -> + siexec_path path f st = Some st' -> + ssem_internal ge sp st' rs0 m0 is. +Proof. + induction path; simpl. + - intros. congruence. + - intros ICF st SSEM STEQ'. + destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate]. + destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate]. + eapply IHpath. 3: eapply STEQ'. eauto. + unfold ssem_internal in SSEM. rewrite ICF in SSEM. + destruct SSEM as (ext & lx & SEXIT & ALLFU). + unfold ssem_internal. rewrite ICF. exists ext, lx. + constructor; auto. eapply siexec_inst_preserves_allfu; eauto. +Qed. + +Lemma siexec_path_preserves_sabort ge sp path f rs0 m0 st': forall st, + siexec_path path f st = Some st' -> + sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. +Proof. + Local Hint Resolve siexec_inst_preserves_sabort: core. + induction path; simpl. + + unfold sist_set_local; try_simplify_someHyps. + + intros st; inversion_SOME i. + inversion_SOME st1; eauto. +Qed. + +Lemma siexec_path_WF path f: forall st, + siexec_path path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None. +Proof. + induction path; simpl. + + unfold sist_set_local. intuition congruence. + + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try tauto. + destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl. + - intros; erewrite siexec_inst_default_succ; eauto. + - intros; erewrite siexec_inst_WF; eauto. +Qed. + +Lemma siexec_path_default_succ path f st': forall st, + siexec_path path f st = Some st' -> nth_default_succ (fn_code f) path st.(si_pc) = Some st'.(si_pc). +Proof. + induction path; simpl. + + unfold sist_set_local. intros st H. inversion_clear H; simpl; try congruence. + + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try congruence. + destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl; try congruence. + intros; erewrite siexec_inst_default_succ; eauto. +Qed. + +Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is, + is.(icontinue)=true -> + ssem_internal ge sp st rs0 m0 is -> + nth_default_succ (fn_code f) path st.(si_pc) <> None -> + ssem_internal_opt2 ge sp (siexec_path path f st) rs0 m0 + (isteps ge path f sp is.(irs) is.(imem) is.(ipc)) + . +Proof. + Local Hint Resolve siexec_path_correct_false siexec_path_preserves_sabort siexec_path_WF: core. + induction path; simpl. + + intros st is CONT INV WF; + unfold ssem_internal, sist_set_local in * |- *; + try_simplify_someHyps. simpl. + destruct is; simpl in * |- *; subst; intuition auto. + + intros st is CONT; unfold ssem_internal at 1; rewrite CONT. + intros (LOCAL & PC & NYE) WF. + rewrite <- PC. + inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. + exploit siexec_inst_correct; eauto. + inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. + - inversion_SOME is1; intros His1;rewrite His1; simpl. + * destruct (icontinue is1) eqn:CONT1. + (* icontinue is0 = true *) + intros; eapply IHpath; eauto. + destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps. + (* icontinue is0 = false -> EARLY EXIT *) + destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto. + destruct WF. erewrite siexec_inst_default_succ; eauto. + (* try_simplify_someHyps; eauto. *) + * destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto. + - intros His1;rewrite His1; simpl; auto. +Qed. + +(** REM: in the following two unused lemmas *) + +Lemma siexec_path_right_assoc_decompose f path: forall st st', + siexec_path (S path) f st = Some st' -> + exists st0, siexec_path path f st = Some st0 /\ siexec_path 1%nat f st0 = Some st'. +Proof. + induction path; simpl; eauto. + intros st st'. + inversion_SOME i1. + inversion_SOME st1. + try_simplify_someHyps; eauto. +Qed. + +Lemma siexec_path_right_assoc_compose f path: forall st st0 st', + siexec_path path f st = Some st0 -> + siexec_path 1%nat f st0 = Some st' -> + siexec_path (S path) f st = Some st'. +Proof. + induction path. + + intros st st0 st' H. simpl in H. + try_simplify_someHyps; auto. + + intros st st0 st'. + assert (X:exists x, x=(S path)); eauto. + destruct X as [x X]. + intros H1 H2. rewrite <- X. + generalize H1; clear H1. simpl. + inversion_SOME i1. intros Hi1; rewrite Hi1. + inversion_SOME st1. intros Hst1; rewrite Hst1. + subst; eauto. +Qed. + + +(** * Main function of the symbolic execution *) + +Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. + +Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}. + +Lemma init_ssem_internal ge sp pc rs m: ssem_internal ge sp (init_sistate pc) rs m (mk_istate true pc rs m). +Proof. + unfold ssem_internal, ssem_local, all_fallthrough; simpl. intuition. +Qed. + +Definition sexec (f: function) (pc:node): option sstate := + SOME path <- (fn_path f)!pc IN + SOME st <- siexec_path path.(psize) f (init_sistate pc) IN + SOME i <- (fn_code f)!(st.(si_pc)) IN + Some (match siexec_inst i st with + | Some st' => {| internal := st'; final := Snone |} + | None => {| internal := st; final := sexec_final i st.(si_local) |} + end). + +Lemma final_node_path_simpl f path pc: + (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(psize) pc <> None. +Proof. + intros; exploit final_node_path; eauto. + intros (i & NTH & DUM). + congruence. +Qed. + +Lemma symb_path_last_step i st st' ge pge stack (f:function) sp pc rs m t s: + (fn_code f) ! pc = Some i -> + pc = st.(si_pc) -> + siexec_inst i st = Some st' -> + path_last_step ge pge stack f sp pc rs m t s -> + exists mk_istate, + istep ge i sp rs m = Some mk_istate + /\ t = E0 + /\ s = (State stack f sp mk_istate.(ipc) mk_istate.(RTLpath.irs) mk_istate.(imem)). +Proof. + intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl. +Qed. + +(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) +(sexec is a correct over-approximation) +*) +Theorem sexec_correct f pc pge ge sp path stack rs m t s: + (fn_path f)!pc = Some path -> + path_step ge pge path.(psize) stack f sp rs m pc t s -> + exists st, sexec f pc = Some st /\ ssem pge ge sp st stack f rs m t s. +Proof. + Local Hint Resolve init_ssem_internal: core. + intros PATH STEP; unfold sexec; rewrite PATH; simpl. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. + { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. } + (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]). + destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST]; + (* intro Hst *) + (rewrite STEPS; unfold ssem_internal_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence); + (* intro SEM *) + (simpl; unfold ssem_internal; simpl; rewrite CONT; intro SEM); + (* intro Hi' *) + ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); + [ unfold nth_default_succ_inst in Hi; + exploit siexec_path_default_succ; eauto; simpl; + intros DEF; rewrite DEF in Hi; auto + | clear Hi; rewrite Hi' ]); + (* eexists *) + (eexists; constructor; eauto). + - (* early *) + eapply ssem_early; eauto. + unfold ssem_internal; simpl; rewrite CONT. + destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl; eauto. + destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx. + constructor; auto. eapply siexec_inst_preserves_allfu; eauto. + - destruct SEM as (SEM & PC & HNYE). + destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl. + + (* normal on Snone *) + rewrite <- PC in LAST. + exploit symb_path_last_step; eauto; simpl. + intros (mk_istate & ISTEP & Ht & Hs); subst. + exploit siexec_inst_correct; eauto. simpl. + erewrite Hst', ISTEP; simpl. + clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i. + intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT. + { (* icontinue mk_istate = true *) + eapply ssem_normal; simpl; eauto. + unfold ssem_internal in SEM. + rewrite CONT in SEM.eval_sval + destruct SEM as (SEM & PC & HNYE). + rewrite <- PC. + eapply exec_Snone. } + { eapply ssem_early; eauto. } + + (* normal non-Snone instruction *) + eapply ssem_normal; eauto. + * unfold ssem_internal; simpl; rewrite CONT; intuition. + * simpl. eapply sexec_final_correct; eauto. + rewrite PC; auto. +Qed. + +(* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *) +Inductive equiv_stackframe: stackframe -> stackframe -> Prop := + | equiv_stackframe_intro res f sp pc rs1 rs2 + (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + +Inductive equiv_state: state -> state -> Prop := + | State_equiv stack f sp pc rs1 m rs2 + (EQUIV: forall r, rs1#r = rs2#r): + equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) + | Call_equiv stk stk' f args m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Callstate stk f args m) (Callstate stk' f args m) + | Return_equiv stk stk' v m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Returnstate stk v m) (Returnstate stk' v m). + +Lemma equiv_stackframe_refl stf: equiv_stackframeeval_sval stf stf. +Proof. + destruct stf. constructor; auto. +Qed. + +Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. +Proof. + Local Hint Resolve equiv_stackframe_refl: core. + induction stk; simpl; constructor; auto. +Qed.eval_sval + +Lemma equiv_state_refl s: equiv_state s s. +Proof. + Local Hint Resolve equiv_stack_refl: core. + induction s; simpl; constructor; auto. +Qed. + +(* +Lemma equiv_stackframe_trans stf1 stf2 stf3: + equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3. +Proof. + destruct 1; intros EQ; inv EQ; try econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +Lemma equiv_stack_trans stk1 stk2: + list_forall2 equiv_stackframe stk1 stk2 -> + forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> + list_forall2 equiv_stackframe stk1 stk3. +Proof. + Local Hint Resolve equiv_stackframe_trans. + induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. +Qed. + +Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3. +Proof. + Local Hint Resolve equiv_stack_trans. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. +*) + +Lemma regmap_setres_eq (rs rs': regset) res vres: + (forall r, rs # r = rs' # r) -> + forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r. +Proof. + intros RSEQ r. destruct res; simpl; try congruence. + destruct (peq x r). + - subst. repeat (rewrite Regmap.gss). reflexivity. + - repeat (rewrite Regmap.gso); auto. +Qed. + +Lemma sem_sfval_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + sem_sfval pge ge sp st stack f rs0 m0 sv rs1 m t s -> + (forall r, rs1#r = rs2#r) -> + exists s', equiv_state s s' /\ sem_sfval pge ge sp st stack f rs0 m0 sv rs2 m t s'. +Proof. eval_sval + Local Hint Resolve equiv_stack_refl: core. + destruct 1. + - (* Snone *) intros; eexists; econstructor. + + eapply State_equiv; eauto. + + eapply exec_Snone. + - (* Scall *) + intros; eexists; econstructor. + 2: { eapply exec_Scall; eauto. } + apply Call_equiv; auto. + repeat (constructor; auto). + - (* Stailcall *) + intros; eexists; econstructor; [| eapply exec_Stailcall; eauto]. + apply Call_equiv; auto. + - (* Sbuiltin *) + intros; eexists; econstructor; [| eapply exec_Sbuiltin; eauto]. + constructor. eapply regmap_setres_eq; eauto. + - (* Sjumptable *) + intros; eexists; econstructor; [| eapply exec_Sjumptable; eauto]. + constructor. assumption. + - (* Sreturn *) + intros; eexists; econstructor; [| eapply exec_Sreturn; eauto]. + eapply equiv_state_refl; eauto. +Qed. + +Lemma siexec_inst_early_exit_absurd i st st' ge sp rs m rs' m' pc': + siexec_inst i st = Some st' -> + (exists ext lx, ssem_exit ge sp ext rs m rs' m' pc' /\ + all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m) -> + all_fallthrough ge sp (si_exits st') rs m -> + False. +Proof. + intros SIEXEC (ext & lx & SSEME & ALLFU) ALLF. destruct ALLFU as (TAIL & _). + exploit siexec_inst_add_exits; eauto. destruct 1 as [SIEQ | (ext0 & SIEQ)]. + - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. eassumption. + - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. + constructor. eassumption. +Qed. +eval_sval +Lemma is_tail_false {A: Type}: forall (l: list A) a, is_tail (a::l) nil -> False. +Proof. + intros. eapply is_tail_incl in H. unfold incl in H. pose (H a). + assert (In a (a::l)) by (constructor; auto). assert (In a nil) by auto. apply in_nil in H1. + contradiction. +Qed. + +Lemma cons_eq_false {A: Type}: forall (l: list A) a, + a :: l = l -> False. +Proof. + induction l; intros. + - discriminate. + - inv H. apply IHl in H2. contradiction. +Qed. + +Lemma app_cons_nil_eq {A: Type}: forall l' l (a:A), + (l' ++ a :: nil) ++ l = l' ++ a::l. +Proof. + induction l'; intros. + - simpl. reflexivity. + - simpl. rewrite IHl'. reflexivity. +Qed. + +Lemma app_eq_false {A: Type}: forall l (l': list A) a, + l' ++ a :: l = l -> False. +Proof. + induction l; intros. + - apply app_eq_nil in H. destruct H as (_ & H). apply cons_eq_false in H. contradiction. + - destruct l' as [|a' l']. + + simpl in H. apply cons_eq_false in H. contradiction. + + rewrite <- app_comm_cons in H. inv H. + apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption. +Qed. + +Lemma is_tail_false_gen {A: Type}: forall (l: list A) l' a, is_tail (l'++(a::l)) l -> False. +Proof. + induction l. + - intros. destruct l' as [|a' l']. + + simpl in H. apply is_tail_false in H. contradiction. + + rewrite <- app_comm_cons in H. apply is_tail_false in H. contradiction. + - intros. inv H. + + apply app_eq_false in H2. contradiction. + + apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption. +Qed. + +Lemma is_tail_eq {A: Type}: forall (l l': list A), + is_tail l' l -> + is_tail l l' -> + l = l'. +Proof. + destruct l as [|a l]; intros l' ITAIL ITAIL'. + - destruct l' as [|i' l']; auto. apply is_tail_false in ITAIL. contradiction. + - inv ITAIL; auto. + destruct l' as [|i' l']. { apply is_tail_false in ITAIL'. contradiction. } + exploit is_tail_trans. eapply ITAIL'. eauto. intro ABSURD. + apply (is_tail_false_gen l nil a) in ABSURD. contradiction. +Qed. + +(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution + (sexec is exact). +*) +Theorem sexec_exact f pc pge ge sp path stack st rs m t s1: + (fn_path f)!pc = Some path -> + sexec f pc = Some st -> + ssem pge ge sp st stack f rs m t s1 -> + exists s2, path_step ge pge path.(psize) stack f sp rs m pc t s2 /\ + equiv_state s1 s2. +Proof. + Local Hint Resolve init_ssem_internal: core. + unfold sexec; intros PATH SSTEP SEM; rewrite PATH in SSTEP. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. + { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. } + (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]). + unfold nth_default_succ_inst in Hi. + destruct (siexec_path path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl. + 2:{ (* absurd case *) + exploit siexec_path_WF; eauto. + simpl; intros NDS; rewrite NDS in Hi; congruence. } + exploit siexec_path_default_succ;eval_sval eauto; simpl. + intros NDS; rewrite NDS in Hi. + rewrite Hi in SSTEP. + intros ISTEPS. try_simplify_someHyps. + destruct (siexec_inst i st0) as [st'|] eqn:Hst'; simpl. + + (* exit on Snone instruction *) + assert (SEM': t = E0 /\ exists is, ssem_internal ge sp st' rs m is + /\ s1 = (State stack f sp (if (icontinue is) then (si_pc st') else (ipc is)) (irs is) (imem is))). + { destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. + - repeat (econstructor; eauto). + rewrite CONT; eauto. + - inversion SEM2. repeat (econstructor; eauto). + rewrite CONT; eauto. } + clear SEM; subst. destruct SEM' as [X (is & SEM & X')]; subst. + intros. + destruct (isteps ge (psize path) f sp rs m pc) as [is0|] eqn:RISTEPS; simpl in *. + * unfold ssem_internal in ISTEPS. destruct (icontinue is0) eqn: ICONT0. + ** (* icontinue is0=true: path_step by normal_exit *) + destruct ISTEPS as (SEMis0&H1&H2). + rewrite H1 in * |-. + exploit siexec_inst_correct; eauto. + rewrite Hst'; simpl. + intros; exploit ssem_internal_opt_determ; eauto. + destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). + eexists. econstructor 1. + *** eapply exec_normal_exit; eauto. + eapply exec_istate; eauto. + *** rewrite EQ1. + enough ((ipc st) = (if icontinue st then si_pc st' else ipc is)) as ->. + { rewrite EQ2, EQ4. eapply State_equiv; auto. } + destruct (icontinue st) eqn:ICONT; auto. + exploit siexec_inst_default_succ; eauto. + erewrite istep_normal_exit; eauto. + try_simplify_someHyps. + ** (* The concrete execution has not reached "i" => early exit *) + unfold ssem_internal in SEM. + destruct (icontinue is) eqn:ICONT. + { destruct SEM as (SEML & SIPC & ALLF). + exploit siexec_inst_early_exit_absurd; eauto. contradiction. } + + eexists. econstructor 1. + *** eapply exec_early_exit; eauto. + *** destruct ISTEPS as (ext & lx & SSEME & ALLFU). destruct SEM as (ext' & lx' & SSEME' & ALLFU'). + eapply siexec_inst_preserves_allfu in ALLFU; eauto. + exploit ssem_exit_fallthrough_upto_exit; eauto. + exploit ssem_exit_fallthrough_upto_exit. eapply SSEME. eapply ALLFU. eapply ALLFU'. + intros ITAIL ITAIL'. apply is_tail_eq in ITAIL; auto. clear ITAIL'. + inv ITAIL. exploit ssem_exit_determ. eapply SSEME. eapply SSEME'. intros (IPCEQ & IRSEQ & IMEMEQ). + rewrite <- IPCEQ. rewrite <- IMEMEQ. constructor. congruence. + * (* The concrete execution has not reached "i" => abort case *) + eapply siexec_inst_preserves_sabort in ISTEPS; eauto. + exploit ssem_internal_exclude_sabort; eauto. contradiction. + + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. + - (* early exit *) + intros. + exploit ssem_internal_opt_determ; eauto. + destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). + eexists. econstructor 1. + * eapply exec_early_exit; eauto. + * rewrite EQ2, EQ4; eapply State_equiv. auto. + - (* normal exit non-Snone instruction *) + intros. + exploit ssem_internal_opt_determ; eauto. + destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). + unfold ssem_internal in SEM1. + rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). + exploit sem_sfval_equiv; eauto. + clear SEM2; destruct 1 as (s' & Ms' & SEM2). + rewrite ! EQ4 in * |-; clear EQ4. + rewrite ! EQ2 in * |-; clear EQ2. + exists s'; intuition. + eapply exec_normal_exit; eauto. + eapply sexec_final_complete; eauto. + * congruence. + * unfold ssem_local in * |- *. + destruct SEM1 as (A & B & C). constructor; [|constructor]; eauto. + intro r. congruence. + * congruence. +Qed. + +Require Import RTLpathLivegen RTLpathLivegenproof. + +(** * DEFINITION OF SIMULATION BETWEEN (ABSTRACT) SYMBOLIC EXECUTIONS +*) + +Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := + is1.(icontinue) = is2.(icontinue) + /\ eqlive_reg alive is1.(irs) is2.(irs) + /\ is1.(imem) = is2.(imem). + +Definition istate_simu f (srce: PTree.t node) outframe is1 is2: Prop := + if is1.(icontinue) then + istate_simulive (fun r => Regset.In r outframe) srce is1 is2 + else + exists path, f.(fn_path)!(is1.(ipc)) = Some path + /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 + /\ srce!(is2.(ipc)) = Some is1.(ipc). + +Record simu_proof_context {f1: RTLpath.function} := { + liveness_hyps: liveness_ok_function f1; + the_ge1: RTL.genv; + the_ge2: RTL.genv; + genv_match: forall s, Genv.find_symbol the_ge1 s = Genv.find_symbol the_ge2 s; + the_sp: val; + the_rs0: regset; + the_m0: mem +}. +Arguments simu_proof_context: clear implicits. + +(* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) +Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) outframe (st1 st2: sistate) (ctx: simu_proof_context f): Prop := + forall is1, ssem_internal (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) is1 -> + exists is2, ssem_internal (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) is2 + /\ istate_simu f dm outframe is1 is2. + +Inductive svident_simu (f: RTLpath.function) (ctx: simu_proof_context f): (sval + ident) -> (sval + ident) -> Prop := + | Sleft_simu sv1 sv2: + (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) + -> svident_simu f ctx (inl sv1) (inl sv2) + | Sright_simu id1 id2: + id1 = id2 + -> svident_simu f ctx (inr id1) (inr id2) + . + + +Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := + match lp with + | nil => Some nil + | p1::lp => SOME p2 <- pt!p1 IN + SOME lp2 <- (ptree_get_list pt lp) IN + Some (p2 :: lp2) + end. + +Lemma ptree_get_list_nth dm p2: forall lp2 lp1, + ptree_get_list dm lp2 = Some lp1 -> + forall n, list_nth_z lp2 n = Some p2 -> + exists p1, + list_nth_z lp1 n = Some p1 /\ dm ! p2 = Some p1. +Proof. + induction lp2. + - simpl. intros. inv H. simpl in *. discriminate. + - intros lp1 PGL n LNZ. simpl in PGL. explore. + inv LNZ. destruct (zeq n 0) eqn:ZEQ. + + subst. inv H0. exists n0. simpl; constructor; auto. + + exploit IHlp2; eauto. intros (p1 & LNZ & DMEQ). + eexists. simpl. rewrite ZEQ. + constructor; eauto. +Qed. + +Lemma ptree_get_list_nth_rev dm p1: forall lp2 lp1, + ptree_get_list dm lp2 = Some lp1 -> + forall n, list_nth_z lp1 n = Some p1 -> + exists p2, + list_nth_z lp2 n = Some p2 /\ dm ! p2 = Some p1. +Proof. + induction lp2. + - simpl. intros. inv H. simpl in *. discriminate. + - intros lp1 PGL n LNZ. simpl in PGL. explore. + inv LNZ. destruct (zeq n 0) eqn:ZEQ. + + subst. inv H0. exists a. simpl; constructor; auto. + + exploit IHlp2; eauto. intros (p2 & LNZ & DMEQ). + eexists. simpl. rewrite ZEQ. + constructor; eauto. congruence. +Qed. + +(* NOTE: we need to mix semantical simulation and syntactic definition on [sfval] in order to abstract the [match_states] *) +Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node) (ctx: simu_proof_context f): sfval -> sfval -> Prop := + | Snone_simu: + dm!opc2 = Some opc1 -> + sfval_simu dm f opc1 opc2 ctx Snone Snone + | Scall_simu sig svos1 svos2 lsv1 lsv2 res pc1 pc2: + dm!pc2 = Some pc1 -> + svident_simu f ctx svos1 svos2 -> + (eval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) + = (eval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) -> + sfval_simu dm f opc1 opc2 ctx (Scall sig svos1 lsv1 res pc1) (Scall sig svos2 lsv2 res pc2) + | Stailcall_simu sig svos1 svos2 lsv1 lsv2: + svident_simu f ctx svos1 svos2 -> + (eval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) + = (eval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) -> + sfval_simu dm f opc1 opc2 ctx (Stailcall sig svos1 lsv1) (Stailcall sig svos2 lsv2) + | Sbuiltin_simu ef lbs1 lbs2 br pc1 pc2: + dm!pc2 = Some pc1 -> + (eval_list_builtin_sval (the_ge1 ctx) (the_sp ctx) lbs1 (the_rs0 ctx) (the_m0 ctx)) + = (eval_list_builtin_sval (the_ge2 ctx) (the_sp ctx) lbs2 (the_rs0 ctx) (the_m0 ctx)) -> + sfval_simu dm f opc1 opc2 ctx (Sbuiltin ef lbs1 br pc1) (Sbuiltin ef lbs2 br pc2) + | Sjumptable_simu sv1 sv2 lpc1 lpc2: + ptree_get_list dm lpc2 = Some lpc1 -> + (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) + = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) -> + sfval_simu dm f opc1 opc2 ctx (Sjumptable sv1 lpc1) (Sjumptable sv2 lpc2) + | Sreturn_simu_none: sfval_simu dm f opc1 opc2 ctx (Sreturn None) (Sreturn None) + | Sreturn_simu_some sv1 sv2: + (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) + = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) -> + sfval_simu dm f opc1 opc2 ctx (Sreturn (Some sv1)) (Sreturn (Some sv2)). + +Definition sstate_simu dm f outframe (s1 s2: sstate) (ctx: simu_proof_context f): Prop := + sistate_simu dm f outframe s1.(internal) s2.(internal) ctx + /\ forall is1, + ssem_internal (the_ge1 ctx) (the_sp ctx) s1 (the_rs0 ctx) (the_m0 ctx) is1 -> + is1.(icontinue) = true -> + sfval_simu dm f s1.(si_pc) s2.(si_pc) ctx s1.(final) s2.(final). + +Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := + forall st1, sexec f1 pc1 = Some st1 -> + exists path st2, (fn_path f1)!pc1 = Some path /\ sexec f2 pc2 = Some st2 + /\ forall ctx, sstate_simu dm f1 path.(pre_output_regs) st1 st2 ctx. + +*) \ No newline at end of file -- cgit From fd74f68479c340351641093e5aa9a884f74d3651 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 6 May 2021 18:58:28 +0200 Subject: refactorisation + 1ere version de sstate --- scheduling/BTL_SEtheory.v | 1630 +++++---------------------------------------- 1 file changed, 169 insertions(+), 1461 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 53c00c8b..7025b90c 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -13,6 +13,16 @@ Require Import RTL BTL OptionMonad. Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) + +Record iblock_exec_context := { + cge: BTL.genv; + csp: val; + cstk: list stackframe; + cf: function; + crs0: regset; + cm0: mem +}. + (** * Syntax and semantics of symbolic values *) (* symbolic value *) @@ -40,26 +50,26 @@ Fixpoint list_sval_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. -Fixpoint eval_sval (ge: BTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := +Fixpoint eval_sval ctx (sv: sval): option val := match sv with - | Sinput r => Some (rs0#r) + | Sinput r => Some ((crs0 ctx)#r) | Sop op l sm => - SOME args <- eval_list_sval ge sp l rs0 m0 IN - SOME m <- eval_smem ge sp sm rs0 m0 IN - eval_operation ge sp op args m + SOME args <- eval_list_sval ctx l IN + SOME m <- eval_smem ctx sm IN + eval_operation (cge ctx) (csp ctx) op args m | Sload sm trap chunk addr lsv => match trap with | TRAP => - SOME args <- eval_list_sval ge sp lsv rs0 m0 IN - SOME a <- eval_addressing ge sp addr args IN - SOME m <- eval_smem ge sp sm rs0 m0 IN + SOME args <- eval_list_sval ctx lsv IN + SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN + SOME m <- eval_smem ctx sm IN Mem.loadv chunk m a | NOTRAP => - SOME args <- eval_list_sval ge sp lsv rs0 m0 IN - match (eval_addressing ge sp addr args) with + SOME args <- eval_list_sval ctx lsv IN + match (eval_addressing (cge ctx) (csp ctx) addr args) with | None => Some (default_notrap_load_value chunk) | Some a => - SOME m <- eval_smem ge sp sm rs0 m0 IN + SOME m <- eval_smem ctx sm IN match (Mem.loadv chunk m a) with | None => Some (default_notrap_load_value chunk) | Some val => Some val @@ -67,52 +77,49 @@ Fixpoint eval_sval (ge: BTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): o end end end -with eval_list_sval (ge: BTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := +with eval_list_sval ctx (lsv: list_sval): option (list val) := match lsv with | Snil => Some nil | Scons sv lsv' => - SOME v <- eval_sval ge sp sv rs0 m0 IN - SOME lv <- eval_list_sval ge sp lsv' rs0 m0 IN + SOME v <- eval_sval ctx sv IN + SOME lv <- eval_list_sval ctx lsv' IN Some (v::lv) end -with eval_smem (ge: BTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem := +with eval_smem ctx (sm: smem): option mem := match sm with - | Sinit => Some m0 + | Sinit => Some (cm0 ctx) | Sstore sm chunk addr lsv srce => - SOME args <- eval_list_sval ge sp lsv rs0 m0 IN - SOME a <- eval_addressing ge sp addr args IN - SOME m <- eval_smem ge sp sm rs0 m0 IN - SOME sv <- eval_sval ge sp srce rs0 m0 IN + SOME args <- eval_list_sval ctx lsv IN + SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN + SOME m <- eval_smem ctx sm IN + SOME sv <- eval_sval ctx srce IN Mem.storev chunk m a sv end. -Lemma eval_list_sval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: - (forall r : reg, eval_sval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> - eval_list_sval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Lemma eval_list_sval_inj ctx l (sreg: reg -> sval) rs: + (forall r : reg, eval_sval ctx (sreg r) = Some (rs # r)) -> + eval_list_sval ctx (list_sval_inj (map sreg l)) = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; repeat autodestruct; auto. Qed. -Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := - SOME args <- eval_list_sval ge sp lsv rs0 m0 IN - SOME m <- eval_smem ge sp sm rs0 m0 IN +Definition seval_condition ctx (cond: condition) (lsv: list_sval) (sm: smem) : option bool := + SOME args <- eval_list_sval ctx lsv IN + SOME m <- eval_smem ctx sm IN eval_condition cond args m. (** * Auxiliary definitions on Builtins *) -(* TODO: clean this. Some generic stuffs could be put in [AST.v] *) +(* TODO: clean this. Some (cge ctx)neric stuffs could be put in [AST.v] *) Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *) -Variable ge: BTL.genv. -Variable sp: val. +Variable ctx: iblock_exec_context. Variable m: mem. -Variable rs0: regset. -Variable m0: mem. Inductive seval_builtin_arg: builtin_arg sval -> val -> Prop := | seval_BA: forall x v, - eval_sval ge sp x rs0 m0 = Some v -> + eval_sval ctx x = Some v -> seval_builtin_arg (BA x) v | seval_BA_int: forall n, seval_builtin_arg (BA_int n) (Vint n) @@ -123,22 +130,23 @@ Inductive seval_builtin_arg: builtin_arg sval -> val -> Prop := | seval_BA_single: forall n, seval_builtin_arg (BA_single n) (Vsingle n) | seval_BA_loadstack: forall chunk ofs v, - Mem.loadv chunk m (Val.offset_ptr sp ofs) = Some v -> + Mem.loadv chunk m (Val.offset_ptr (csp ctx) ofs) = Some v -> seval_builtin_arg (BA_loadstack chunk ofs) v | seval_BA_addrstack: forall ofs, - seval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr sp ofs) + seval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr (csp ctx) ofs) | seval_BA_loadglobal: forall chunk id ofs v, - Mem.loadv chunk m (Senv.symbol_address ge id ofs) = Some v -> + Mem.loadv chunk m (Senv.symbol_address (cge ctx) id ofs) = Some v -> seval_builtin_arg (BA_loadglobal chunk id ofs) v | seval_BA_addrglobal: forall id ofs, - seval_builtin_arg (BA_addrglobal id ofs) (Senv.symbol_address ge id ofs) + seval_builtin_arg (BA_addrglobal id ofs) (Senv.symbol_address (cge ctx) id ofs) | seval_BA_splitlong: forall hi lo vhi vlo, seval_builtin_arg hi vhi -> seval_builtin_arg lo vlo -> seval_builtin_arg (BA_splitlong hi lo) (Val.longofwords vhi vlo) | seval_BA_addptr: forall a1 a2 v1 v2, seval_builtin_arg a1 v1 -> seval_builtin_arg a2 v2 -> seval_builtin_arg (BA_addptr a1 a2) - (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2). + (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2) +. Definition seval_builtin_args (al: list (builtin_arg sval)) (vl: list val) : Prop := list_forall2 seval_builtin_arg al vl. @@ -159,7 +167,7 @@ Qed. End SEVAL_BUILTIN_ARG. -(* NB: generic function that could be put into [AST] file *) +(* NB: (cge ctx)neric function that could be put into [AST] file *) Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B := match arg with | BA x => BA (f x) @@ -175,10 +183,10 @@ Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2) end. -Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg: forall arg varg, - (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - eval_builtin_arg ge (fun r => rs # r) sp m arg varg -> - seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg. +Lemma seval_builtin_arg_correct ctx rs m sreg: forall arg varg, + (forall r, eval_sval ctx (sreg r) = Some rs # r) -> + eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg -> + seval_builtin_arg ctx m (builtin_arg_map sreg arg) varg. Proof. induction arg. all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence). @@ -188,10 +196,10 @@ Proof. eapply IHarg1; eauto. eapply IHarg2; eauto. Qed. -Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs: - (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - eval_builtin_args ge (fun r => rs # r) sp m args vargs -> - seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs. +Lemma seval_builtin_args_correct ctx rs m sreg args vargs: + (forall r, eval_sval ctx (sreg r) = Some rs # r) -> + eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs -> + seval_builtin_args ctx m (map (builtin_arg_map sreg) args) vargs. Proof. induction 2. - constructor. @@ -199,10 +207,10 @@ Proof. eapply seval_builtin_arg_correct; eauto. Qed. -Lemma seval_builtin_arg_complete ge sp rs m rs0 m0 sreg: forall arg varg, - (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg -> - eval_builtin_arg ge (fun r => rs # r) sp m arg varg. +Lemma seval_builtin_arg_complete ctx rs m sreg: forall arg varg, + (forall r, eval_sval ctx (sreg r) = Some rs # r) -> + seval_builtin_arg ctx m (builtin_arg_map sreg arg) varg -> + eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg. Proof. induction arg. all: intros varg SEVAL BARG; try (inv BARG; constructor; congruence). @@ -213,10 +221,10 @@ Proof. eapply IHarg1; eauto. eapply IHarg2; eauto. Qed. -Lemma seval_builtin_args_complete ge sp rs m rs0 m0 sreg: forall args vargs, - (forall r, eval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs -> - eval_builtin_args ge (fun r => rs # r) sp m args vargs. +Lemma seval_builtin_args_complete ctx rs m sreg: forall args vargs, + (forall r, eval_sval ctx (sreg r) = Some rs # r) -> + seval_builtin_args ctx m (map (builtin_arg_map sreg) args) vargs -> + eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs. Proof. induction args. - simpl. intros. inv H0. constructor. @@ -225,16 +233,16 @@ Proof. eapply seval_builtin_arg_complete; eauto. Qed. -Fixpoint seval_builtin_sval ge sp bsv rs0 m0 := +Fixpoint seval_builtin_sval ctx bsv := match bsv with - | BA sv => SOME v <- eval_sval ge sp sv rs0 m0 IN Some (BA v) + | BA sv => SOME v <- eval_sval ctx sv IN Some (BA v) | BA_splitlong sv1 sv2 => - SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN - SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN + SOME v1 <- seval_builtin_sval ctx sv1 IN + SOME v2 <- seval_builtin_sval ctx sv2 IN Some (BA_splitlong v1 v2) | BA_addptr sv1 sv2 => - SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN - SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN + SOME v1 <- seval_builtin_sval ctx sv1 IN + SOME v2 <- seval_builtin_sval ctx sv2 IN Some (BA_addptr v1 v2) | BA_int i => Some (BA_int i) | BA_long l => Some (BA_long l) @@ -246,56 +254,54 @@ Fixpoint seval_builtin_sval ge sp bsv rs0 m0 := | BA_addrglobal id ptr => Some (BA_addrglobal id ptr) end. -Fixpoint eval_list_builtin_sval ge sp lbsv rs0 m0 := +Fixpoint eval_list_builtin_sval ctx lbsv := match lbsv with | nil => Some nil - | bsv::lbsv => SOME v <- seval_builtin_sval ge sp bsv rs0 m0 IN - SOME lv <- eval_list_builtin_sval ge sp lbsv rs0 m0 IN + | bsv::lbsv => SOME v <- seval_builtin_sval ctx bsv IN + SOME lv <- eval_list_builtin_sval ctx lbsv IN Some (v::lv) end. -Lemma eval_list_builtin_sval_nil ge sp rs0 m0 lbs2: - eval_list_builtin_sval ge sp lbs2 rs0 m0 = Some nil -> +Lemma eval_list_builtin_sval_nil ctx lbs2: + eval_list_builtin_sval ctx lbs2 = Some nil -> lbs2 = nil. Proof. - destruct lbs2; simpl; auto. - intros. destruct (seval_builtin_sval _ _ _ _ _); - try destruct (eval_list_builtin_sval _ _ _ _ _); discriminate. + destruct lbs2; simpl; repeat autodestruct; congruence. Qed. -Lemma seval_builtin_sval_arg ge sp rs0 m0 bs: +Lemma seval_builtin_sval_arg ctx bs: forall ba m v, - seval_builtin_sval ge sp bs rs0 m0 = Some ba -> - eval_builtin_arg ge (fun id => id) sp m ba v -> - seval_builtin_arg ge sp m rs0 m0 bs v. + seval_builtin_sval ctx bs = Some ba -> + eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v -> + seval_builtin_arg ctx m bs v. Proof. induction bs; simpl; try (intros ba m v H; inversion H; subst; clear H; intros H; inversion H; subst; econstructor; auto; fail). - - intros ba m v; destruct (eval_sval _ _ _ _ _) eqn: SV; + - intros ba m v; destruct (eval_sval _ _) eqn: SV; intros H; inversion H; subst; clear H. intros H; inversion H; subst. econstructor; auto. - intros ba m v. - destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence. - destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence. + destruct (seval_builtin_sval _ bs1) eqn: SV1; try congruence. + destruct (seval_builtin_sval _ bs2) eqn: SV2; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst. econstructor; eauto. - intros ba m v. - destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence. - destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence. + destruct (seval_builtin_sval _ bs1) eqn: SV1; try congruence. + destruct (seval_builtin_sval _ bs2) eqn: SV2; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst. econstructor; eauto. Qed. -Lemma seval_builtin_arg_sval ge sp m rs0 m0 v: forall bs, - seval_builtin_arg ge sp m rs0 m0 bs v -> +Lemma seval_builtin_arg_sval ctx m v: forall bs, + seval_builtin_arg ctx m bs v -> exists ba, - seval_builtin_sval ge sp bs rs0 m0 = Some ba - /\ eval_builtin_arg ge (fun id => id) sp m ba v. + seval_builtin_sval ctx bs = Some ba + /\ eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v. Proof. induction 1. all: try (eexists; constructor; [simpl; reflexivity | constructor]). @@ -315,28 +321,28 @@ Proof. + constructor; assumption. Qed. -Lemma seval_builtin_sval_args ge sp rs0 m0 lbs: +Lemma seval_builtin_sval_args ctx lbs: forall lba m v, - eval_list_builtin_sval ge sp lbs rs0 m0 = Some lba -> - list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba v -> - seval_builtin_args ge sp m rs0 m0 lbs v. + eval_list_builtin_sval ctx lbs = Some lba -> + list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba v -> + seval_builtin_args ctx m lbs v. Proof. unfold seval_builtin_args; induction lbs; simpl; intros lba m v. - intros H; inversion H; subst; clear H. intros H; inversion H. econstructor. - - destruct (seval_builtin_sval _ _ _ _ _) eqn:SV; try congruence. - destruct (eval_list_builtin_sval _ _ _ _ _) eqn: SVL; try congruence. + - destruct (seval_builtin_sval _ _) eqn:SV; try congruence. + destruct (eval_list_builtin_sval _ _) eqn: SVL; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst; clear H. econstructor; eauto. eapply seval_builtin_sval_arg; eauto. Qed. -Lemma seval_builtin_args_sval ge sp m rs0 m0 lv: forall lbs, - seval_builtin_args ge sp m rs0 m0 lbs lv -> +Lemma seval_builtin_args_sval ctx m lv: forall lbs, + seval_builtin_args ctx m lbs lv -> exists lba, - eval_list_builtin_sval ge sp lbs rs0 m0 = Some lba - /\ list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba lv. + eval_list_builtin_sval ctx lbs = Some lba + /\ list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba lv. Proof. induction 1. - eexists. constructor. @@ -349,10 +355,10 @@ Proof. + constructor; assumption. Qed. -Lemma seval_builtin_sval_correct ge sp m rs0 m0: forall bs1 v bs2, - seval_builtin_arg ge sp m rs0 m0 bs1 v -> - (seval_builtin_sval ge sp bs1 rs0 m0) = (seval_builtin_sval ge sp bs2 rs0 m0) -> - seval_builtin_arg ge sp m rs0 m0 bs2 v. +Lemma seval_builtin_sval_correct ctx m: forall bs1 v bs2, + seval_builtin_arg ctx m bs1 v -> + (seval_builtin_sval ctx bs1) = (seval_builtin_sval ctx bs2) -> + seval_builtin_arg ctx m bs2 v. Proof. intros. exploit seval_builtin_arg_sval; eauto. intros (ba & X1 & X2). @@ -360,10 +366,10 @@ Proof. congruence. Qed. -Lemma eval_list_builtin_sval_correct ge sp m rs0 m0 vargs: forall lbs1, - seval_builtin_args ge sp m rs0 m0 lbs1 vargs -> - forall lbs2, (eval_list_builtin_sval ge sp lbs1 rs0 m0) = (eval_list_builtin_sval ge sp lbs2 rs0 m0) -> - seval_builtin_args ge sp m rs0 m0 lbs2 vargs. +Lemma eval_list_builtin_sval_correct ctx m vargs: forall lbs1, + seval_builtin_args ctx m lbs1 vargs -> + forall lbs2, (eval_list_builtin_sval ctx lbs1) = (eval_list_builtin_sval ctx lbs2) -> + seval_builtin_args ctx m lbs2 vargs. Proof. intros. exploit seval_builtin_args_sval; eauto. intros (ba & X1 & X2). @@ -383,164 +389,64 @@ Inductive sfval := | Sreturn: option sval -> sfval . -Definition sfind_function (ge: BTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := +Definition sfind_function ctx (svos : sval + ident): option fundef := match svos with - | inl sv => SOME v <- eval_sval ge sp sv rs0 m0 IN Genv.find_funct ge v - | inr symb => SOME b <- Genv.find_symbol ge symb IN Genv.find_funct_ptr ge b + | inl sv => SOME v <- eval_sval ctx sv IN Genv.find_funct (cge ctx) v + | inr symb => SOME b <- Genv.find_symbol (cge ctx) symb IN Genv.find_funct_ptr (cge ctx) b end . -Inductive sem_sfval (ge: BTL.genv) (sp:val) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := | exec_Sgoto pc rs m: - sem_sfval ge sp stack f rs0 m0 (Sgoto pc) rs m E0 (State stack f sp pc rs m) + sem_sfval ctx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc rs m) | exec_Sreturn stk osv rs m m' v: - sp = (Vptr stk Ptrofs.zero) -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - match osv with Some sv => eval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> - sem_sfval ge sp stack f rs0 m0 (Sreturn osv) rs m - E0 (Returnstate stack v m') + (csp ctx) = (Vptr stk Ptrofs.zero) -> + Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> + match osv with Some sv => eval_sval ctx sv | None => Some Vundef end = Some v -> + sem_sfval ctx (Sreturn osv) rs m + E0 (Returnstate (cstk ctx) v m') | exec_Scall rs m sig svos lsv args res pc fd: - sfind_function ge sp svos rs0 m0 = Some fd -> + sfind_function ctx svos = Some fd -> funsig fd = sig -> - eval_list_sval ge sp lsv rs0 m0 = Some args -> - sem_sfval ge sp stack f rs0 m0 (Scall sig svos lsv res pc) rs m - E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) + eval_list_sval ctx lsv = Some args -> + sem_sfval ctx (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc rs :: (cstk ctx)) fd args m) | exec_Stailcall stk rs m sig svos args fd m' lsv: - sfind_function ge sp svos rs0 m0 = Some fd -> + sfind_function ctx svos = Some fd -> funsig fd = sig -> - sp = Vptr stk Ptrofs.zero -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - eval_list_sval ge sp lsv rs0 m0 = Some args -> - sem_sfval ge sp stack f rs0 m0 (Stailcall sig svos lsv) rs m - E0 (Callstate stack fd args m') + (csp ctx) = Vptr stk Ptrofs.zero -> + Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> + eval_list_sval ctx lsv = Some args -> + sem_sfval ctx (Stailcall sig svos lsv) rs m + E0 (Callstate (cstk ctx) fd args m') | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: - seval_builtin_args ge sp m rs0 m0 sargs vargs -> - external_call ef ge vargs m t vres m' -> - sem_sfval ge sp stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m - t (State stack f sp pc (regmap_setres res vres rs) m') + seval_builtin_args ctx m sargs vargs -> + external_call ef (cge ctx) vargs m t vres m' -> + sem_sfval ctx (Sbuiltin ef sargs res pc) rs m + t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres rs) m') | exec_Sjumptable sv tbl pc' n rs m: - eval_sval ge sp sv rs0 m0 = Some (Vint n) -> + eval_sval ctx sv = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - sem_sfval ge sp stack f rs0 m0 (Sjumptable sv tbl) rs m - E0 (State stack f sp pc' rs m) + sem_sfval ctx (Sjumptable sv tbl) rs m + E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' rs m) . -(** * Preservation properties *) - -Section SymbValPreserved. - -Variable ge ge': BTL.genv. - -Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. - -Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. - -Lemma senv_find_symbol_preserved id: - Senv.find_symbol ge id = Senv.find_symbol ge' id. -Proof. - destruct senv_preserved_BTL as (A & B & C). congruence. -Qed. - -Lemma senv_symbol_address_preserved id ofs: - Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. -Proof. - unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. - reflexivity. -Qed. - -Lemma eval_sval_preserved sp sv rs0 m0: - eval_sval ge sp sv rs0 m0 = eval_sval ge' sp sv rs0 m0. -Proof. - induction sv using sval_mut with (P0 := fun lsv => eval_list_sval ge sp lsv rs0 m0 = eval_list_sval ge' sp lsv rs0 m0) - (P1 := fun sm => eval_smem ge sp sm rs0 m0 = eval_smem ge' sp sm rs0 m0); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _ _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _ _ _); auto. - erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _ _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (eval_sval _ _ _ _); auto. - rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _ _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (eval_smem _ _ _ _); auto. - rewrite IHsv1; auto. -Qed. - -Lemma seval_builtin_arg_preserved sp m rs0 m0: - forall bs varg, - seval_builtin_arg ge sp m rs0 m0 bs varg -> - seval_builtin_arg ge' sp m rs0 m0 bs varg. -Proof. - induction 1. - all: try (constructor; auto). - - rewrite <- eval_sval_preserved. assumption. - - rewrite <- senv_symbol_address_preserved. assumption. - - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. -Qed. - -Lemma seval_builtin_args_preserved sp m rs0 m0 lbs vargs: - seval_builtin_args ge sp m rs0 m0 lbs vargs -> - seval_builtin_args ge' sp m rs0 m0 lbs vargs. -Proof. - induction 1; constructor; eauto. - eapply seval_builtin_arg_preserved; auto. -Qed. - -Lemma list_sval_eval_preserved sp lsv rs0 m0: - eval_list_sval ge sp lsv rs0 m0 = eval_list_sval ge' sp lsv rs0 m0. -Proof. - induction lsv; simpl; auto. - rewrite eval_sval_preserved. destruct (eval_sval _ _ _ _); auto. - rewrite IHlsv; auto. -Qed. - -Lemma smem_eval_preserved sp sm rs0 m0: - eval_smem ge sp sm rs0 m0 = eval_smem ge' sp sm rs0 m0. -Proof. - induction sm; simpl; auto. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _ _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsm; clear IHsm. destruct (eval_smem _ _ _ _); auto. - rewrite eval_sval_preserved; auto. -Qed. - -Lemma seval_condition_preserved sp cond lsv sm rs0 m0: - seval_condition ge sp cond lsv sm rs0 m0 = seval_condition ge' sp cond lsv sm rs0 m0. -Proof. - unfold seval_condition. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _ _ _); auto. - rewrite smem_eval_preserved; auto. -Qed. - -End SymbValPreserved. - (* Syntax and Semantics of symbolic internal states *) -(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) -Record sistate := { si_pre: BTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }. +(* [si_pre] is a precondition on initial context *) +Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg: reg -> sval; si_smem: smem }. -(* Predicate on which (rs, m) is a possible final state after evaluating [st] on (rs0, m0) *) -Definition sem_sistate (ge: BTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := - st.(si_pre) ge sp rs0 m0 - /\ eval_smem ge sp st.(si_smem) rs0 m0 = Some m - /\ forall (r:reg), eval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r). +(* Predicate on which (rs, m) is a possible final state after evaluating [st] on ((crs0 ctx), (cm0 ctx)) *) +Definition sem_sistate ctx (st: sistate) (rs: regset) (m: mem): Prop := + st.(si_pre) ctx + /\ eval_smem ctx st.(si_smem) = Some m + /\ forall (r:reg), eval_sval ctx (st.(si_sreg) r) = Some (rs#r). -Definition abort_sistate (ge: BTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := - ~(st.(si_pre) ge sp rs0 m0) - \/ eval_smem ge sp st.(si_smem) rs0 m0 = None - \/ exists (r: reg), eval_sval ge sp (st.(si_sreg) r) rs0 m0 = None. - -Record sstate_exit := { se_sis:> sistate; se_sfv: sfval }. - -Definition sem_sexit ge (sp:val) (st: sstate_exit) stack f (rs0: regset) (m0: mem) rs m (t: trace) (s:BTL.state) := - sem_sistate ge sp st rs0 m0 rs m /\ - sem_sfval ge sp stack f rs0 m0 st.(se_sfv) rs m t s - . +Definition abort_sistate ctx (st: sistate): Prop := + ~(st.(si_pre) ctx) + \/ eval_smem ctx st.(si_smem) = None + \/ exists (r: reg), eval_sval ctx (st.(si_sreg) r) = None. (** * Symbolic execution of final step *) Definition sexec_final_sfv (i: final) (sis: sistate): sfval := @@ -567,10 +473,10 @@ Definition sexec_final_sfv (i: final) (sis: sistate): sfval := Local Hint Constructors sem_sfval: core. -Lemma sexec_final_svf_correct ge sp i f sis stack rs0 m0 t rs m s: - sem_sistate ge sp sis rs0 m0 rs m -> - final_step ge stack f sp rs m i t s -> - sem_sfval ge sp stack f rs0 m0 (sexec_final_sfv i sis) rs m t s. +Lemma sexec_final_svf_correct ctx i sis t rs m s: + sem_sistate ctx sis rs m -> + final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> + sem_sfval ctx (sexec_final_sfv i sis) rs m t s. Proof. intros (PRE&MEM®). destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. @@ -578,9 +484,9 @@ Proof. - destruct ros; simpl in * |- *; auto. rewrite REG; auto. - erewrite eval_list_sval_inj; simpl; auto. - + (* Btailcall *) intros. eapply exec_Stailcall; auto. - - destruct ros; simpl in * |- *; auto. - rewrite REG; auto. + + (* Btailcall *) intros. eapply exec_Stailcall; eauto. + - destruct ros; simpl in * |- *; eauto. + rewrite REG; eauto. - erewrite eval_list_sval_inj; simpl; auto. + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto. eapply seval_builtin_args_correct; eauto. @@ -590,10 +496,10 @@ Qed. Local Hint Constructors final_step: core. Local Hint Resolve seval_builtin_args_complete: core. -Lemma sexec_final_svf_complete ge sp i f sis stack rs0 m0 t rs m s: - sem_sistate ge sp sis rs0 m0 rs m -> - sem_sfval ge sp stack f rs0 m0 (sexec_final_sfv i sis) rs m t s - -> final_step ge stack f sp rs m i t s. +Lemma sexec_final_svf_complete ctx i sis t rs m s: + sem_sistate ctx sis rs m -> + sem_sfval ctx (sexec_final_sfv i sis) rs m t s + -> final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. Proof. intros (PRE&MEM®). destruct i; simpl; intros LAST; inv LAST; eauto. @@ -615,1217 +521,19 @@ Proof. congruence. Qed. - -(* TODO: clean/recover this COPY-PASTE OF RTLpathSE_theory.v - - -(* Syntax and semantics of symbolic exit states *) - -Definition all_fallthrough ge sp (lx: list sistate_exit) rs0 m0: Prop := - forall ext, List.In ext lx -> - seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some false. - -Lemma all_fallthrough_revcons ge sp ext rs m lx: - all_fallthrough ge sp (ext::lx) rs m -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false - /\ all_fallthrough ge sp lx rs m. -Proof. - intros ALLFU. constructor. - - assert (In ext (ext::lx)) by (constructor; auto). apply ALLFU in H. assumption. - - intros ext' INEXT. assert (In ext' (ext::lx)) by (apply in_cons; auto). - apply ALLFU in H. assumption. -Qed. - -(** Semantic of an exit in pseudo code: - if si_cond (si_condargs) - si_elocal; goto if_so - else () -*) - -Definition ssem_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true - /\ ssem_local ge sp (si_elocal ext) rs m rs' m' - /\ (si_ifso ext) = pc'. - -(* Either an abort on the condition evaluation OR an abort on the sistate_local IF the condition was true *) -Definition sabort_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) : Prop := - let sev_cond := seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m in - sev_cond = None - \/ (sev_cond = Some true /\ sabort_local ge sp ext.(si_elocal) rs m). - -(** * Syntax and Semantics of symbolic internal state *) -Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. - -Definition all_fallthrough_upto_exit ge sp ext lx' lx rs m : Prop := - is_tail (ext::lx') lx /\ all_fallthrough ge sp lx' rs m. - -(** Semantic of a sistate in pseudo code: - si_exit1; si_exit2; ...; si_exitn; - si_local; goto si_pc *) - -(* Note: in RTLpath, is.(icontinue) = false iff we took an early exit *) - -Definition ssem_internal (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop := - if (is.(icontinue)) - then - ssem_local ge sp st.(si_local) rs m is.(irs) is.(imem) - /\ st.(si_pc) = is.(ipc) - /\ all_fallthrough ge sp st.(si_exits) rs m - else exists ext lx, - ssem_exit ge sp ext rs m is.(irs) is.(imem) is.(ipc) - /\ all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m. - -Definition sabort (ge: RTL.genv) (sp: val) (st: sistate) (rs: regset) (m: mem): Prop := - (* No early exit was met but we aborted on the si_local *) - (all_fallthrough ge sp st.(si_exits) rs m /\ sabort_local ge sp st.(si_local) rs m) - (* OR we aborted on an evaluation of one of the early exits *) - \/ (exists ext lx, all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m /\ sabort_exit ge sp ext rs m). - -Definition ssem_internal_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := - match ois with - | Some is => ssem_internal ge sp st rs0 m0 is - | None => sabort ge sp st rs0 m0 - end. - -Definition ssem_internal_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := - match ost with - | Some st => ssem_internal_opt ge sp st rs0 m0 ois - | None => ois=None - end. - -(** * An internal state represents a parallel program ! - - We prove below that the semantics [ssem_internal_opt] is deterministic. - - *) - -Definition istate_eq ist1 ist2 := - ist1.(icontinue) = ist2.(icontinue) /\ - ist1.(ipc) = ist2.(ipc) /\ - (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\ - ist1.(imem) = ist2.(imem). - -Lemma all_fallthrough_noexit ge sp ext lx rs0 m0 rs m pc: - ssem_exit ge sp ext rs0 m0 rs m pc -> - In ext lx -> - all_fallthrough ge sp lx rs0 m0 -> - False. -Proof. - Local Hint Resolve is_tail_in: core. - intros SSEM INE ALLF. - destruct SSEM as (SSEM & SSEM'). - unfold all_fallthrough in ALLF. rewrite ALLF in SSEM; eauto. - discriminate. -Qed. - -Lemma ssem_internal_exclude_incompatible_continue ge sp st rs m is1 is2: - is1.(icontinue) = true -> - is2.(icontinue) = false -> - ssem_internal ge sp st rs m is1 -> - ssem_internal ge sp st rs m is2 -> - False. -Proof. - Local Hint Resolve all_fallthrough_noexit: core. - unfold ssem_internal. - intros CONT1 CONT2. - rewrite CONT1, CONT2; simpl. - intuition eauto. - destruct H0 as (ext & lx & SSEME & ALLFU). - destruct ALLFU as (ALLFU & ALLFU'). - eapply all_fallthrough_noexit; eauto. -Qed. - -Lemma ssem_internal_determ_continue ge sp st rs m is1 is2: - ssem_internal ge sp st rs m is1 -> - ssem_internal ge sp st rs m is2 -> - is1.(icontinue) = is2.(icontinue). -Proof. - Local Hint Resolve ssem_internal_exclude_incompatible_continue: core. - destruct (Bool.bool_dec is1.(icontinue) is2.(icontinue)) as [|H]; auto. - intros H1 H2. assert (absurd: False); intuition. - destruct (icontinue is1) eqn: His1, (icontinue is2) eqn: His2; eauto. -Qed. - -Lemma ssem_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: - ssem_local ge sp st rs0 m0 rs1 m1 -> - ssem_local ge sp st rs0 m0 rs2 m2 -> - (forall r, rs1#r = rs2#r) /\ m1 = m2. -Proof. - unfold ssem_local. intuition try congruence. - generalize (H5 r); rewrite H4; congruence. -Qed. - -(* TODO: lemma to move in Coqlib *) -Lemma is_tail_bounded_total {A} (l1 l2 l3: list A): is_tail l1 l3 -> is_tail l2 l3 - -> is_tail l1 l2 \/ is_tail l2 l1. -Proof. - Local Hint Resolve is_tail_cons: core. - induction 1 as [|i l1 l3 T1 IND]; simpl; auto. - intros T2; inversion T2; subst; auto. -Qed. - -Lemma exit_cond_determ ge sp rs0 m0 l1 l2: - is_tail l1 l2 -> forall ext1 lx1 ext2 lx2, - l1=(ext1 :: lx1) -> - l2=(ext2 :: lx2) -> - all_fallthrough ge sp lx1 rs0 m0 -> - seval_condition ge sp (si_cond ext1) (si_scondargs ext1) (si_smem (si_elocal ext1)) rs0 m0 = Some true -> - all_fallthrough ge sp lx2 rs0 m0 -> - ext1=ext2. -Proof. - destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst; - inversion EQ2; subst; auto. - intros D1 EVAL NYE. - Local Hint Resolve is_tail_in: core. - unfold all_fallthrough in NYE. - rewrite NYE in EVAL; eauto. - try congruence. -Qed. - -Lemma ssem_exit_determ ge sp ext rs0 m0 rs1 m1 pc1 rs2 m2 pc2: - ssem_exit ge sp ext rs0 m0 rs1 m1 pc1 -> - ssem_exit ge sp ext rs0 m0 rs2 m2 pc2 -> - pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. -Proof. - Local Hint Resolve exit_cond_determ eq_sym: core. - intros SSEM1 SSEM2. destruct SSEM1 as (SEVAL1 & SLOC1 & PCEQ1). destruct SSEM2 as (SEVAL2 & SLOC2 & PCEQ2). subst. - destruct (ssem_local_determ ge sp (si_elocal ext) rs0 m0 rs1 m1 rs2 m2); auto. -Qed. - -Remark is_tail_inv_left {A: Type} (a a': A) l l': - is_tail (a::l) (a'::l') -> - (a = a' /\ l = l') \/ (In a l' /\ is_tail l (a'::l')). -Proof. - intros. inv H. - - left. eauto. - - right. econstructor. - + eapply is_tail_in; eauto. - + eapply is_tail_cons_left; eauto. -Qed. - -Lemma ssem_internal_determ ge sp st rs m is1 is2: - ssem_internal ge sp st rs m is1 -> - ssem_internal ge sp st rs m is2 -> - istate_eq is1 is2. -Proof. - unfold istate_eq. - intros SEM1 SEM2. - exploit (ssem_internal_determ_continue ge sp st rs m is1 is2); eauto. - intros CONTEQ. unfold ssem_internal in * |-. rewrite CONTEQ in * |- *. - destruct (icontinue is2). - - destruct (ssem_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); - intuition (try congruence). - - destruct SEM1 as (ext1 & lx1 & SSEME1 & ALLFU1). destruct SEM2 as (ext2 & lx2 & SSEME2 & ALLFU2). - destruct ALLFU1 as (ALLFU1 & ALLFU1'). destruct ALLFU2 as (ALLFU2 & ALLFU2'). - destruct SSEME1 as (SSEME1 & SSEME1' & SSEME1''). destruct SSEME2 as (SSEME2 & SSEME2' & SSEME2''). - assert (X:ext1=ext2). - { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2 :: lx2) (si_exits st)) as [TAIL|TAIL]; eauto. } - subst. destruct (ssem_local_determ ge sp (si_elocal ext2) rs m (irs is1) (imem is1) (irs is2) (imem is2)); auto. - intuition. congruence. -Qed. - -Lemma ssem_local_exclude_sabort_local ge sp loc rs m rs' m': - ssem_local ge sp loc rs m rs' m' -> - sabort_local ge sp loc rs m -> - False. -Proof. - intros SIML ABORT. inv SIML. destruct H0 as (H0 & H0'). - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. -Qed. - -Lemma ssem_local_exclude_sabort ge sp st rs m rs' m': - ssem_local ge sp (si_local st) rs m rs' m' -> - all_fallthrough ge sp (si_exits st) rs m -> - sabort ge sp st rs m -> - False. -Proof. - intros SIML ALLF ABORT. - inv ABORT. - - intuition; eapply ssem_local_exclude_sabort_local; eauto. - - destruct H as (ext & lx & ALLFU & SABORT). - destruct ALLFU as (TAIL & _). eapply is_tail_in in TAIL. - eapply ALLF in TAIL. - destruct SABORT as [CONDFAIL | (CONDTRUE & ABORTL)]; congruence. -Qed. - -Lemma ssem_exit_fallthrough_upto_exit ge sp ext ext' lx lx' exits rs m rs' m' pc': - ssem_exit ge sp ext rs m rs' m' pc' -> - all_fallthrough_upto_exit ge sp ext lx exits rs m -> - all_fallthrough_upto_exit ge sp ext' lx' exits rs m -> - is_tail (ext'::lx') (ext::lx). -Proof. - intros SSEME ALLFU ALLFU'. - destruct ALLFU as (ISTAIL & ALLFU). destruct ALLFU' as (ISTAIL' & ALLFU'). - destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto. - inv H. - - econstructor; eauto. - - eapply is_tail_in in H2. eapply ALLFU' in H2. - destruct SSEME as (SEVAL & _). congruence. -Qed. - -Lemma ssem_exit_exclude_sabort_exit ge sp ext rs m rs' m' pc': - ssem_exit ge sp ext rs m rs' m' pc' -> - sabort_exit ge sp ext rs m -> - False. -Proof. - intros A B. destruct A as (A & A' & A''). inv B. - - congruence. - - destruct H as (_ & H). eapply ssem_local_exclude_sabort_local; eauto. -Qed. - -Lemma ssem_exit_exclude_sabort ge sp ext st lx rs m rs' m' pc': - ssem_exit ge sp ext rs m rs' m' pc' -> - all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m -> - sabort ge sp st rs m -> - False. -Proof. - intros SSEM ALLFU ABORT. - inv ABORT. - - destruct H as (ALLF & _). destruct ALLFU as (TAIL & _). - eapply is_tail_in in TAIL. - destruct SSEM as (SEVAL & _ & _). - eapply ALLF in TAIL. congruence. - - destruct H as (ext' & lx' & ALLFU' & ABORT). - exploit ssem_exit_fallthrough_upto_exit; eauto. intros ITAIL. - destruct ALLFU as (ALLFU1 & ALLFU2). destruct ALLFU' as (ALLFU1' & ALLFU2'). - exploit (is_tail_inv_left ext' ext lx' lx); eauto. intro. inv H. - + inv H0. eapply ssem_exit_exclude_sabort_exit; eauto. - + destruct H0 as (INE & TAIL). eapply ALLFU2 in INE. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence. -Qed. - -Lemma ssem_internal_exclude_sabort ge sp st rs m is: - sabort ge sp st rs m -> - ssem_internal ge sp st rs m is -> False. -Proof. - intros ABORT SEM. - unfold ssem_internal in SEM. destruct icontinue. - - destruct SEM as (SEM1 & SEM2 & SEM3). - eapply ssem_local_exclude_sabort; eauto. - - destruct SEM as (ext & lx & SEM1 & SEM2). eapply ssem_exit_exclude_sabort; eauto. -Qed. - -Definition istate_eq_opt ist1 oist := - exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. - -Lemma ssem_internal_opt_determ ge sp st rs m ois is: - ssem_internal_opt ge sp st rs m ois -> - ssem_internal ge sp st rs m is -> - istate_eq_opt is ois. -Proof. - destruct ois as [is1|]; simpl; eauto. - - intros; eexists; intuition; eapply ssem_internal_determ; eauto. - - intros; exploit ssem_internal_exclude_sabort; eauto. destruct 1. -Qed. - -(** * Symbolic execution of one internal step *) - -Definition slocal_set_sreg (st:sistate_local) (r:reg) (sv:sval) := - {| si_pre:=(fun ge sp rs m => eval_sval ge sp (st.(si_sreg) r) rs m <> None /\ (st.(si_pre) ge sp rs m)); - si_sreg:=fun y => if Pos.eq_dec r y then sv else st.(si_sreg) y; - si_smem:= st.(si_smem)|}. - -Definition slocal_set_smem (st:sistate_local) (sm:smem) := - {| si_pre:=(fun ge sp rs m => eval_smem ge sp st.(si_smem) rs m <> None /\ (st.(si_pre) ge sp rs m)); - si_sreg:= st.(si_sreg); - si_smem:= sm |}. - -Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): sistate := - {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. - -Definition slocal_store st chunk addr args src : sistate_local := - let args := list_sval_inj (List.map (si_sreg st) args) in - let src := si_sreg st src in - let sm := Sstore (si_smem st) chunk addr args src - in slocal_set_smem st sm. - -Definition siexec_inst (i: instruction) (st: sistate): option sistate := - match i with - | Inop pc' => - Some (sist_set_local st pc' st.(si_local)) - | Iop op args dst pc' => - let prev := st.(si_local) in - let vargs := list_sval_inj (List.map prev.(si_sreg) args) in - let next := slocal_set_sreg prev dst (Sop op vargs prev.(si_smem)) in - Some (sist_set_local st pc' next) - | Iload trap chunk addr args dst pc' => - let prev := st.(si_local) in - let vargs := list_sval_inj (List.map prev.(si_sreg) args) in - let next := slocal_set_sreg prev dst (Sload prev.(si_smem) trap chunk addr vargs) in - Some (sist_set_local st pc' next) - | Istore chunk addr args src pc' => - let next := slocal_store st.(si_local) chunk addr args src in - Some (sist_set_local st pc' next) - | Icond cond args ifso ifnot _ => - let prev := st.(si_local) in - let vargs := list_sval_inj (List.map prev.(si_sreg) args) in - let ex := {| si_cond:=cond; si_scondargs:=vargs; si_elocal := prev; si_ifso := ifso |} in - Some {| si_pc := ifnot; si_exits := ex::st.(si_exits); si_local := prev |} - | _ => None - end. - -Lemma slocal_set_sreg_preserves_sabort_local ge sp st rs0 m0 r sv: - sabort_local ge sp st rs0 m0 -> - sabort_local ge sp (slocal_set_sreg st r sv) rs0 m0. -Proof. - unfold sabort_local. simpl; intuition. - destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST. - - subst; rewrite H; intuition. - - right. right. exists r1. rewrite HTEST. auto. -Qed. - -Lemma slocal_set_smem_preserves_sabort_local ge sp st rs0 m0 m: - sabort_local ge sp st rs0 m0 -> - sabort_local ge sp (slocal_set_smem st m) rs0 m0. -Proof. - unfold sabort_local. simpl; intuition. -Qed. - -Lemma all_fallthrough_upto_exit_cons ge sp ext lx ext' exits rs m: - all_fallthrough_upto_exit ge sp ext lx exits rs m -> - all_fallthrough_upto_exit ge sp ext lx (ext'::exits) rs m. -Proof. - intros. inv H. econstructor; eauto. -Qed. - -Lemma all_fallthrough_cons ge sp exits rs m ext: - all_fallthrough ge sp exits rs m -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false -> - all_fallthrough ge sp (ext::exits) rs m. -Proof. - intros. unfold all_fallthrough in *. intros. - inv H1; eauto. -Qed. - -Lemma siexec_inst_preserves_sabort i ge sp rs m st st': - siexec_inst i st = Some st' -> - sabort ge sp st rs m -> sabort ge sp st' rs m. -Proof. - intros SISTEP ABORT. - destruct i; simpl in SISTEP; try discriminate; inv SISTEP; unfold sabort; simpl. - (* NOP *) - * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - left. constructor; eauto. - - right. exists ext0, lx0. constructor; eauto. - (* OP *) - * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto. - - right. exists ext0, lx0. constructor; eauto. - (* LOAD *) - * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto. - - right. exists ext0, lx0. constructor; eauto. - (* STORE *) - * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - left. constructor; eauto. eapply slocal_set_smem_preserves_sabort_local; eauto. - - right. exists ext0, lx0. constructor; eauto. - (* COND *) - * remember ({| si_cond := _; si_scondargs := _; si_elocal := _; si_ifso := _ |}) as ext. - destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - destruct (seval_condition ge sp (si_cond ext) (si_scondargs ext) - (si_smem (si_elocal ext)) rs m) eqn:SEVAL; [destruct b|]. - (* case true *) - + right. exists ext, (si_exits st). - constructor. - ++ constructor. econstructor; eauto. eauto. - ++ unfold sabort_exit. right. constructor; eauto. - subst. simpl. eauto. - (* case false *) - + left. constructor; eauto. eapply all_fallthrough_cons; eauto. - (* case None *) - + right. exists ext, (si_exits st). constructor. - ++ constructor. econstructor; eauto. eauto. - ++ unfold sabort_exit. left. eauto. - - right. exists ext0, lx0. constructor; eauto. eapply all_fallthrough_upto_exit_cons; eauto. -Qed. - -Lemma siexec_inst_WF i st: - siexec_inst i st = None -> default_succ i = None. -Proof. - destruct i; simpl; unfold sist_set_local; simpl; congruence. -Qed. - -Lemma siexec_inst_default_succ i st st': - siexec_inst i st = Some st' -> default_succ i = Some (st'.(si_pc)). -Proof. - destruct i; simpl; unfold sist_set_local; simpl; try congruence; - intro H; inversion_clear H; simpl; auto. -Qed. - - -Lemma eval_list_sval_inj_not_none ge sp st rs0 m0: forall l, - (forall r, List.In r l -> eval_sval ge sp (si_sreg st r) rs0 m0 = None -> False) -> - eval_list_sval ge sp (list_sval_inj (map (si_sreg st) l)) rs0 m0 = None -> False. -Proof. - induction l. - - intuition discriminate. - - intros ALLR. simpl. - inversion_SOME v. - + intro SVAL. inversion_SOME lv; [discriminate|]. - assert (forall r : reg, In r l -> eval_sval ge sp (si_sreg st r) rs0 m0 = None -> False). - { intros r INR. eapply ALLR. right. assumption. } - intro SVALLIST. intro. eapply IHl; eauto. - + intros. exploit (ALLR a); simpl; eauto. -Qed. - -Lemma siexec_inst_correct ge sp i st rs0 m0 rs m: - ssem_local ge sp st.(si_local) rs0 m0 rs m -> - all_fallthrough ge sp st.(si_exits) rs0 m0 -> - ssem_internal_opt2 ge sp (siexec_inst i st) rs0 m0 (istep ge i sp rs m). -Proof. - intros (PRE & MEM & REG) NYE. - destruct i; simpl; auto. - + (* Nop *) - constructor; [|constructor]; simpl; auto. - constructor; auto. - + (* Op *) - inversion_SOME v; intros OP; simpl. - - constructor; [|constructor]; simpl; auto. - constructor; simpl; auto. - * constructor; auto. congruence. - * constructor; auto. - intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst. rewrite Regmap.gss; simpl; auto. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - - left. constructor; simpl; auto. - unfold sabort_local. right. right. - simpl. exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - + (* LOAD *) - inversion_SOME a0; intro ADD. - { inversion_SOME v; intros LOAD; simpl. - - explore_destruct; unfold ssem_internal, ssem_local; simpl; intuition. - * unfold ssem_internal. simpl. constructor; [|constructor]; auto. - constructor; constructor; simpl; auto. congruence. intro r0. - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - * unfold ssem_internal. simpl. constructor; [|constructor]; auto. - constructor; constructor; simpl; auto. congruence. intro r0. - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl. - inversion_SOME args; intros ARGS. - 2: { exploit eval_list_sval_inj_not_none; eauto; intuition congruence. } - exploit eval_list_sval_inj; eauto. intro ARGS'. erewrite ARGS in ARGS'. inv ARGS'. rewrite ADD. - inversion_SOME m2. intro SMEM. - assert (m = m2) by congruence. subst. rewrite LOAD. reflexivity. - - explore_destruct; unfold sabort, sabort_local; simpl. - * unfold sabort. simpl. left. constructor; auto. - right. right. exists r. simpl. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite eval_list_sval_inj; simpl; auto. - rewrite ADD; simpl; auto. try_simplify_eval_svalsomeHyps. - * unfold ssem_internal. simpl. constructor; [|constructor]; auto. - constructor; constructor; simpl; auto. congruence. intro r0. - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - } { rewrite ADD. destruct t. - - simpl. left; eauto. simpl. econstructor; eauto. - right. right. simpl. exists r. destruct (Pos.eq_dec r r); [|contradiction]. - simpl. inversion_SOME args. intro SLS. - eapply eval_list_sval_inj in REG. rewrite REG in SLS. inv SLS. - rewrite ADD. reflexivity. - - simpl. constructor; [|constructor]; simpl; auto. - constructor; simpl; constructor; auto; [congruence|]. - intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst. simpl. rewrite Regmap.gss. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - } - + (* STORE *) - inversion_SOME a0; intros ADD. - { inversion_SOME m'; intros STORE; simpl. - - unfold ssem_internal, ssem_local; simpl; intuition. - * congruence. - * erewrite eval_list_sval_inj; simpl; auto. - erewrite REG. - try_simplify_someHyps. - - unfold sabort, sabort_local; simpl. - left. constructor; auto. right. left. - erewrite eval_list_sval_inj; simpl; auto. - erewrite REG. - try_simplify_someHyps. } - { unfold sabort, sabort_local; simpl. - left. constructor; auto. right. left.eval_sval - erewrite eval_list_sval_inj; simpl; auto. - erewrite ADD; simpl; auto. } - + (* COND *) - Local Hint Resolve is_tail_refl: core. - Local Hint Unfold ssem_local: core. - inversion_SOME b; intros COND. - { destruct b; simpl; unfold ssem_internal, ssem_local; simpl. - - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st). - constructor; constructor; subst; simpl; auto. - unfold seval_condition. subst; simpl. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - - intuition. unfold all_fallthrough in * |- *. simpl. - intuition. subst. simpl. - unfold seval_condition. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. } - { unfold sabort. simpl. right. - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st). - constructor; [constructor; subst; simpl; auto|]. - left. subst; simpl; auto. - unfold seval_condition. - erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. } -Qed. - - -Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m: - ssem_local ge sp (st.(si_local)) rs0 m0 rs m -> - siexec_inst i st = None -> - istep ge i sp rs m = None. -Proof. - intros (PRE & MEM & REG). - destruct i; simpl; unfold sist_set_local, ssem_internal, ssem_local; simpl; try_simplify_someHyps. -Qed. - -(** * Symbolic execution of the internal steps of a path *) -Fixpoint siexec_path (path:nat) (f: function) (st: sistate): option sistate := - match path with - | O => Some st - | S p => - SOME i <- (fn_code f)!(st.(si_pc)) IN - SOME st1 <- siexec_inst i st IN - siexec_path p f st1 - end. - -Lemma siexec_inst_add_exits i st st': - siexec_inst i st = Some st' -> - ( si_exits st' = si_exits st \/ exists ext, si_exits st' = ext :: si_exits st ). -Proof. - destruct i; simpl; intro SISTEP; inversion_clear SISTEP; unfold siexec_inst; simpl; (discriminate || eauto). -Qed. - -Lemma siexec_inst_preserves_allfu ge sp ext lx rs0 m0 st st' i: - all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs0 m0 -> - siexec_inst i st = Some st' -> - all_fallthrough_upto_exit ge sp ext lx (si_exits st') rs0 m0. -Proof. - intros ALLFU SISTEP. destruct ALLFU as (ISTAIL & ALLF). - constructor; eauto. - destruct i; simpl in SISTEP; inversion_clear SISTEP; simpl; (discriminate || eauto). -Qed. - -Lemma siexec_path_correct_false ge sp f rs0 m0 st' is: - forall path, - is.(icontinue)=false -> - forall st, ssem_internal ge sp st rs0 m0 is -> - siexec_path path f st = Some st' -> - ssem_internal ge sp st' rs0 m0 is. -Proof. - induction path; simpl. - - intros. congruence. - - intros ICF st SSEM STEQ'. - destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate]. - destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate]. - eapply IHpath. 3: eapply STEQ'. eauto. - unfold ssem_internal in SSEM. rewrite ICF in SSEM. - destruct SSEM as (ext & lx & SEXIT & ALLFU). - unfold ssem_internal. rewrite ICF. exists ext, lx. - constructor; auto. eapply siexec_inst_preserves_allfu; eauto. -Qed. - -Lemma siexec_path_preserves_sabort ge sp path f rs0 m0 st': forall st, - siexec_path path f st = Some st' -> - sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. -Proof. - Local Hint Resolve siexec_inst_preserves_sabort: core. - induction path; simpl. - + unfold sist_set_local; try_simplify_someHyps. - + intros st; inversion_SOME i. - inversion_SOME st1; eauto. -Qed. - -Lemma siexec_path_WF path f: forall st, - siexec_path path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None. -Proof. - induction path; simpl. - + unfold sist_set_local. intuition congruence. - + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try tauto. - destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl. - - intros; erewrite siexec_inst_default_succ; eauto. - - intros; erewrite siexec_inst_WF; eauto. -Qed. - -Lemma siexec_path_default_succ path f st': forall st, - siexec_path path f st = Some st' -> nth_default_succ (fn_code f) path st.(si_pc) = Some st'.(si_pc). -Proof. - induction path; simpl. - + unfold sist_set_local. intros st H. inversion_clear H; simpl; try congruence. - + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try congruence. - destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl; try congruence. - intros; erewrite siexec_inst_default_succ; eauto. -Qed. - -Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is, - is.(icontinue)=true -> - ssem_internal ge sp st rs0 m0 is -> - nth_default_succ (fn_code f) path st.(si_pc) <> None -> - ssem_internal_opt2 ge sp (siexec_path path f st) rs0 m0 - (isteps ge path f sp is.(irs) is.(imem) is.(ipc)) - . -Proof. - Local Hint Resolve siexec_path_correct_false siexec_path_preserves_sabort siexec_path_WF: core. - induction path; simpl. - + intros st is CONT INV WF; - unfold ssem_internal, sist_set_local in * |- *; - try_simplify_someHyps. simpl. - destruct is; simpl in * |- *; subst; intuition auto. - + intros st is CONT; unfold ssem_internal at 1; rewrite CONT. - intros (LOCAL & PC & NYE) WF. - rewrite <- PC. - inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. - exploit siexec_inst_correct; eauto. - inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - - inversion_SOME is1; intros His1;rewrite His1; simpl. - * destruct (icontinue is1) eqn:CONT1. - (* icontinue is0 = true *) - intros; eapply IHpath; eauto. - destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps. - (* icontinue is0 = false -> EARLY EXIT *) - destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto. - destruct WF. erewrite siexec_inst_default_succ; eauto. - (* try_simplify_someHyps; eauto. *) - * destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto. - - intros His1;rewrite His1; simpl; auto. -Qed. - -(** REM: in the following two unused lemmas *) - -Lemma siexec_path_right_assoc_decompose f path: forall st st', - siexec_path (S path) f st = Some st' -> - exists st0, siexec_path path f st = Some st0 /\ siexec_path 1%nat f st0 = Some st'. -Proof. - induction path; simpl; eauto. - intros st st'. - inversion_SOME i1. - inversion_SOME st1. - try_simplify_someHyps; eauto. -Qed. - -Lemma siexec_path_right_assoc_compose f path: forall st st0 st', - siexec_path path f st = Some st0 -> - siexec_path 1%nat f st0 = Some st' -> - siexec_path (S path) f st = Some st'. -Proof. - induction path. - + intros st st0 st' H. simpl in H. - try_simplify_someHyps; auto. - + intros st st0 st'. - assert (X:exists x, x=(S path)); eauto. - destruct X as [x X]. - intros H1 H2. rewrite <- X. - generalize H1; clear H1. simpl. - inversion_SOME i1. intros Hi1; rewrite Hi1. - inversion_SOME st1. intros Hst1; rewrite Hst1. - subst; eauto. -Qed. - - -(** * Main function of the symbolic execution *) - -Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. - -Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}. - -Lemma init_ssem_internal ge sp pc rs m: ssem_internal ge sp (init_sistate pc) rs m (mk_istate true pc rs m). -Proof. - unfold ssem_internal, ssem_local, all_fallthrough; simpl. intuition. -Qed. - -Definition sexec (f: function) (pc:node): option sstate := - SOME path <- (fn_path f)!pc IN - SOME st <- siexec_path path.(psize) f (init_sistate pc) IN - SOME i <- (fn_code f)!(st.(si_pc)) IN - Some (match siexec_inst i st with - | Some st' => {| internal := st'; final := Snone |} - | None => {| internal := st; final := sexec_final i st.(si_local) |} - end). - -Lemma final_node_path_simpl f path pc: - (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(psize) pc <> None. -Proof. - intros; exploit final_node_path; eauto. - intros (i & NTH & DUM). - congruence. -Qed. - -Lemma symb_path_last_step i st st' ge pge stack (f:function) sp pc rs m t s: - (fn_code f) ! pc = Some i -> - pc = st.(si_pc) -> - siexec_inst i st = Some st' -> - path_last_step ge pge stack f sp pc rs m t s -> - exists mk_istate, - istep ge i sp rs m = Some mk_istate - /\ t = E0 - /\ s = (State stack f sp mk_istate.(ipc) mk_istate.(RTLpath.irs) mk_istate.(imem)). -Proof. - intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl. -Qed. - -(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) -(sexec is a correct over-approximation) -*) -Theorem sexec_correct f pc pge ge sp path stack rs m t s: - (fn_path f)!pc = Some path -> - path_step ge pge path.(psize) stack f sp rs m pc t s -> - exists st, sexec f pc = Some st /\ ssem pge ge sp st stack f rs m t s. -Proof. - Local Hint Resolve init_ssem_internal: core. - intros PATH STEP; unfold sexec; rewrite PATH; simpl. - lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. - { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. } - (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]). - destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST]; - (* intro Hst *) - (rewrite STEPS; unfold ssem_internal_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence); - (* intro SEM *) - (simpl; unfold ssem_internal; simpl; rewrite CONT; intro SEM); - (* intro Hi' *) - ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); - [ unfold nth_default_succ_inst in Hi; - exploit siexec_path_default_succ; eauto; simpl; - intros DEF; rewrite DEF in Hi; auto - | clear Hi; rewrite Hi' ]); - (* eexists *) - (eexists; constructor; eauto). - - (* early *) - eapply ssem_early; eauto. - unfold ssem_internal; simpl; rewrite CONT. - destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl; eauto. - destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx. - constructor; auto. eapply siexec_inst_preserves_allfu; eauto. - - destruct SEM as (SEM & PC & HNYE). - destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl. - + (* normal on Snone *) - rewrite <- PC in LAST. - exploit symb_path_last_step; eauto; simpl. - intros (mk_istate & ISTEP & Ht & Hs); subst. - exploit siexec_inst_correct; eauto. simpl. - erewrite Hst', ISTEP; simpl. - clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i. - intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT. - { (* icontinue mk_istate = true *) - eapply ssem_normal; simpl; eauto. - unfold ssem_internal in SEM. - rewrite CONT in SEM.eval_sval - destruct SEM as (SEM & PC & HNYE). - rewrite <- PC. - eapply exec_Snone. } - { eapply ssem_early; eauto. } - + (* normal non-Snone instruction *) - eapply ssem_normal; eauto. - * unfold ssem_internal; simpl; rewrite CONT; intuition. - * simpl. eapply sexec_final_correct; eauto. - rewrite PC; auto. -Qed. - -(* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *) -Inductive equiv_stackframe: stackframe -> stackframe -> Prop := - | equiv_stackframe_intro res f sp pc rs1 rs2 - (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): - equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). - -Inductive equiv_state: state -> state -> Prop := - | State_equiv stack f sp pc rs1 m rs2 - (EQUIV: forall r, rs1#r = rs2#r): - equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) - | Call_equiv stk stk' f args m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Callstate stk f args m) (Callstate stk' f args m) - | Return_equiv stk stk' v m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Returnstate stk v m) (Returnstate stk' v m). - -Lemma equiv_stackframe_refl stf: equiv_stackframeeval_sval stf stf. -Proof. - destruct stf. constructor; auto. -Qed. - -Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. -Proof. - Local Hint Resolve equiv_stackframe_refl: core. - induction stk; simpl; constructor; auto. -Qed.eval_sval - -Lemma equiv_state_refl s: equiv_state s s. -Proof. - Local Hint Resolve equiv_stack_refl: core. - induction s; simpl; constructor; auto. -Qed. - -(* -Lemma equiv_stackframe_trans stf1 stf2 stf3: - equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3. -Proof. - destruct 1; intros EQ; inv EQ; try econstructor; eauto. - intros; eapply eq_trans; eauto. -Qed. - -Lemma equiv_stack_trans stk1 stk2: - list_forall2 equiv_stackframe stk1 stk2 -> - forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> - list_forall2 equiv_stackframe stk1 stk3. -Proof. - Local Hint Resolve equiv_stackframe_trans. - induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. -Qed. - -Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3. -Proof. - Local Hint Resolve equiv_stack_trans. - destruct 1; intros EQ; inv EQ; econstructor; eauto. - intros; eapply eq_trans; eauto. -Qed. -*) - -Lemma regmap_setres_eq (rs rs': regset) res vres: - (forall r, rs # r = rs' # r) -> - forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r. -Proof. - intros RSEQ r. destruct res; simpl; try congruence. - destruct (peq x r). - - subst. repeat (rewrite Regmap.gss). reflexivity. - - repeat (rewrite Regmap.gso); auto. -Qed. - -Lemma sem_sfval_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: - sem_sfval pge ge sp st stack f rs0 m0 sv rs1 m t s -> - (forall r, rs1#r = rs2#r) -> - exists s', equiv_state s s' /\ sem_sfval pge ge sp st stack f rs0 m0 sv rs2 m t s'. -Proof. eval_sval - Local Hint Resolve equiv_stack_refl: core. - destruct 1. - - (* Snone *) intros; eexists; econstructor. - + eapply State_equiv; eauto. - + eapply exec_Snone. - - (* Scall *) - intros; eexists; econstructor. - 2: { eapply exec_Scall; eauto. } - apply Call_equiv; auto. - repeat (constructor; auto). - - (* Stailcall *) - intros; eexists; econstructor; [| eapply exec_Stailcall; eauto]. - apply Call_equiv; auto. - - (* Sbuiltin *) - intros; eexists; econstructor; [| eapply exec_Sbuiltin; eauto]. - constructor. eapply regmap_setres_eq; eauto. - - (* Sjumptable *) - intros; eexists; econstructor; [| eapply exec_Sjumptable; eauto]. - constructor. assumption. - - (* Sreturn *) - intros; eexists; econstructor; [| eapply exec_Sreturn; eauto]. - eapply equiv_state_refl; eauto. -Qed. - -Lemma siexec_inst_early_exit_absurd i st st' ge sp rs m rs' m' pc': - siexec_inst i st = Some st' -> - (exists ext lx, ssem_exit ge sp ext rs m rs' m' pc' /\ - all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m) -> - all_fallthrough ge sp (si_exits st') rs m -> - False. -Proof. - intros SIEXEC (ext & lx & SSEME & ALLFU) ALLF. destruct ALLFU as (TAIL & _). - exploit siexec_inst_add_exits; eauto. destruct 1 as [SIEQ | (ext0 & SIEQ)]. - - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. eassumption. - - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. - constructor. eassumption. -Qed. -eval_sval -Lemma is_tail_false {A: Type}: forall (l: list A) a, is_tail (a::l) nil -> False. -Proof. - intros. eapply is_tail_incl in H. unfold incl in H. pose (H a). - assert (In a (a::l)) by (constructor; auto). assert (In a nil) by auto. apply in_nil in H1. - contradiction. -Qed. - -Lemma cons_eq_false {A: Type}: forall (l: list A) a, - a :: l = l -> False. -Proof. - induction l; intros. - - discriminate. - - inv H. apply IHl in H2. contradiction. -Qed. - -Lemma app_cons_nil_eq {A: Type}: forall l' l (a:A), - (l' ++ a :: nil) ++ l = l' ++ a::l. -Proof. - induction l'; intros. - - simpl. reflexivity. - - simpl. rewrite IHl'. reflexivity. -Qed. - -Lemma app_eq_false {A: Type}: forall l (l': list A) a, - l' ++ a :: l = l -> False. -Proof. - induction l; intros. - - apply app_eq_nil in H. destruct H as (_ & H). apply cons_eq_false in H. contradiction. - - destruct l' as [|a' l']. - + simpl in H. apply cons_eq_false in H. contradiction. - + rewrite <- app_comm_cons in H. inv H. - apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption. -Qed. - -Lemma is_tail_false_gen {A: Type}: forall (l: list A) l' a, is_tail (l'++(a::l)) l -> False. -Proof. - induction l. - - intros. destruct l' as [|a' l']. - + simpl in H. apply is_tail_false in H. contradiction. - + rewrite <- app_comm_cons in H. apply is_tail_false in H. contradiction. - - intros. inv H. - + apply app_eq_false in H2. contradiction. - + apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption. -Qed. - -Lemma is_tail_eq {A: Type}: forall (l l': list A), - is_tail l' l -> - is_tail l l' -> - l = l'. -Proof. - destruct l as [|a l]; intros l' ITAIL ITAIL'. - - destruct l' as [|i' l']; auto. apply is_tail_false in ITAIL. contradiction. - - inv ITAIL; auto. - destruct l' as [|i' l']. { apply is_tail_false in ITAIL'. contradiction. } - exploit is_tail_trans. eapply ITAIL'. eauto. intro ABSURD. - apply (is_tail_false_gen l nil a) in ABSURD. contradiction. -Qed. - -(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution - (sexec is exact). -*) -Theorem sexec_exact f pc pge ge sp path stack st rs m t s1: - (fn_path f)!pc = Some path -> - sexec f pc = Some st -> - ssem pge ge sp st stack f rs m t s1 -> - exists s2, path_step ge pge path.(psize) stack f sp rs m pc t s2 /\ - equiv_state s1 s2. -Proof. - Local Hint Resolve init_ssem_internal: core. - unfold sexec; intros PATH SSTEP SEM; rewrite PATH in SSTEP. - lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. - { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. } - (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]). - unfold nth_default_succ_inst in Hi. - destruct (siexec_path path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl. - 2:{ (* absurd case *) - exploit siexec_path_WF; eauto. - simpl; intros NDS; rewrite NDS in Hi; congruence. } - exploit siexec_path_default_succ;eval_sval eauto; simpl. - intros NDS; rewrite NDS in Hi. - rewrite Hi in SSTEP. - intros ISTEPS. try_simplify_someHyps. - destruct (siexec_inst i st0) as [st'|] eqn:Hst'; simpl. - + (* exit on Snone instruction *) - assert (SEM': t = E0 /\ exists is, ssem_internal ge sp st' rs m is - /\ s1 = (State stack f sp (if (icontinue is) then (si_pc st') else (ipc is)) (irs is) (imem is))). - { destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - - repeat (econstructor; eauto). - rewrite CONT; eauto. - - inversion SEM2. repeat (econstructor; eauto). - rewrite CONT; eauto. } - clear SEM; subst. destruct SEM' as [X (is & SEM & X')]; subst. - intros. - destruct (isteps ge (psize path) f sp rs m pc) as [is0|] eqn:RISTEPS; simpl in *. - * unfold ssem_internal in ISTEPS. destruct (icontinue is0) eqn: ICONT0. - ** (* icontinue is0=true: path_step by normal_exit *) - destruct ISTEPS as (SEMis0&H1&H2). - rewrite H1 in * |-. - exploit siexec_inst_correct; eauto. - rewrite Hst'; simpl. - intros; exploit ssem_internal_opt_determ; eauto. - destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - eexists. econstructor 1. - *** eapply exec_normal_exit; eauto. - eapply exec_istate; eauto. - *** rewrite EQ1. - enough ((ipc st) = (if icontinue st then si_pc st' else ipc is)) as ->. - { rewrite EQ2, EQ4. eapply State_equiv; auto. } - destruct (icontinue st) eqn:ICONT; auto. - exploit siexec_inst_default_succ; eauto. - erewrite istep_normal_exit; eauto. - try_simplify_someHyps. - ** (* The concrete execution has not reached "i" => early exit *) - unfold ssem_internal in SEM. - destruct (icontinue is) eqn:ICONT. - { destruct SEM as (SEML & SIPC & ALLF). - exploit siexec_inst_early_exit_absurd; eauto. contradiction. } - - eexists. econstructor 1. - *** eapply exec_early_exit; eauto. - *** destruct ISTEPS as (ext & lx & SSEME & ALLFU). destruct SEM as (ext' & lx' & SSEME' & ALLFU'). - eapply siexec_inst_preserves_allfu in ALLFU; eauto. - exploit ssem_exit_fallthrough_upto_exit; eauto. - exploit ssem_exit_fallthrough_upto_exit. eapply SSEME. eapply ALLFU. eapply ALLFU'. - intros ITAIL ITAIL'. apply is_tail_eq in ITAIL; auto. clear ITAIL'. - inv ITAIL. exploit ssem_exit_determ. eapply SSEME. eapply SSEME'. intros (IPCEQ & IRSEQ & IMEMEQ). - rewrite <- IPCEQ. rewrite <- IMEMEQ. constructor. congruence. - * (* The concrete execution has not reached "i" => abort case *) - eapply siexec_inst_preserves_sabort in ISTEPS; eauto. - exploit ssem_internal_exclude_sabort; eauto. contradiction. - + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - - (* early exit *) - intros. - exploit ssem_internal_opt_determ; eauto. - destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - eexists. econstructor 1. - * eapply exec_early_exit; eauto. - * rewrite EQ2, EQ4; eapply State_equiv. auto. - - (* normal exit non-Snone instruction *) - intros. - exploit ssem_internal_opt_determ; eauto. - destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - unfold ssem_internal in SEM1. - rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). - exploit sem_sfval_equiv; eauto. - clear SEM2; destruct 1 as (s' & Ms' & SEM2). - rewrite ! EQ4 in * |-; clear EQ4. - rewrite ! EQ2 in * |-; clear EQ2. - exists s'; intuition. - eapply exec_normal_exit; eauto. - eapply sexec_final_complete; eauto. - * congruence. - * unfold ssem_local in * |- *. - destruct SEM1 as (A & B & C). constructor; [|constructor]; eauto. - intro r. congruence. - * congruence. -Qed. - -Require Import RTLpathLivegen RTLpathLivegenproof. - -(** * DEFINITION OF SIMULATION BETWEEN (ABSTRACT) SYMBOLIC EXECUTIONS -*) - -Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := - is1.(icontinue) = is2.(icontinue) - /\ eqlive_reg alive is1.(irs) is2.(irs) - /\ is1.(imem) = is2.(imem). - -Definition istate_simu f (srce: PTree.t node) outframe is1 is2: Prop := - if is1.(icontinue) then - istate_simulive (fun r => Regset.In r outframe) srce is1 is2 - else - exists path, f.(fn_path)!(is1.(ipc)) = Some path - /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 - /\ srce!(is2.(ipc)) = Some is1.(ipc). - -Record simu_proof_context {f1: RTLpath.function} := { - liveness_hyps: liveness_ok_function f1; - the_ge1: RTL.genv; - the_ge2: RTL.genv; - genv_match: forall s, Genv.find_symbol the_ge1 s = Genv.find_symbol the_ge2 s; - the_sp: val; - the_rs0: regset; - the_m0: mem -}. -Arguments simu_proof_context: clear implicits. - -(* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) -Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) outframe (st1 st2: sistate) (ctx: simu_proof_context f): Prop := - forall is1, ssem_internal (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) is1 -> - exists is2, ssem_internal (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) is2 - /\ istate_simu f dm outframe is1 is2. - -Inductive svident_simu (f: RTLpath.function) (ctx: simu_proof_context f): (sval + ident) -> (sval + ident) -> Prop := - | Sleft_simu sv1 sv2: - (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) - -> svident_simu f ctx (inl sv1) (inl sv2) - | Sright_simu id1 id2: - id1 = id2 - -> svident_simu f ctx (inr id1) (inr id2) +(* symbolic state *) +Inductive sstate := + | Sfinal (sis: sistate) (sfv: sfval) + | Scond (cond: condition) (args: list_sval) (sm: smem) (ifso ifnot: sstate) + . + +Inductive sem_internal_sstate ctx rs m t s: sstate -> Prop := + | sem_Sfinal sis sfv + (SIS: sem_sistate ctx sis rs m) + (SFV: sem_sfval ctx sfv rs m t s) + : sem_internal_sstate ctx rs m t s (Sfinal sis sfv) + | sem_Scond b cond args sm ifso ifnot + (SEVAL: seval_condition ctx cond args sm = Some b) + (SELECT: sem_internal_sstate ctx rs m t s (if b then ifso else ifnot)) + : sem_internal_sstate ctx rs m t s (Scond cond args sm ifso ifnot) . - - -Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := - match lp with - | nil => Some nil - | p1::lp => SOME p2 <- pt!p1 IN - SOME lp2 <- (ptree_get_list pt lp) IN - Some (p2 :: lp2) - end. - -Lemma ptree_get_list_nth dm p2: forall lp2 lp1, - ptree_get_list dm lp2 = Some lp1 -> - forall n, list_nth_z lp2 n = Some p2 -> - exists p1, - list_nth_z lp1 n = Some p1 /\ dm ! p2 = Some p1. -Proof. - induction lp2. - - simpl. intros. inv H. simpl in *. discriminate. - - intros lp1 PGL n LNZ. simpl in PGL. explore. - inv LNZ. destruct (zeq n 0) eqn:ZEQ. - + subst. inv H0. exists n0. simpl; constructor; auto. - + exploit IHlp2; eauto. intros (p1 & LNZ & DMEQ). - eexists. simpl. rewrite ZEQ. - constructor; eauto. -Qed. - -Lemma ptree_get_list_nth_rev dm p1: forall lp2 lp1, - ptree_get_list dm lp2 = Some lp1 -> - forall n, list_nth_z lp1 n = Some p1 -> - exists p2, - list_nth_z lp2 n = Some p2 /\ dm ! p2 = Some p1. -Proof. - induction lp2. - - simpl. intros. inv H. simpl in *. discriminate. - - intros lp1 PGL n LNZ. simpl in PGL. explore. - inv LNZ. destruct (zeq n 0) eqn:ZEQ. - + subst. inv H0. exists a. simpl; constructor; auto. - + exploit IHlp2; eauto. intros (p2 & LNZ & DMEQ). - eexists. simpl. rewrite ZEQ. - constructor; eauto. congruence. -Qed. - -(* NOTE: we need to mix semantical simulation and syntactic definition on [sfval] in order to abstract the [match_states] *) -Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node) (ctx: simu_proof_context f): sfval -> sfval -> Prop := - | Snone_simu: - dm!opc2 = Some opc1 -> - sfval_simu dm f opc1 opc2 ctx Snone Snone - | Scall_simu sig svos1 svos2 lsv1 lsv2 res pc1 pc2: - dm!pc2 = Some pc1 -> - svident_simu f ctx svos1 svos2 -> - (eval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) - = (eval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) -> - sfval_simu dm f opc1 opc2 ctx (Scall sig svos1 lsv1 res pc1) (Scall sig svos2 lsv2 res pc2) - | Stailcall_simu sig svos1 svos2 lsv1 lsv2: - svident_simu f ctx svos1 svos2 -> - (eval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) - = (eval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) -> - sfval_simu dm f opc1 opc2 ctx (Stailcall sig svos1 lsv1) (Stailcall sig svos2 lsv2) - | Sbuiltin_simu ef lbs1 lbs2 br pc1 pc2: - dm!pc2 = Some pc1 -> - (eval_list_builtin_sval (the_ge1 ctx) (the_sp ctx) lbs1 (the_rs0 ctx) (the_m0 ctx)) - = (eval_list_builtin_sval (the_ge2 ctx) (the_sp ctx) lbs2 (the_rs0 ctx) (the_m0 ctx)) -> - sfval_simu dm f opc1 opc2 ctx (Sbuiltin ef lbs1 br pc1) (Sbuiltin ef lbs2 br pc2) - | Sjumptable_simu sv1 sv2 lpc1 lpc2: - ptree_get_list dm lpc2 = Some lpc1 -> - (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) - = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) -> - sfval_simu dm f opc1 opc2 ctx (Sjumptable sv1 lpc1) (Sjumptable sv2 lpc2) - | Sreturn_simu_none: sfval_simu dm f opc1 opc2 ctx (Sreturn None) (Sreturn None) - | Sreturn_simu_some sv1 sv2: - (eval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) - = (eval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) -> - sfval_simu dm f opc1 opc2 ctx (Sreturn (Some sv1)) (Sreturn (Some sv2)). - -Definition sstate_simu dm f outframe (s1 s2: sstate) (ctx: simu_proof_context f): Prop := - sistate_simu dm f outframe s1.(internal) s2.(internal) ctx - /\ forall is1, - ssem_internal (the_ge1 ctx) (the_sp ctx) s1 (the_rs0 ctx) (the_m0 ctx) is1 -> - is1.(icontinue) = true -> - sfval_simu dm f s1.(si_pc) s2.(si_pc) ctx s1.(final) s2.(final). - -Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, sexec f1 pc1 = Some st1 -> - exists path st2, (fn_path f1)!pc1 = Some path /\ sexec f2 pc2 = Some st2 - /\ forall ctx, sstate_simu dm f1 path.(pre_output_regs) st1 st2 ctx. - -*) \ No newline at end of file -- cgit From cb683f0bd30799ceb60bd04cb7a6561b8da5cc97 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 6 May 2021 19:51:16 +0200 Subject: starting RTL->BTL proof --- scheduling/RTLtoBTLproof.v | 50 ++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 46 insertions(+), 4 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index ed661280..aac7772c 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -80,13 +80,26 @@ Let tge := Genv.globalenv tprog. Local Open Scope nat_scope. -(* TODO: +(** * Match relation from a RTL state to a BTL state -Le (option iblock) en paramètre représente l'état de l'exécution côté BTL -depuis le début de l'exécution du block +The "option iblock" parameter represents the current BTL execution state. +Thus, each RTL single step is symbolized by a new BTL "option iblock" +starting at the equivalent PC. - faire un dessin ASCII art de la simulation avec match_states_intro +The simulation diagram for match_states_intro is as follows: +<< + + RTL state match_states_intro BTL state + [pc0,rs0,m0] ---------------------------- [pc0',rs0,m0] + | | + | | + RTL_RUN | *E0 | BTL_RUN + | | + | MIB | + [pc,rs,m] ---------------------------------- [ib] + +>> *) Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := @@ -209,6 +222,14 @@ Cette mesure décroit strictement quand on exécute un "petit-pas" de iblock. Par exemple, le nombre de noeuds (ou d'instructions "RTL") dans le iblock devrait convenir. La hauteur de l'arbre aussi. *) +(** Representing an intermediate BTL state + +We keep a measure of code that remains to be executed with the omeasure +type defined below. Intuitively, each RTL step corresponds to either + - a single BTL step if we are on the last instruction of the block + - no BTL step (as we use a "big step" semantics) but a change in + the measure which represents the new intermediate state of the BTL code + *) Parameter measure: iblock -> nat. Definition omeasure (oib: option iblock): nat := @@ -224,6 +245,27 @@ Theorem opt_simu s1 t s1' oib s2: (exists s2', step tge s2 t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure oib /\ t=E0 /\ match_states oib' s1' s2) . +Proof. + induction 2. + - admit. + - (* Callstate *) + inv H. + + (* Internal function *) + inv TRANSF. + rename H0 into MF. + eapply dupmap_entrypoint in MF as ENTRY. + eapply dupmap_correct in MF as DMC. unfold match_cfg in DMC. + apply DMC in ENTRY as DMC'. destruct DMC' as [ib [CENTRY MI]]. + eexists; left; eexists; split. + * eapply exec_function_internal. + erewrite preserv_fnstacksize; eauto. + * econstructor; eauto. + eapply code_right_assoc; eauto. + all: erewrite preserv_fnparams; eauto. + constructor. + + (* External function *) + admit. + - admit. Admitted. Local Hint Resolve plus_one star_refl: core. -- cgit From 9c6213a5265408426b94416cf3807e891e54569a Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 7 May 2021 07:26:22 +0200 Subject: fix SymbValPreserved section --- scheduling/BTL_SEtheory.v | 99 ++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 98 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 7025b90c..4d5314f2 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -14,7 +14,7 @@ Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestr Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) -Record iblock_exec_context := { +Record iblock_exec_context := Bctx { cge: BTL.genv; csp: val; cstk: list stackframe; @@ -431,6 +431,103 @@ Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' rs m) . +(** * Preservation properties *) + +Section SymbValPreserved. + +Variable ge ge': BTL.genv. + +Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. + +Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol ge id = Senv.find_symbol ge' id. +Proof. + destruct senv_preserved_BTL as (A & B & C). congruence. +Qed. + +Lemma senv_symbol_address_preserved id ofs: + Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. +Proof. + unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. + reflexivity. +Qed. + +Variable sp: val. +Variable stk stk': list stackframe. +Variable f f': function. +Variable rs0: regset. +Variable m0: mem. + +Lemma eval_sval_preserved sv: + eval_sval (Bctx ge sp stk f rs0 m0) sv = eval_sval (Bctx ge' sp stk' f' rs0 m0) sv. +Proof. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge sp stk f rs0 m0) lsv = eval_list_sval (Bctx ge' sp stk' f' rs0 m0) lsv) + (P1 := fun sm => eval_smem (Bctx ge sp stk f rs0 m0) sm = eval_smem (Bctx ge' sp stk' f' rs0 m0) sm); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma seval_builtin_arg_preserved m: forall bs varg, + seval_builtin_arg (Bctx ge sp stk f rs0 m0) m bs varg -> + seval_builtin_arg (Bctx ge' sp stk' f' rs0 m0) m bs varg. +Proof. + induction 1; simpl. + all: try (constructor; auto). + - rewrite <- eval_sval_preserved. assumption. + - rewrite <- senv_symbol_address_preserved. assumption. + - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. +Qed. + +Lemma seval_builtin_args_preserved m lbs vargs: + seval_builtin_args (Bctx ge sp stk f rs0 m0) m lbs vargs -> + seval_builtin_args (Bctx ge' sp stk' f' rs0 m0) m lbs vargs. +Proof. + induction 1; constructor; eauto. + eapply seval_builtin_arg_preserved; auto. +Qed. + +Lemma list_sval_eval_preserved lsv: + eval_list_sval (Bctx ge sp stk f rs0 m0) lsv = eval_list_sval (Bctx ge' sp stk' f' rs0 m0) lsv. +Proof. + induction lsv; simpl; auto. + rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma smem_eval_preserved sm: + eval_smem (Bctx ge sp stk f rs0 m0) sm = eval_smem (Bctx ge' sp stk' f' rs0 m0) sm. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. + rewrite eval_sval_preserved; auto. +Qed. + +Lemma seval_condition_preserved cond lsv sm: + seval_condition (Bctx ge sp stk f rs0 m0) cond lsv sm = seval_condition (Bctx ge' sp stk' f' rs0 m0) cond lsv sm. +Proof. + unfold seval_condition. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + +End SymbValPreserved. (* Syntax and Semantics of symbolic internal states *) -- cgit From 65a1029a0e2c1b1678e522f485b1e914b6e6d52a Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 7 May 2021 10:31:29 +0200 Subject: proof OK for Call and Return states --- scheduling/RTLtoBTLproof.v | 27 +++++++++++++++++++++------ 1 file changed, 21 insertions(+), 6 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index aac7772c..79b7ca1a 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -252,10 +252,11 @@ Proof. inv H. + (* Internal function *) inv TRANSF. - rename H0 into MF. - eapply dupmap_entrypoint in MF as ENTRY. - eapply dupmap_correct in MF as DMC. unfold match_cfg in DMC. - apply DMC in ENTRY as DMC'. destruct DMC' as [ib [CENTRY MI]]. + rename H0 into TRANSF. + eapply dupmap_entrypoint in TRANSF as ENTRY. + eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. + apply DMC in ENTRY as DMC'. + destruct DMC' as [ib [CENTRY MI]]; clear DMC. eexists; left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. @@ -264,8 +265,22 @@ Proof. all: erewrite preserv_fnparams; eauto. constructor. + (* External function *) - admit. - - admit. + inv TRANSF. + eexists; left; eexists; split. + * eapply exec_function_external. + eapply external_call_symbols_preserved. + eapply senv_preserved. eauto. + * econstructor; eauto. + - (* Returnstate *) + inv H. inv STACKS. inv H1. + eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. + remember DUPLIC as ODUPLIC; clear HeqODUPLIC. + apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. + eexists; left; eexists; split. + + eapply exec_return. + + econstructor; eauto. + eapply code_right_assoc; eauto. + constructor. Admitted. Local Hint Resolve plus_one star_refl: core. -- cgit From 72da07fabaea1139ce1a5bd26e907c7aea68c73f Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 7 May 2021 12:12:20 +0200 Subject: a draft for the Bnop case --- scheduling/RTLtoBTLproof.v | 36 +++++++++++++++++++++++++----------- 1 file changed, 25 insertions(+), 11 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 79b7ca1a..7d58de5a 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -91,29 +91,29 @@ The simulation diagram for match_states_intro is as follows: << RTL state match_states_intro BTL state - [pc0,rs0,m0] ---------------------------- [pc0',rs0,m0] + [pcR0,rs0,m0 ---------------------------- [pcB0,rs0,m0] | | | | RTL_RUN | *E0 | BTL_RUN | | | MIB | - [pc,rs,m] ---------------------------------- [ib] + [pcR1,rs,m] -------------------------------- [ib] >> *) Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) - ib dupmap st f sp pc rs m st' f' pc0' pc0 rs0 m0 isfst ib0 + ib dupmap st f sp rs m st' f' pcB0 pcR0 pcR1 rs0 m0 isfst ib0 (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') - (ATpc0: (fn_code f')!pc0' = Some ib0) - (DUPLIC: dupmap!pc0' = Some pc0) - (MIB: match_iblock dupmap (RTL.fn_code f) isfst pc ib None) + (ATpc0: (fn_code f')!pcB0 = Some ib0) + (DUPLIC: dupmap!pcB0 = Some pcR0) + (MIB: match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) (RIGHTA: is_right_assoc ib) - (RTL_RUN: star RTL.step ge (RTL.State st f sp pc0 rs0 m0) E0 (RTL.State st f sp pc rs m)) + (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) - : match_states (Some ib) (RTL.State st f sp pc rs m) (State st' f' sp pc0' rs0 m0) + : match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) | match_states_call st st' f f' args m (STACKS: list_forall2 match_stackframes st st') @@ -230,7 +230,12 @@ type defined below. Intuitively, each RTL step corresponds to either - no BTL step (as we use a "big step" semantics) but a change in the measure which represents the new intermediate state of the BTL code *) -Parameter measure: iblock -> nat. +Fixpoint measure ib: nat := + match ib with + | Bseq ib1 ib2 + | Bcond _ _ ib1 ib2 _ => measure ib1 + measure ib2 + | ib => 1 + end. Definition omeasure (oib: option iblock): nat := match oib with @@ -246,8 +251,17 @@ Theorem opt_simu s1 t s1' oib s2: \/ (omeasure oib' < omeasure oib /\ t=E0 /\ match_states oib' s1' s2) . Proof. - induction 2. - - admit. + inversion 2; subst; clear H0. + - (* State *) + inv H. + + 1: { (* Inop *) + eexists; right; split. + 2: split; trivial. + 2: { + econstructor; eauto. + all: admit. } admit. } + all: admit. - (* Callstate *) inv H. + (* Internal function *) -- cgit From 89aececb825a04dbc1982ec9e61331ef3272c228 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 7 May 2021 13:35:35 +0200 Subject: intermediate lemma for opt_simu intro case --- scheduling/RTLtoBTLproof.v | 31 +++++++++++++++++++++---------- 1 file changed, 21 insertions(+), 10 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7d58de5a..f5ea9fdc 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -243,6 +243,24 @@ Definition omeasure (oib: option iblock): nat := | Some ib => measure ib end. +Lemma opt_simu_intro + sp f f' st st' dupmap isfst pcR0 pcR1 pcB0 ib ib0 rs m rs0 m0 t s1' + (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t s1') + (STACKS : list_forall2 match_stackframes st st') + (TRANSF : match_function dupmap f f') + (ATpc0 : (fn_code f') ! pcB0 = Some ib0) + (DUPLIC : dupmap ! pcB0 = Some pcR0) + (MIB : match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) + (RIGHTA : is_right_assoc ib) + (RTL_RUN : star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 + (RTL.State st f sp pcR1 rs m)) + (BTL_RUN : iblock_istep_run tge sp (entry ib0) rs0 m0 = + iblock_istep_run tge sp ib rs m) + : exists (oib : option iblock), + omeasure oib < omeasure (Some ib) /\ + t = E0 /\ match_states oib s1' (State st' f' sp pcB0 rs0 m0). +Admitted. + Theorem opt_simu s1 t s1' oib s2: RTL.step ge s1 t s1' -> match_states oib s1 s2 -> @@ -253,15 +271,8 @@ Theorem opt_simu s1 t s1' oib s2: Proof. inversion 2; subst; clear H0. - (* State *) - inv H. - - 1: { (* Inop *) - eexists; right; split. - 2: split; trivial. - 2: { - econstructor; eauto. - all: admit. } admit. } - all: admit. + exploit opt_simu_intro; eauto. + intros (oib' & SIMU); exists oib'; right; assumption. - (* Callstate *) inv H. + (* Internal function *) @@ -295,7 +306,7 @@ Proof. + econstructor; eauto. eapply code_right_assoc; eauto. constructor. -Admitted. +Qed. Local Hint Resolve plus_one star_refl: core. -- cgit From 388a5077d75e6021611abf1301e16ad39e9f3770 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 7 May 2021 15:52:44 +0200 Subject: start a model of symbolic execution in Continuation-Passing-Style --- scheduling/BTL_SEtheory.v | 251 ++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 233 insertions(+), 18 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 4d5314f2..8aa8f32b 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -13,12 +13,11 @@ Require Import RTL BTL OptionMonad. Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) - Record iblock_exec_context := Bctx { cge: BTL.genv; - csp: val; cstk: list stackframe; cf: function; + csp: val; crs0: regset; cm0: mem }. @@ -454,17 +453,17 @@ Proof. reflexivity. Qed. -Variable sp: val. Variable stk stk': list stackframe. Variable f f': function. +Variable sp: val. Variable rs0: regset. Variable m0: mem. Lemma eval_sval_preserved sv: - eval_sval (Bctx ge sp stk f rs0 m0) sv = eval_sval (Bctx ge' sp stk' f' rs0 m0) sv. + eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. Proof. - induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge sp stk f rs0 m0) lsv = eval_list_sval (Bctx ge' sp stk' f' rs0 m0) lsv) - (P1 := fun sm => eval_smem (Bctx ge sp stk f rs0 m0) sm = eval_smem (Bctx ge' sp stk' f' rs0 m0) sm); simpl; auto. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) + (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. erewrite eval_operation_preserved; eauto. @@ -482,8 +481,8 @@ Proof. Qed. Lemma seval_builtin_arg_preserved m: forall bs varg, - seval_builtin_arg (Bctx ge sp stk f rs0 m0) m bs varg -> - seval_builtin_arg (Bctx ge' sp stk' f' rs0 m0) m bs varg. + seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> + seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. Proof. induction 1; simpl. all: try (constructor; auto). @@ -493,15 +492,15 @@ Proof. Qed. Lemma seval_builtin_args_preserved m lbs vargs: - seval_builtin_args (Bctx ge sp stk f rs0 m0) m lbs vargs -> - seval_builtin_args (Bctx ge' sp stk' f' rs0 m0) m lbs vargs. + seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> + seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. Proof. induction 1; constructor; eauto. eapply seval_builtin_arg_preserved; auto. Qed. Lemma list_sval_eval_preserved lsv: - eval_list_sval (Bctx ge sp stk f rs0 m0) lsv = eval_list_sval (Bctx ge' sp stk' f' rs0 m0) lsv. + eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. Proof. induction lsv; simpl; auto. rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. @@ -509,7 +508,7 @@ Proof. Qed. Lemma smem_eval_preserved sm: - eval_smem (Bctx ge sp stk f rs0 m0) sm = eval_smem (Bctx ge' sp stk' f' rs0 m0) sm. + eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. Proof. induction sm; simpl; auto. rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. @@ -520,7 +519,7 @@ Proof. Qed. Lemma seval_condition_preserved cond lsv sm: - seval_condition (Bctx ge sp stk f rs0 m0) cond lsv sm = seval_condition (Bctx ge' sp stk' f' rs0 m0) cond lsv sm. + seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. Proof. unfold seval_condition. rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. @@ -618,19 +617,235 @@ Proof. congruence. Qed. + +(** * symbolic execution of basic instructions *) + +Definition sis_init : sistate := {| si_pre:= fun _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. + +Lemma sis_init_correct ctx: + sem_sistate ctx sis_init (crs0 ctx) (cm0 ctx). +Proof. + unfold sis_init, sem_sistate; simpl; intuition eauto. +Qed. + +Definition set_sreg (r:reg) (sv:sval) (sis:sistate): sistate := + {| si_pre:=(fun ctx => eval_sval ctx (sis.(si_sreg) r) <> None /\ (sis.(si_pre) ctx)); + si_sreg:=fun y => if Pos.eq_dec r y then sv else sis.(si_sreg) y; + si_smem:= sis.(si_smem)|}. + +Lemma set_sreg_correct ctx dst sv sis (rs rs': regset) m: + sem_sistate ctx sis rs m -> + (eval_sval ctx sv = Some rs' # dst) -> + (forall r, r <> dst -> rs'#r = rs#r) -> + sem_sistate ctx (set_sreg dst sv sis) rs' m. +Proof. + intros (PRE&MEM®) NEW OLD. + unfold sem_sistate, set_sreg; simpl. + intuition. + - rewrite REG in *; congruence. + - destruct (Pos.eq_dec dst r); simpl; subst; eauto. + rewrite REG in *. rewrite OLD; eauto. +Qed. + +Definition set_smem (sm:smem) (sis:sistate): sistate := + {| si_pre:=(fun ctx => eval_smem ctx sis.(si_smem) <> None /\ (sis.(si_pre) ctx)); + si_sreg:= sis.(si_sreg); + si_smem:= sm |}. + +Lemma set_smem_correct ctx sm sis rs m m': + sem_sistate ctx sis rs m -> + eval_smem ctx sm = Some m' -> + sem_sistate ctx (set_smem sm sis) rs m'. +Proof. + intros (PRE&MEM®) NEW. + unfold sem_sistate, set_smem; simpl. + intuition. + rewrite MEM in *; congruence. +Qed. + +Definition sexec_op op args dst sis: sistate := + let args := list_sval_inj (List.map sis.(si_sreg) args) in + set_sreg dst (Sop op args sis.(si_smem)) sis. + +Lemma sexec_op_correct ctx op args dst sis rs m v + (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = Some v) + (SIS: sem_sistate ctx sis rs m) + :(sem_sistate ctx (sexec_op op args dst sis) (rs#dst <- v) m). +Proof. + eapply set_sreg_correct; eauto. + - simpl. destruct SIS as (PRE&MEM®). + rewrite Regmap.gss; simpl; auto. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + - intros; rewrite Regmap.gso; auto. +Qed. + +Definition sexec_load trap chunk addr args dst sis: sistate := + let args := list_sval_inj (List.map sis.(si_sreg) args) in + set_sreg dst (Sload sis.(si_smem) trap chunk addr args) sis. + +Lemma sexec_load_TRAP_correct ctx chunk addr args dst sis rs m a v + (EVAL: eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a) + (LOAD: Mem.loadv chunk m a = Some v) + (SIS: sem_sistate ctx sis rs m) + :(sem_sistate ctx (sexec_load TRAP chunk addr args dst sis) (rs#dst <- v) m). +Proof. + eapply set_sreg_correct; eauto. + - simpl. destruct SIS as (PRE&MEM®). + rewrite Regmap.gss; simpl; auto. + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + - intros; rewrite Regmap.gso; auto. +Qed. + +Definition sexec_store chunk addr args src sis: sistate := + let args := list_sval_inj (List.map sis.(si_sreg) args) in + let src := sis.(si_sreg) src in + let sm := Sstore sis.(si_smem) chunk addr args src in + set_smem sm sis. + +Lemma sexec_store_correct ctx chunk addr args src sis rs m m' a + (EVAL: eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a) + (STORE: Mem.storev chunk m a (rs # src) = Some m') + (SIS: sem_sistate ctx sis rs m) + :(sem_sistate ctx (sexec_store chunk addr args src sis) rs m'). +Proof. + eapply set_smem_correct; eauto. + simpl. destruct SIS as (PRE&MEM®). + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + rewrite REG; auto. +Qed. + +Lemma seval_condition_eq ctx cond args sis rs m + (SIS : sem_sistate ctx sis rs m) + :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) (si_smem sis) = eval_condition cond rs ## args m. +Proof. + destruct SIS as (PRE&MEM®); unfold seval_condition; simpl. + erewrite eval_list_sval_inj; simpl; auto. + erewrite MEM; auto. +Qed. + +(** * symbolic execution of blocks *) + (* symbolic state *) Inductive sstate := | Sfinal (sis: sistate) (sfv: sfval) | Scond (cond: condition) (args: list_sval) (sm: smem) (ifso ifnot: sstate) + | Sdead . -Inductive sem_internal_sstate ctx rs m t s: sstate -> Prop := - | sem_Sfinal sis sfv +(* transition (t,s) produced by a sstate in initial context ctx *) +Inductive sem_sstate ctx t s: sstate -> Prop := + | sem_Sfinal sis sfv rs m (SIS: sem_sistate ctx sis rs m) (SFV: sem_sfval ctx sfv rs m t s) - : sem_internal_sstate ctx rs m t s (Sfinal sis sfv) + : sem_sstate ctx t s (Sfinal sis sfv) | sem_Scond b cond args sm ifso ifnot (SEVAL: seval_condition ctx cond args sm = Some b) - (SELECT: sem_internal_sstate ctx rs m t s (if b then ifso else ifnot)) - : sem_internal_sstate ctx rs m t s (Scond cond args sm ifso ifnot) + (SELECT: sem_sstate ctx t s (if b then ifso else ifnot)) + : sem_sstate ctx t s (Scond cond args sm ifso ifnot) + (* NB: Sdead: fails to produce a transition *) . + +(** * Idée de l'execution symbolique en Continuation Passing Style + +[k] ci-dessous est la continuation (c-a-d. la suite de la construction de l'arbre qu'on va appliquer dans chaque branche) + +Rem: si manipuler une telle continuation s'avère compliqué dans les preuves, +il faudra faire autrement dans le modèle -- par exemple avec une structure de donnée +pour représenter l'ensemble des chemins. +(même si on peut conserver le CPS dans l'implem finale, pour l'efficacité). + +*) + +Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := + match ib with + | BF fin => Sfinal sis (sexec_final_sfv fin sis) + (* basic instructions *) + | Bnop => k sis + | Bop op args res => k (sexec_op op args res sis) + | Bload TRAP chunk addr args dst => k (sexec_load TRAP chunk addr args dst sis) + | Bload NOTRAP chunk addr args dst => Sdead (* TODO *) + | Bstore chunk addr args src => k (sexec_store chunk addr args src sis) + (* composed instructions *) + | Bseq ib1 ib2 => + sexec_rec ib1 sis (fun sis2 => sexec_rec ib2 sis2 k) + | Bcond cond args ifso ifnot _ => + let args := list_sval_inj (List.map sis.(si_sreg) args) in + let ifso := sexec_rec ifso sis k in + let ifnot := sexec_rec ifnot sis k in + Scond cond args sis.(si_smem) ifso ifnot + end + . + +Definition sexec ib := sexec_rec ib sis_init (fun _ => Sdead). + +Local Hint Constructors sem_sstate: core. +Local Hint Resolve sexec_final_svf_correct sis_init_correct: core. + +Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin + (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): + forall k + (CONT: forall sis, ofin = None -> sem_sistate ctx sis rs1 m1 -> sem_sstate ctx t s (k sis)) + sis + (SIS: sem_sistate ctx sis rs m) + , match ofin with + | Some fin => + final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s -> + sem_sstate ctx t s (sexec_rec ib sis k) + | None => sem_sstate ctx t s (sexec_rec ib sis k) + end. +Proof. + induction ISTEP; simpl; eauto. + - (* op *) + intros; eapply CONT; eauto. + eapply sexec_op_correct; eauto. + - (* load *) + intros; eapply CONT; eauto. + eapply sexec_load_TRAP_correct; eauto. + - (* store *) + intros; eapply CONT; eauto. + eapply sexec_store_correct; eauto. + - (* seq_stop *) + intros; eapply IHISTEP; eauto. + try congruence. + - (* seq_continue *) + intros. destruct ofin as [fin|]. + + (* Some *) + clear CONT. + intros; eapply IHISTEP1; eauto. + intros; eapply IHISTEP2; eauto. + try congruence. + + (* None *) + intros; eapply IHISTEP1; eauto. + - (* cond *) + intros. + assert (IFEQ: (if b then sexec_rec ifso sis k else sexec_rec ifnot sis k) = sexec_rec (if b then ifso else ifnot) sis k). + { destruct b; simpl; auto. } + destruct ofin as [fin|]. + + (* Some *) + clear CONT. + intro FSTEP; eapply sem_Scond; eauto. + * erewrite seval_condition_eq; eauto. + * rewrite IFEQ. + eapply IHISTEP; eauto. + try congruence. + + (* None *) + eapply sem_Scond; eauto. + * erewrite seval_condition_eq; eauto. + * rewrite IFEQ. + eapply IHISTEP; eauto. +Qed. + +(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) + (sexec is a correct over-approximation) +*) +Theorem sexec_correct ctx ib t s: + iblock_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> + sem_sstate ctx t s (sexec ib). +Proof. + destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). + unfold sexec. + exploit sexec_rec_correct; simpl; eauto || congruence. +Qed. -- cgit From 1d32156c5bc43f966abe2c12c317a27e092355c4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 7 May 2021 17:30:43 +0200 Subject: idea of a new scheme to define the semantics of LOAD NOTRAP --- scheduling/BTL.v | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index fb67d150..8ccab024 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -175,6 +175,23 @@ Definition find_function (ros: reg + ident) (rs: regset) : option fundef := Local Open Scope option_monad_scope. +(* (* TODO: a new (hopefully simpler) scheme to support "NOTRAP" wrt current scheme of RTL *) + +Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop := + | has_loaded_normal a trap + (EVAL: eval_addressing ge sp addr rs##args = Some a) + (LOAD: Mem.loadv chunk m a = Some v) + : has_loaded sp rs m chunk addr args v trap + | has_loaded_default + (LOAD: forall a, eval_addressing ge sp addr rs##args = Some a -> Mem.loadv chunk m a = None) + (DEFAULT: v = default_notrap_load_value chunk) + : has_loaded sp rs m chunk addr args v NOTRAP + . + +(* TODO: move this scheme in "Memory" module if this scheme is useful ! *) + +*) + (** 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) @@ -186,6 +203,11 @@ Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option fi (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 +(* 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' (EVAL: eval_addressing ge sp addr rs##args = Some a) (STORE: Mem.storev chunk m a rs#src = Some m') -- cgit From 08b39b2025951a9655939c1377f6e53b346c6821 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 7 May 2021 19:58:02 +0200 Subject: Some advance in proof and NOTRAP loads fix --- scheduling/BTL.v | 13 ++++++----- scheduling/RTLtoBTLproof.v | 57 ++++++++++++++++++++++++++++++++++++++++++---- 2 files changed, 60 insertions(+), 10 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index fb67d150..c8dfa0af 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -400,9 +400,9 @@ Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> i | mib_op isfst pc pc' op lr r: 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' tm m a lr r: - cfg!pc = Some (Iload tm m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bload tm m a lr r) (Some pc') + | mib_load isfst pc pc' m a lr r: + 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: cfg!pc = Some (Istore m a lr r pc') -> match_iblock dupmap cfg isfst pc (Bstore m a lr r) (Some pc') @@ -581,7 +581,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) | Bload tm m a lr r => match cfg!pc with | Some (Iload tm' m' a' lr' r' pc') => - if (trapping_mode_eq tm tm') then + if (trapping_mode_eq tm TRAP && trapping_mode_eq tm' TRAP) then if (chunk_eq m m') then if (eq_addressing a a') then if (list_eq_dec Pos.eq_dec lr lr') then @@ -591,7 +591,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "verify_block: different lr in Bload") else Error (msg "verify_block: different addressing in Bload") else Error (msg "verify_block: different mchunk in Bload") - else Error (msg "verify_block: different trapping_mode in Bload") + else Error (msg "verify_block: NOTRAP trapping_mode unsupported in Bload") | _ => Error (msg "verify_block: incorrect cfg Bload") end | Bstore m a lr r => @@ -691,7 +691,8 @@ Proof. constructor; assumption. - (* Bload *) destruct i; try discriminate. - destruct (trapping_mode_eq _ _); try discriminate. + do 2 (destruct (trapping_mode_eq _ _); try discriminate). + simpl in H. destruct (chunk_eq _ _); try discriminate. destruct (eq_addressing _ _); try discriminate. destruct (list_eq_dec _ _ _); try discriminate. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index f5ea9fdc..023837f9 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -256,9 +256,59 @@ Lemma opt_simu_intro (RTL.State st f sp pcR1 rs m)) (BTL_RUN : iblock_istep_run tge sp (entry ib0) rs0 m0 = iblock_istep_run tge sp ib rs m) - : exists (oib : option iblock), - omeasure oib < omeasure (Some ib) /\ - t = E0 /\ match_states oib s1' (State st' f' sp pcB0 rs0 m0). + : exists (oib' : option iblock), + (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') + \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). +Proof. + inv MIB. + - (* mib_BF *) + admit. + - (* mib_exit *) + (* destruct t. + 2: { + eexists; left. eexists; split. + + eapply exec_iblock; eauto. + econstructor; repeat eexists. + eapply iblock_istep_run_equiv; eauto. + inv STEP. + eapply exec_Bgoto.*) admit. + - (* mib_seq *) + inv H. + + (* Bnop *) + inversion STEP; subst; try_simplify_someHyps. (* TODO maybe move that for each subcase *) + exists (Some b2); right; repeat (split; auto). + econstructor; eauto. inv RIGHTA; auto; discriminate. + eapply star_right; eauto. + + (* Bop *) + inversion STEP; subst; try_simplify_someHyps. + exists (Some b2); right; repeat (split; auto). + econstructor; eauto. inv RIGHTA; auto; discriminate. + eapply star_right; eauto. + erewrite eval_operation_preserved in H10. + erewrite H10 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. + + (* Bload *) + inversion STEP; subst; try_simplify_someHyps. + exists (Some (b2)); right; repeat (split; auto). + econstructor; eauto. inv RIGHTA; auto; discriminate. + eapply star_right; eauto. + erewrite eval_addressing_preserved in H10. + erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. + + (* Bstore *) + inversion STEP; subst; try_simplify_someHyps. + exists (Some b2); right; repeat (split; auto). + econstructor; eauto. inv RIGHTA; auto; discriminate. + eapply star_right; eauto. + erewrite eval_addressing_preserved in H10. + erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. + + (* Absurd case *) + inv RIGHTA. inv H4. inv H. + + (* Bcond *) + admit. + - (* mib_cond *) + admit. Admitted. Theorem opt_simu s1 t s1' oib s2: @@ -272,7 +322,6 @@ Proof. inversion 2; subst; clear H0. - (* State *) exploit opt_simu_intro; eauto. - intros (oib' & SIMU); exists oib'; right; assumption. - (* Callstate *) inv H. + (* Internal function *) -- cgit From a1cde0686dfb40595423f40ccc40e18de3539e52 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 8 May 2021 07:31:05 +0200 Subject: idee pour simplifier la preuve: restreindre le "right_assoc" en "expand" ! --- scheduling/BTL.v | 66 ++++++++++++++++++++++++++++++++++++++++++---- scheduling/RTLtoBTL.v | 3 +++ scheduling/RTLtoBTLproof.v | 2 ++ 3 files changed, 66 insertions(+), 5 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 8ccab024..7f3d327b 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -785,11 +785,6 @@ Definition verify_function dupmap f f' : res unit := do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); verify_cfg dupmap (fn_code f) (RTL.fn_code f'). -Definition is_seq (ib: iblock): bool := - match ib with - | Bseq _ _ => true - | _ => false - end. Definition is_atom (ib: iblock): bool := match ib with @@ -797,6 +792,67 @@ Definition is_atom (ib: iblock): bool := | _ => true end. +Inductive is_expand: iblock -> Prop := + | exp_Bseq ib1 ib2: + is_atom ib1 = true -> + is_expand ib2 -> + is_expand (Bseq ib1 ib2) + | exp_Bcond cond args ib1 ib2 i: + is_expand ib1 -> + is_expand ib2 -> + is_expand (Bcond cond args ib1 ib2 i) + | exp_others ib: + is_atom ib = true -> + is_expand ib + . +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 + | ib => + match k with + | None => ib + | Some rem => Bseq ib rem + end + end. + +Lemma expand_correct ib: forall k, + (match k with Some rem => is_expand rem | None => True end) + -> is_expand (expand ib k). +Proof. + induction ib; simpl; intros; try autodestruct; auto. +Qed. + +Lemma expand_None_correct ib: + is_expand (expand ib None). +Proof. + apply expand_correct; simpl; auto. +Qed. + +Definition expand_code (cfg: code): code := + (PTree.map (fun _ ib => {| entry:=expand ib.(entry) None; input_regs := ib.(input_regs) |}) cfg). + +Lemma expand_code_correct cfg pc ib : + (expand_code cfg)!pc=Some ib -> is_expand (ib.(entry)). +Proof. + unfold expand_code. + rewrite PTree.gmap. + destruct (cfg!pc); simpl; intros; try_simplify_someHyps. + apply expand_None_correct; auto. +Qed. + + +(* TODO: eliminer toute la suite quand "right_assoc" sera remplace par "expand" *) + +Definition is_seq (ib: iblock): bool := + match ib with + | Bseq _ _ => true + | _ => false + end. + Inductive is_right_assoc: iblock -> Prop := | ira_Bseq ib1 ib2: diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index 14aecb21..2b2bd15c 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -12,6 +12,9 @@ Local Open Scope error_monad_scope. Definition transf_function (f: RTL.function) : res BTL.function := let (tcte, dupmap) := rtl2btl f in let (tc, te) := tcte in + (* TODO, pour finir la preuve: remplacer ci-dessous "right_assoc_code" par "expand_code" + on s'arrangera pour éliminer cette transformation "coûteuse" à la fin ! + *) let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) (right_assoc_code tc) te in do u <- verify_function dupmap f' f; OK f'. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7d58de5a..7c9fb456 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -6,6 +6,8 @@ Require Import Errors Linking RTLtoBTL. Require Import Linking. +(* TODO: remplacer "right_assoc" par "expand" ci-dessous *) + Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); -- cgit From d422c63cbcad7ba156d5d324e0221db9d13f9559 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 8 May 2021 07:56:20 +0200 Subject: sexec: renommage Sdead -> Sabort + simplification de sexec_correct --- scheduling/BTL_SEtheory.v | 78 +++++++++++++---------------------------------- 1 file changed, 21 insertions(+), 57 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 8aa8f32b..0a2c0a04 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -732,7 +732,7 @@ Qed. Inductive sstate := | Sfinal (sis: sistate) (sfv: sfval) | Scond (cond: condition) (args: list_sval) (sm: smem) (ifso ifnot: sstate) - | Sdead + | Sabort . (* transition (t,s) produced by a sstate in initial context ctx *) @@ -745,7 +745,7 @@ Inductive sem_sstate ctx t s: sstate -> Prop := (SEVAL: seval_condition ctx cond args sm = Some b) (SELECT: sem_sstate ctx t s (if b then ifso else ifnot)) : sem_sstate ctx t s (Scond cond args sm ifso ifnot) - (* NB: Sdead: fails to produce a transition *) + (* NB: Sabort: fails to produce a transition *) . (** * Idée de l'execution symbolique en Continuation Passing Style @@ -766,7 +766,7 @@ Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := | Bnop => k sis | Bop op args res => k (sexec_op op args res sis) | Bload TRAP chunk addr args dst => k (sexec_load TRAP chunk addr args dst sis) - | Bload NOTRAP chunk addr args dst => Sdead (* TODO *) + | Bload NOTRAP chunk addr args dst => Sabort (* TODO *) | Bstore chunk addr args src => k (sexec_store chunk addr args src sis) (* composed instructions *) | Bseq ib1 ib2 => @@ -779,63 +779,28 @@ Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := end . -Definition sexec ib := sexec_rec ib sis_init (fun _ => Sdead). +Definition sexec ib := sexec_rec ib sis_init (fun _ => Sabort). Local Hint Constructors sem_sstate: core. -Local Hint Resolve sexec_final_svf_correct sis_init_correct: core. +Local Hint Resolve sexec_op_correct sexec_final_svf_correct + sexec_load_TRAP_correct sexec_store_correct sis_init_correct: core. Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin - (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): - forall k - (CONT: forall sis, ofin = None -> sem_sistate ctx sis rs1 m1 -> sem_sstate ctx t s (k sis)) - sis + (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): forall sis k (SIS: sem_sistate ctx sis rs m) - , match ofin with - | Some fin => - final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s -> - sem_sstate ctx t s (sexec_rec ib sis k) - | None => sem_sstate ctx t s (sexec_rec ib sis k) - end. -Proof. - induction ISTEP; simpl; eauto. - - (* op *) - intros; eapply CONT; eauto. - eapply sexec_op_correct; eauto. - - (* load *) - intros; eapply CONT; eauto. - eapply sexec_load_TRAP_correct; eauto. - - (* store *) - intros; eapply CONT; eauto. - eapply sexec_store_correct; eauto. - - (* seq_stop *) - intros; eapply IHISTEP; eauto. - try congruence. - - (* seq_continue *) - intros. destruct ofin as [fin|]. - + (* Some *) - clear CONT. - intros; eapply IHISTEP1; eauto. - intros; eapply IHISTEP2; eauto. - try congruence. - + (* None *) - intros; eapply IHISTEP1; eauto. - - (* cond *) - intros. - assert (IFEQ: (if b then sexec_rec ifso sis k else sexec_rec ifnot sis k) = sexec_rec (if b then ifso else ifnot) sis k). - { destruct b; simpl; auto. } - destruct ofin as [fin|]. - + (* Some *) - clear CONT. - intro FSTEP; eapply sem_Scond; eauto. - * erewrite seval_condition_eq; eauto. - * rewrite IFEQ. - eapply IHISTEP; eauto. - try congruence. - + (* None *) - eapply sem_Scond; eauto. - * erewrite seval_condition_eq; eauto. - * rewrite IFEQ. - eapply IHISTEP; eauto. + (CONT: match ofin with + | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx t s (k sis') + | Some fin => final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s + end), + sem_sstate ctx t s (sexec_rec ib sis k). +Proof. + induction ISTEP; simpl; try autodestruct; eauto. + (* condition *) + all: intros; + eapply sem_Scond; eauto; [ + erewrite seval_condition_eq; eauto | + replace (if b then sexec_rec ifso sis k else sexec_rec ifnot sis k) with (sexec_rec (if b then ifso else ifnot) sis k); + try autodestruct; eauto ]. Qed. (* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) @@ -846,6 +811,5 @@ Theorem sexec_correct ctx ib t s: sem_sstate ctx t s (sexec ib). Proof. destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). - unfold sexec. - exploit sexec_rec_correct; simpl; eauto || congruence. + eapply sexec_rec_correct; simpl; eauto. Qed. -- cgit From c78393a5d4d13d0f2cd7f0a73756d6bb598b5ae4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sun, 9 May 2021 22:41:40 +0200 Subject: mib_exit draft --- scheduling/RTLtoBTLproof.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 023837f9..1db088d2 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -264,12 +264,13 @@ Proof. - (* mib_BF *) admit. - (* mib_exit *) - (* destruct t. - 2: { + (* + simpl in *. + eapply iblock_istep_run_equiv in BTL_RUN. + inv BTL_RUN. eexists; left. eexists; split. + eapply exec_iblock; eauto. econstructor; repeat eexists. - eapply iblock_istep_run_equiv; eauto. inv STEP. eapply exec_Bgoto.*) admit. - (* mib_seq *) -- cgit From 8ae28f605b558447bb7d31542f9fbaecfa857ecc Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 10 May 2021 09:26:02 +0200 Subject: précision du roadmap MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- scheduling/BTLroadmap.md | 234 ++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 230 insertions(+), 4 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 27e74de5..05dcb95e 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -16,6 +16,29 @@ Examples of optimizations that we aim to support: We sketch below the various kinds of supported optimizations in development order... +## Introduction + +En gros la syntaxe d'un bloc BTL est définie par: + + Inductive iblock: Type := + | ... (* instructions basiques ou "finales" (call, return, goto, etc) *) + | Bseq (b1 b2: iblock) (* séquence de deux blocs *) + | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (* if-then-else *) + +Le modèle de base de l'exécution symbolique représente un tel bloc par un état symbolique de type: + + Inductive sstate := + | Sfinal (sis: sistate) (sfv: sfval) + | Scond (cond: condition) (args: list_sval) (ifso ifnot: sstate) + | Sabort + . + +où `sistate` est un PPA (preconditioned parallel assignment) des registres et `sfval` représente un branchement (call, return, goto, etc). + +Autrement dit, un état symbolique représente tous les chemins +d'exécution possibles par une sorte de gros BDD ayant sur les feuilles +un `Sfinal` (ou un `Sabort` s'il manque une instruction de branchement sur ce chemin). + ## Block boundaries, branch duplication or factorization Two possibility of branch duplications (e.g tail-duplication, loop unrolling, etc): @@ -25,12 +48,21 @@ Two possibility of branch duplications (e.g tail-duplication, loop unrolling, et Branch factorization should also be possible in BTL -> RTL pass. Example: revert "loop-unrolling". -**IMPLEM NOTE:** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. +**IMPLEM NOTE:** in a first step a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. + **CURRENT STATUS** -- verifier: specified but not yet implemented. -- BTL -> RTL: proof started. +- verifier: implemented and proved w.r.t match_iblock specification. +- BTL -> RTL: done. +- RTL -> BTL: started. + +**TODO** + +- lien BTL/RTL: autoriser le BTL à avoir des Bnop en plus du RTL, e.g. pour autoriser des "if-then" sans else. +Pour faciliter le vérificateur, faire comme le Bgoto: les Bnop en "trop" ne sont autorisés que s'il y a eu une instruction RTL avant. +Ajouter aussi un booléen (positionné par l'oracle) sur le Bnop qui indique si le "nop" existe ou pas au niveau RTL. + ## "Basic" symbolic execution / symbolic simulation @@ -41,7 +73,9 @@ IMPLEM PLAN in two steps: 2. Support a "SVundef" symbolic value, and implement symbolic simulation wrt "less_undef" relation. This will allow more general rewriting rules and will generalize modulo-liveness (see next section). -**IMPLEM NOTE:** use a symbolic execution in Continuation-Passing Style, in order to have of simple handling of "side-exits" within "branch unfolding". +**CURRENT STATUS** + +- model of symbolic execution in Continuation-Passing Style, with basic "correctness" thm ("completness" not yet done). ## Simulation modulo liveness @@ -67,6 +101,198 @@ Semantically, the destruction of registers happens "in parallel" of the final st - Port rewriting rules of RTLpath. - Justus's "poor man SSA" + if-lifting. +## Efficient comparison (guided by oracles) of "if-then-else" sequences. + +Le pb est complexe. Comparer des expressions booleennes contenant juste des variables booleennes est déjà NP-complet, avec "explosion exponentielle" dans le pire cas. + +Approche proposée: utiliser un mécanisme de vérification "simple", basée sur une comparaison par execution symbolique de "tous" les chemins d'execution (cf Intro ci-dessus). +Ça pète exponentiellement dans le pire cas: mais on pourra contrôler ce risque d'explosion par les oracles... + +Ci-dessous, on liste quelques "techniques" de collaboration oracle/vérificateur pour contrôler l'explosion des chemins. +Idée: les conditions comportent une liste d'annotations qui permet le guidage du vérificateur par l'oracle. + +### Contrôle des "joins internes" dans le bloc. + +Si dans le bloc, toute condition a au plus un "predecesseur" (au sens +du CFG RTL) dans le bloc: alors le nombre de "chemins sémantiques" +(explorés par l'exécution symbolique) est identique au nombre de +"branches syntaxiques" (présents dans le code). Une façon simple de +contrôler l'explosion de l'exécution symbolique est de fabriquer (avec +les oracles) des blocs avec un nombre borné (petit) de "joins +internes". + +**Exemple d'application: généralisation des superblocks** + +On considère le bloc BTL ci-dessous (où les `i*` sont des séquences d'instructions basiques sans branchement): + + i0; + if (c1) { i1 } else { i1'; goto pc1 } + if (c2) { i2; goto pc2 } else { i2' } + if (c3} { i3; goto pc3 } else { i3'; goto pc3' } + +Sur un tel bloc, il n'y a aucun "join interne" (l'exécution symbolique est donc linéaire). +Mais représenter en RTLpath un tel bloc nécessite au moins 4 blocks (1 superbloc et 3 basic-blocs): + + i0; c1; i1; c2; i2'; i3'; goto pc3' + i1'; goto pc1 + i2; goto pc2 + i3; goto pc3 + +La vérification BTL sur le gros bloc ne prendra à priori pas plus de +temps que la vérification RTLpath cumulée des 4 "petits" blocs. Mais +la vérification BTL sera plus *puissante*, puisque que quand on va +vérifier les chemins d'exécutions correspondant à ceux des 3 +basic-blocs, on aura le `i0` en plus dans l'état symbolique (i.e. un +"contexte d'exécution" plus précis). + +**Autre exemple d'application: le if-lifting à la Justus** + +Le superblock suivant: + + y1 = e1(x) + x = e2(a) + y2 = e3(x) + if (c[x]) { goto pc } else { i4; goto pc' } + +peut être directement montré équivalent à + + x' = e2(a) // x' un registre "frais" (pas dans les "liveout") + if (c[x']) { + y1 = e1(x); + x = x'; + y2 = e3(x); + goto pc + } else { + y1 = e1(x); + x = x'; + y2 = e3(x); + i4; + goto pc' + } + +Ici, la duplication de branche a donc lieu en BTL. + +L'exécution symbolique de ces deux blocs va produire deux BDD comparables (avec comparaison des feuilles modulo liveness). + +### Comparaison des BDD (modulo réordonnancement des conditions ?) + +On peut avoir envie de montrer que les deux blocs ci-dessous sont équivalents (si les dépendences sur les variables le permettent): + + if (c1) { i1 } else { i1' } + if (c2) { i2 } else { i2' } + +et + + if (c2) { i2 } else { i2' } + if (c1) { i1 } else { i1' } + +Ça revient (en gros) à comparer les BDD: + + if (c1) { if (c2) {i1;i2} else {i1;i2'} } else { if (c2) {i1';i2} else {i1';i2'} } + +et + + if (c2) { if (c1) {i2;i1} else {i2;i1'} } else { if (c1) {i2';i1} else {i2';i1'} } + +Pour gérer ça, on peut faire des "Ordered BDD": l'oracle du **bloc +transformé** annote (certaines) conditions avec des numéros de façon à +ce l'exécution symbolique du bloc transformé produise un "BDD" qui +correspond à celui du bloc d'origine (cf. "Principe" +ci-dessous). Cependant, il semble difficile d'appliquer complètement +les techniques de mémoïsation des BDD ayant des booléens sur les +feuilles. Car ici on veut effectuer une comparaison sur des feuilles +2 à 2 qui n'est pas une égalité, mais une inclusion ! + +**Principe du réordonnancement de BDD:** l'exécution symbolique du **bloc transformé** réordonne le BDD de +façon à respecter la numérotation: un pére doit avoir un numéro inférieur à +chacun de ses fils (en l'absence de numéro, on ignore les contraintes +de numérotation entre ce noeud est ses voisins). Exemple ci-dessous +avec trois conditions distinctes (pour order c1 < c2 < c3): + + if (c3) { if (c1) {f1} else {f1'} } else { if (c2} {f2} else {f2'} } + +est réordonné en + + if (c1) { if (c2) { if (c3) {f1} else {f2} } else { if (c3) {f1} else {f2'} } } + else { if (c2) { if (c3) {f1'} else {f2} } else { if (c3) {f1'} else {f2'} } } + +**rem:** on ajoute ici un undefined behavior potentiel à cause l'execution de c2 quand c3 est vrai. +Mais si le **bloc d'origine** est simulé par cet état symbolique réordonné, c'est correct. +Le bloc transformé a juste supprimé un test inutile... + +Bon, à voir si le principe ci-dessus est vraiment utile dans toute sa +généralité. Pour simplifier, on peut aussi restreindre le +réordonnancement du BDD aux cas où il n'y a pas de test redondant +inséré (comme dans l'exemple initial). + +**Version simplifiée: comparaison des BDD sans réordonnancement des conditions** + +Dans un premier temps (jusqu'à ce qu'on trouve une optimisation où ça pourrait être utile): pas de réordonnacement des BDD. +On autorise juste le BDD du bloc transformé à supprimer des conditions par rapport au bloc d'origine. +Autrement dit, dans la comparaison récursive de `{if (c) BDD1 BDD2}` avec `{if (c') BDD1' BDD2}'`: + +- soit `c=c'` et on compare récursivement `BDD1` avec `BDD1'` et `BDD2` avec `BDD2'`. +- soit `c<>c'` et on compare récursivement `BDD1` et `BDD2` avec `{if (c') BDD1' BDD2'}` + +Ce deuxième cas correspond au fait que le test sur `c` dans le bloc d'origine était inutile! + +### Propagation de valeurs symbolique de conditions (et élimination de condition) pendant l'execution symbolique + +L'exécution symbolique se propageant en "avant", on peut propager les valeurs des conditions symboliques, qu'on peut utiliser pour éliminer des conditions redondantes +(et donc limiter l'explosion du nombre de chemin). + +Pour rendre ce mécanisme efficace et puissant, on peut guider ce mécanisme de propagation/élimination avec des annotations introduites par les oracles. + +- une annotation "bind_to x" qui indique de mémoriser la valeur (soit "true" soit "false" suivant la branche) de la condition symbolique avec le nom "x" +- une annotation "eval_to b proof" qui indique que la condition s'évalue sur la constante "b", ce qui est prouvable avec la preuve "proof". + +Ici on peut imaginer un langage plus ou moins compliqué pour prouver l'évaluation des conditions. La version la plus simple: + +- "eq(x)" dit simplement que la condition symbolique est syntaxiquement égale celle appelée "x". +- "eqnot(x)" dit que c'est la négation. + +Dans le cas général, on peut introduire tout un système de réécriture pour éliminer les conditions. + +En fait, il serait sans doute intéressant de mettre en place un +"système de réécriture guidé par oracle" pour toutes les instructions +BTL. Ça permet de concilier "puissance" de l'exécution symbolique et +"efficacité". L'exécution symbolique va pouvoir éventuellement faire +des réécritures compliquées, mais uniquement quand c'est nécessaire. + +**Exemple: une "if-conversion" généralisée** +On aimerait montrer que le bloc d'origine: + + if (c) { + x1 = e1 + x2 = e2 + y = e + x3 = e3 + } else { + x3 = e3' + z = e' + x1 = e1' + x2 = e2' + } + +est simulable par le bloc transformé: + + x1 = (c?e1:e1') + x2 = (c?e2:e2') + x3 = (c?e3:e3') + if (c) { y = e } else { z = e' } + +une solution: ajouter une régle de réécriture `x = (c?e:e')` en `if (c) { x=e } else {x=e'}` +(attention, ce n'est pas une règle de réécriture sur les valeurs +symboliques, mais sur du code BTL lui-même, avant l'exécution +symbolique de ce code). + +L'exécution symbolique ouvre alors deux branches `c=true` et +`c=false` sur `x1 = (c?e1:e1')`, puis la propagation/élimination de la +condition symbolique `c` simplifie les conditionnelles sur `x2`, `x3` et `y`/`z`. +Au final, on obtient deux BDD identiques sur cet exemple (sans explosion combinatoire). + +**Rem** les mécanismes de propagation/réécritures décrits ci-dessus peuvent être aussi très utile pour la simulation symbolique modulo invariants (cf. ci-dessous) ! + ## Invariants at block entry Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. -- cgit From 54d0ce9246ed1be12f150eb3745626a7576bf20b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 10 May 2021 10:28:55 +0200 Subject: is_exp and bcond proof --- scheduling/RTLtoBTL.v | 5 +---- scheduling/RTLtoBTLproof.v | 53 ++++++++++++++++++++++++++++++---------------- 2 files changed, 36 insertions(+), 22 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index 2b2bd15c..3617e691 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -12,10 +12,7 @@ Local Open Scope error_monad_scope. Definition transf_function (f: RTL.function) : res BTL.function := let (tcte, dupmap) := rtl2btl f in let (tc, te) := tcte in - (* TODO, pour finir la preuve: remplacer ci-dessous "right_assoc_code" par "expand_code" - on s'arrangera pour éliminer cette transformation "coûteuse" à la fin ! - *) - let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) (right_assoc_code tc) te in + let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) (expand_code tc) te in do u <- verify_function dupmap f' f; OK f'. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index efc22eb6..ac9cb8d8 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -1,4 +1,4 @@ -Require Import Coqlib Maps. +Require Import Coqlib Maps Lia. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. @@ -6,12 +6,10 @@ Require Import Errors Linking RTLtoBTL. Require Import Linking. -(* TODO: remplacer "right_assoc" par "expand" ci-dessous *) - Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); - code_right_assoc: forall pc ib, (fn_code tf)!pc=Some ib -> is_right_assoc ib.(entry); + code_expand: forall pc ib, (fn_code tf)!pc=Some ib -> is_expand ib.(entry); preserv_fnsig: fn_sig tf = RTL.fn_sig f; preserv_fnparams: fn_params tf = RTL.fn_params f; preserv_fnstacksize: fn_stacksize tf = RTL.fn_stacksize f @@ -33,7 +31,7 @@ Lemma verify_function_correct dupmap f f' tt: fn_sig f' = RTL.fn_sig f -> fn_params f' = RTL.fn_params f -> fn_stacksize f' = RTL.fn_stacksize f -> - (forall pc ib, (fn_code f')!pc=Some ib -> is_right_assoc ib.(entry)) -> + (forall pc ib, (fn_code f')!pc=Some ib -> is_expand ib.(entry)) -> match_function dupmap f f'. Proof. unfold verify_function; intro VERIF. monadInv VERIF. @@ -48,7 +46,7 @@ Proof. unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. - eapply right_assoc_code_correct; eauto. + eapply expand_code_correct; eauto. Qed. Lemma transf_fundef_correct f f': @@ -112,7 +110,7 @@ Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := (ATpc0: (fn_code f')!pcB0 = Some ib0) (DUPLIC: dupmap!pcB0 = Some pcR0) (MIB: match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) - (RIGHTA: is_right_assoc ib) + (IS_EXPD: is_expand ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) : match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) @@ -245,6 +243,12 @@ Definition omeasure (oib: option iblock): nat := | Some ib => measure ib end. +Lemma measure_pos: forall ib, + measure ib > 0. +Proof. + induction ib; simpl; auto; lia. +Qed. + Lemma opt_simu_intro sp f f' st st' dupmap isfst pcR0 pcR1 pcB0 ib ib0 rs m rs0 m0 t s1' (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t s1') @@ -253,7 +257,7 @@ Lemma opt_simu_intro (ATpc0 : (fn_code f') ! pcB0 = Some ib0) (DUPLIC : dupmap ! pcB0 = Some pcR0) (MIB : match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) - (RIGHTA : is_right_assoc ib) + (IS_EXPD : is_expand ib) (RTL_RUN : star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) (BTL_RUN : iblock_istep_run tge sp (entry ib0) rs0 m0 = @@ -280,12 +284,12 @@ Proof. + (* Bnop *) inversion STEP; subst; try_simplify_someHyps. (* TODO maybe move that for each subcase *) exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv RIGHTA; auto; discriminate. + econstructor; eauto. inv IS_EXPD; auto; discriminate. eapply star_right; eauto. + (* Bop *) inversion STEP; subst; try_simplify_someHyps. exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv RIGHTA; auto; discriminate. + econstructor; eauto. inv IS_EXPD; auto; discriminate. eapply star_right; eauto. erewrite eval_operation_preserved in H10. erewrite H10 in BTL_RUN; simpl in BTL_RUN; auto. @@ -293,7 +297,7 @@ Proof. + (* Bload *) inversion STEP; subst; try_simplify_someHyps. exists (Some (b2)); right; repeat (split; auto). - econstructor; eauto. inv RIGHTA; auto; discriminate. + econstructor; eauto. inv IS_EXPD; auto; discriminate. eapply star_right; eauto. erewrite eval_addressing_preserved in H10. erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. @@ -301,17 +305,30 @@ Proof. + (* Bstore *) inversion STEP; subst; try_simplify_someHyps. exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv RIGHTA; auto; discriminate. + econstructor; eauto. inv IS_EXPD; auto; discriminate. eapply star_right; eauto. erewrite eval_addressing_preserved in H10. erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. intros; rewrite <- symbols_preserved; trivial. + (* Absurd case *) - inv RIGHTA. inv H4. inv H. - + (* Bcond *) - admit. + inv IS_EXPD. inv H4. inv H. + + (* Absurd Bcond (Bcond are not allowed in the left part of a Bseq *) + inv IS_EXPD; discriminate. - (* mib_cond *) - admit. + inversion STEP; subst; try_simplify_someHyps. + intros; rewrite H12 in BTL_RUN. destruct b. + * (* Ifso *) + exists (Some bso); right; repeat (split; eauto). + simpl; assert (measure bnot > 0) by apply measure_pos; lia. + inv H2; econstructor; eauto. + 1,3: inv IS_EXPD; auto; discriminate. + all: eapply star_right; eauto. + * (* Ifnot *) + exists (Some bnot); right; repeat (split; eauto). + simpl; assert (measure bso > 0) by apply measure_pos; lia. + inv H2; econstructor; eauto. + 1,3: inv IS_EXPD; auto; discriminate. + all: eapply star_right; eauto. Admitted. Theorem opt_simu s1 t s1' oib s2: @@ -338,7 +355,7 @@ Proof. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. * econstructor; eauto. - eapply code_right_assoc; eauto. + eapply code_expand; eauto. all: erewrite preserv_fnparams; eauto. constructor. + (* External function *) @@ -356,7 +373,7 @@ Proof. eexists; left; eexists; split. + eapply exec_return. + econstructor; eauto. - eapply code_right_assoc; eauto. + eapply code_expand; eauto. constructor. Qed. -- cgit From 895470548b89f00111d1f98ae52d217b9fd15643 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 10 May 2021 12:22:34 +0200 Subject: Some more proof cases, comments and cleanup --- scheduling/BTL.v | 69 +------------------------ scheduling/RTLtoBTLproof.v | 125 ++++++++++++++++++++++++++++++++++----------- 2 files changed, 97 insertions(+), 97 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index d7ffa118..a87f674b 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -793,6 +793,7 @@ Definition is_atom (ib: iblock): bool := | _ => true end. +(** Is expand property to only support atomic instructions on the left part of a Bseq *) Inductive is_expand: iblock -> Prop := | exp_Bseq ib1 ib2: is_atom ib1 = true -> @@ -844,71 +845,3 @@ Proof. destruct (cfg!pc); simpl; intros; try_simplify_someHyps. apply expand_None_correct; auto. Qed. - - -(* TODO: eliminer toute la suite quand "right_assoc" sera remplace par "expand" *) - -Definition is_seq (ib: iblock): bool := - match ib with - | Bseq _ _ => true - | _ => false - end. - - -Inductive is_right_assoc: iblock -> Prop := - | ira_Bseq ib1 ib2: - is_seq ib1 = false -> - is_right_assoc ib1 -> - is_right_assoc ib2 -> - is_right_assoc (Bseq ib1 ib2) - | ira_Bcond cond args ib1 ib2 i: - is_right_assoc ib1 -> - is_right_assoc ib2 -> - is_right_assoc (Bcond cond args ib1 ib2 i) - | ira_others ib: - is_atom ib = true -> - is_right_assoc ib - . -Local Hint Constructors is_right_assoc: core. - -Fixpoint right_assoc (ib: iblock): iblock := - match ib with - | Bseq ib1 ib2 => rec_rotate ib1 (right_assoc ib2) - | Bcond cond args ib1 ib2 i => - Bcond cond args (right_assoc ib1) (right_assoc ib2) i - | ib => ib - end -with rec_rotate (ib: iblock) rem: iblock := - match ib with - | Bseq ib1 ib2 => - rec_rotate ib1 (rec_rotate ib2 rem) - | Bcond cond args ib1 ib2 i => - Bseq (Bcond cond args (right_assoc ib1) (right_assoc ib2) i) rem - | ib => Bseq ib rem -end. - -Lemma right_assoc_rec_rotate_correct ib: - is_right_assoc (right_assoc ib) - /\ (forall rem, is_right_assoc rem -> is_right_assoc (rec_rotate ib rem)). -Proof. - induction ib; simpl; intuition. -Qed. - -Lemma right_assoc_correct ib: - is_right_assoc (right_assoc ib). -Proof. - destruct (right_assoc_rec_rotate_correct ib); auto. -Qed. - -Definition right_assoc_code (cfg: code): code := - (PTree.map (fun _ ib => {| entry:=right_assoc ib.(entry); input_regs := ib.(input_regs) |}) cfg). - -Lemma right_assoc_code_correct cfg pc ib : - (right_assoc_code cfg)!pc=Some ib -> is_right_assoc (ib.(entry)). -Proof. - unfold right_assoc_code. - rewrite PTree.gmap. - destruct (cfg!pc); simpl; intros; try_simplify_someHyps. - apply right_assoc_correct; auto. -Qed. - diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index ac9cb8d8..1fa13242 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -91,7 +91,7 @@ The simulation diagram for match_states_intro is as follows: << RTL state match_states_intro BTL state - [pcR0,rs0,m0 ---------------------------- [pcB0,rs0,m0] + [pcR0,rs0,m0] ---------------------------- [pcB0,rs0,m0] | | | | RTL_RUN | *E0 | BTL_RUN @@ -200,28 +200,6 @@ Proof. rewrite GFS. eassumption. assumption. Qed. -(* Inspired from Duplicateproof.v *) -Lemma list_nth_z_dupmap: - forall dupmap ln ln' (pc pc': node) val, - list_nth_z ln val = Some pc -> - list_forall2 (fun n n' => dupmap!n = Some n') ln ln' -> - exists (pc': node), - list_nth_z ln' val = Some pc' - /\ dupmap!pc = Some pc'. -Proof. - induction ln; intros until val; intros LNZ LFA. - - inv LNZ. - - inv LNZ. destruct (zeq val 0) eqn:ZEQ. - + inv H0. destruct ln'; inv LFA. - simpl. exists n. split; auto. - + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto. -Qed. - -(* TODO: definir une measure sur les iblocks. -Cette mesure décroit strictement quand on exécute un "petit-pas" de iblock. -Par exemple, le nombre de noeuds (ou d'instructions "RTL") dans le iblock devrait convenir. -La hauteur de l'arbre aussi. -*) (** Representing an intermediate BTL state We keep a measure of code that remains to be executed with the omeasure @@ -243,7 +221,7 @@ Definition omeasure (oib: option iblock): nat := | Some ib => measure ib end. -Lemma measure_pos: forall ib, +Remark measure_pos: forall ib, measure ib > 0. Proof. induction ib; simpl; auto; lia. @@ -268,13 +246,73 @@ Lemma opt_simu_intro Proof. inv MIB. - (* mib_BF *) - admit. + inv H0; + inversion STEP; subst; try_simplify_someHyps; intros. + + (* Breturn *) + eexists; left; eexists; split. + * econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + erewrite preserv_fnstacksize; eauto. + * econstructor; eauto. + + (* Bcall *) + rename H10 into FIND. + eapply find_function_preserved in FIND. + destruct FIND as (fd' & FF & TRANSFUN). + eexists; left; eexists; split. + * econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + eapply function_sig_translated; eauto. + * repeat (econstructor; eauto). + eapply transf_fundef_correct; eauto. + + (* Btailcall *) + rename H9 into FIND. + eapply find_function_preserved in FIND. + destruct FIND as (fd' & FF & TRANSFUN). + eexists; left; eexists; split. + * econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + eapply function_sig_translated; eauto. + erewrite preserv_fnstacksize; eauto. + * repeat (econstructor; eauto). + eapply transf_fundef_correct; eauto. + + (* Bbuiltin *) + eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. + remember H1 as ODUPLIC; clear HeqODUPLIC. + apply DMC in H1 as [ib [FNC MI]]; clear DMC. + eexists; left; eexists; split. + * econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + pose symbols_preserved as SYMPRES. + eapply eval_builtin_args_preserved; eauto. + eapply external_call_symbols_preserved; eauto. eapply senv_preserved. + * econstructor; eauto. eapply code_expand; eauto. + apply star_refl. + + (* Bjumptable *) + admit. + (* eexists; left; eexists; split. + * econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + exploit (list_nth_z_dupmap dupmap ln' ln); eauto. + * econstructor; eauto.*) - (* mib_exit *) - (* - simpl in *. + (* exists None; right; repeat (split; auto). + inversion STEP; subst; try_simplify_someHyps. + eexists; left. eexists; split. + econstructor; eauto. econstructor. + eexists; eexists; split. eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor. econstructor; eauto. inv BTL_RUN. - eexists; left. eexists; split. + eapply exec_iblock; eauto. econstructor; repeat eexists. inv STEP. @@ -282,7 +320,7 @@ Proof. - (* mib_seq *) inv H. + (* Bnop *) - inversion STEP; subst; try_simplify_someHyps. (* TODO maybe move that for each subcase *) + inversion STEP; subst; try_simplify_someHyps. exists (Some b2); right; repeat (split; auto). econstructor; eauto. inv IS_EXPD; auto; discriminate. eapply star_right; eauto. @@ -331,6 +369,35 @@ Proof. all: eapply star_right; eauto. Admitted. +(** * Main RTL to BTL simulation theorem + +Two possible executions: + +<< + + **Last instruction:** + + RTL state match_states BTL state + s1 ------------------------------------ s2 + | | + STEP | Classical lockstep simu | + | | + s1' ----------------------------------- s2' + + + **Middle instruction:** + + RTL state match_states [oib] BTL state + s1 ------------------------------------ s2 + | _______/ + STEP | *E0 ___________________/ + | / match_states [oib'] + s1' ______/ + Where omeasure oib' < omeasure oib + +>> +*) + Theorem opt_simu s1 t s1' oib s2: RTL.step ge s1 t s1' -> match_states oib s1 s2 -> -- cgit From e3613e0614ccc93c1013f7c39b58cffb6c21a76c Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 10 May 2021 17:49:29 +0200 Subject: new strong_state predicate and lemma --- scheduling/BTL.v | 5 +++ scheduling/RTLtoBTLproof.v | 110 +++++++++++++++++++++++++++++---------------- 2 files changed, 77 insertions(+), 38 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index a87f674b..bc0f0815 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -786,6 +786,11 @@ Definition verify_function dupmap f f' : res unit := do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); verify_cfg dupmap (fn_code f) (RTL.fn_code f'). +Definition is_goto (ib: iblock): bool := + match ib with + | Bgoto _ => true + | _ => false + end. Definition is_atom (ib: iblock): bool := match ib with diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 1fa13242..eb4734b8 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -102,8 +102,8 @@ The simulation diagram for match_states_intro is as follows: >> *) -Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := - | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) +Inductive match_strong_state: (option iblock) -> RTL.state -> state -> Prop := + | match_strong_state_intro ib dupmap st f sp rs m st' f' pcB0 pcR0 pcR1 rs0 m0 isfst ib0 (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') @@ -113,6 +113,14 @@ Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := (IS_EXPD: is_expand ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) + : match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) + . + +Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := + | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) + ib st f sp rs m st' f' pcB0 pcR1 rs0 m0 + (MSTRONG: match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0)) + (NGOTO: is_goto ib = false) : match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) | match_states_call st st' f f' args m @@ -123,7 +131,24 @@ Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := st st' v m (STACKS: list_forall2 match_stackframes st st') : match_states None (RTL.Returnstate st v m) (Returnstate st' v m) - . + . + +Lemma match_strong_state_equiv sp st st' ib f f' rs m rs0 m0 pcB0 pcR1: + match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) -> + match is_goto ib with + | true => + exists ib' succ, + iblock_istep tge sp rs0 m0 (entry ib') rs m (Some (Bgoto succ)) + | false => match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) + end. +Proof. + destruct (is_goto ib) eqn:EQ. + - intros. destruct ib; try destruct fi; try discriminate. + inv H. simpl in BTL_RUN. + rewrite <- iblock_istep_run_equiv in BTL_RUN. + repeat eexists; eauto. + - intros. econstructor; eauto. +Qed. Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. @@ -227,24 +252,34 @@ Proof. induction ib; simpl; auto; lia. Qed. +Lemma list_nth_z_rev_dupmap: + forall dupmap ln ln' (pc pc': node) val, + list_nth_z ln val = Some pc -> + list_forall2 (fun n' n => dupmap!n' = Some n) ln' ln -> + exists (pc': node), + list_nth_z ln' val = Some pc' + /\ dupmap!pc' = Some pc. +Proof. + induction ln; intros until val; intros LNZ LFA. + - inv LNZ. + - inv LNZ. destruct (zeq val 0) eqn:ZEQ. + + inv H0. destruct ln'; inv LFA. + simpl. exists p. split; auto. + + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto. + intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. +Qed. + Lemma opt_simu_intro - sp f f' st st' dupmap isfst pcR0 pcR1 pcB0 ib ib0 rs m rs0 m0 t s1' + sp f f' st st' pcR1 pcB0 ib rs m rs0 m0 t s1' (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t s1') - (STACKS : list_forall2 match_stackframes st st') - (TRANSF : match_function dupmap f f') - (ATpc0 : (fn_code f') ! pcB0 = Some ib0) - (DUPLIC : dupmap ! pcB0 = Some pcR0) - (MIB : match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) - (IS_EXPD : is_expand ib) - (RTL_RUN : star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 - (RTL.State st f sp pcR1 rs m)) - (BTL_RUN : iblock_istep_run tge sp (entry ib0) rs0 m0 = - iblock_istep_run tge sp ib rs m) + (MSTRONG : match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) + (State st' f' sp pcB0 rs0 m0)) + (NGOTO : is_goto ib = false) : exists (oib' : option iblock), (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). Proof. - inv MIB. + inv MSTRONG. inv MIB. - (* mib_BF *) inv H0; inversion STEP; subst; try_simplify_someHyps; intros. @@ -293,32 +328,31 @@ Proof. pose symbols_preserved as SYMPRES. eapply eval_builtin_args_preserved; eauto. eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - * econstructor; eauto. eapply code_expand; eauto. - apply star_refl. + * econstructor; eauto. + { econstructor; eauto. eapply code_expand; eauto. + apply star_refl. } + inversion MI; subst; try_simplify_someHyps. + inv H3; try_simplify_someHyps. + (* Bjumptable *) - admit. - (* eexists; left; eexists; split. + exploit list_nth_z_rev_dupmap; eauto. + intros (pc'0 & LNZ & DM). + eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. + remember DM as ODUPLIC; clear HeqODUPLIC. + apply DMC in DM as [ib [FNC MI]]; clear DMC. + eexists; left; eexists; split. * econstructor; eauto. econstructor. eexists; eexists; split. eapply iblock_istep_run_equiv in BTL_RUN. eapply BTL_RUN. econstructor; eauto. - exploit (list_nth_z_dupmap dupmap ln' ln); eauto. - * econstructor; eauto.*) + * econstructor; eauto. + { econstructor; eauto. eapply code_expand; eauto. + apply star_refl. } + inversion MI; subst; try_simplify_someHyps. + inv H4; try_simplify_someHyps. - (* mib_exit *) - (* exists None; right; repeat (split; auto). - inversion STEP; subst; try_simplify_someHyps. - eexists; left. eexists; split. - econstructor; eauto. econstructor. - eexists; eexists; split. - eapply iblock_istep_run_equiv in BTL_RUN. - eapply BTL_RUN. econstructor. econstructor; eauto. - inv BTL_RUN. - + eapply exec_iblock; eauto. - econstructor; repeat eexists. - inv STEP. - eapply exec_Bgoto.*) admit. - - (* mib_seq *) - inv H. + discriminate. + - (* mib_seq *) admit. + (* inv H. + (* Bnop *) inversion STEP; subst; try_simplify_someHyps. exists (Some b2); right; repeat (split; auto). @@ -351,8 +385,8 @@ Proof. + (* Absurd case *) inv IS_EXPD. inv H4. inv H. + (* Absurd Bcond (Bcond are not allowed in the left part of a Bseq *) - inv IS_EXPD; discriminate. - - (* mib_cond *) + inv IS_EXPD; discriminate.*) + - (* mib_cond *) admit. (* inversion STEP; subst; try_simplify_someHyps. intros; rewrite H12 in BTL_RUN. destruct b. * (* Ifso *) @@ -366,7 +400,7 @@ Proof. simpl; assert (measure bso > 0) by apply measure_pos; lia. inv H2; econstructor; eauto. 1,3: inv IS_EXPD; auto; discriminate. - all: eapply star_right; eauto. + all: eapply star_right; eauto.*) Admitted. (** * Main RTL to BTL simulation theorem -- cgit From 95b6814a1cff386150a7573cb30e9cb68a18052c Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 10 May 2021 20:07:29 +0200 Subject: prove sexec_exact --- scheduling/BTL.v | 2 +- scheduling/BTL_SEtheory.v | 334 ++++++++++++++++++++++++++++++++++++++++++++-- 2 files changed, 324 insertions(+), 12 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index a87f674b..476e71a7 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -153,7 +153,7 @@ Inductive state : Type := state. (** outcome of a block execution *) -Record outcome := { +Record outcome := out { _rs: regset; _m: mem; _fin: option final diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 0a2c0a04..5f7ecb1b 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -13,6 +13,22 @@ Require Import RTL BTL OptionMonad. Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) + +Ltac depmatch exp := + match exp with + | context f [match ?expr with | _ => _ end] => ltac: (depmatch expr) + | _ => exp + end. + +Ltac autodestruct := + let EQ := fresh "EQ" in + match goal with + | |- context f [match ?expr with | _ => _ end] => + let t := ltac: (depmatch expr) in + destruct t eqn:EQ; generalize EQ; clear EQ; congruence || trivial + end. + + Record iblock_exec_context := Bctx { cge: BTL.genv; cstk: list stackframe; @@ -206,7 +222,7 @@ Proof. eapply seval_builtin_arg_correct; eauto. Qed. -Lemma seval_builtin_arg_complete ctx rs m sreg: forall arg varg, +Lemma seval_builtin_arg_exact ctx rs m sreg: forall arg varg, (forall r, eval_sval ctx (sreg r) = Some rs # r) -> seval_builtin_arg ctx m (builtin_arg_map sreg arg) varg -> eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg. @@ -220,7 +236,7 @@ Proof. eapply IHarg1; eauto. eapply IHarg2; eauto. Qed. -Lemma seval_builtin_args_complete ctx rs m sreg: forall args vargs, +Lemma seval_builtin_args_exact ctx rs m sreg: forall args vargs, (forall r, eval_sval ctx (sreg r) = Some rs # r) -> seval_builtin_args ctx m (map (builtin_arg_map sreg) args) vargs -> eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs. @@ -229,7 +245,7 @@ Proof. - simpl. intros. inv H0. constructor. - intros vargs SEVAL BARG. simpl in BARG. inv BARG. constructor; [| eapply IHargs; eauto]. - eapply seval_builtin_arg_complete; eauto. + eapply seval_builtin_arg_exact; eauto. Qed. Fixpoint seval_builtin_sval ctx bsv := @@ -539,10 +555,19 @@ Definition sem_sistate ctx (st: sistate) (rs: regset) (m: mem): Prop := /\ eval_smem ctx st.(si_smem) = Some m /\ forall (r:reg), eval_sval ctx (st.(si_sreg) r) = Some (rs#r). -Definition abort_sistate ctx (st: sistate): Prop := - ~(st.(si_pre) ctx) - \/ eval_smem ctx st.(si_smem) = None - \/ exists (r: reg), eval_sval ctx (st.(si_sreg) r) = None. +(* Remark that we need to reason modulo "extensionality" wrt Regmap.get about regsets. + And, nothing in their representation as (val * PTree.t val) enforces that + (forall r, rs1#r = rs2#r) -> rs1 = rs2 +*) +Lemma sem_sistate_determ ctx st rs1 m1 rs2 m2: + sem_sistate ctx st rs1 m1 -> + sem_sistate ctx st rs2 m2 -> + (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + intros (_&MEM1®1) (_&MEM2®2). + intuition try congruence. + generalize (REG1 r); rewrite REG2; congruence. +Qed. (** * Symbolic execution of final step *) Definition sexec_final_sfv (i: final) (sis: sistate): sfval := @@ -590,9 +615,9 @@ Proof. Qed. Local Hint Constructors final_step: core. -Local Hint Resolve seval_builtin_args_complete: core. +Local Hint Resolve seval_builtin_args_exact: core. -Lemma sexec_final_svf_complete ctx i sis t rs m s: +Lemma sexec_final_svf_exact ctx i sis t rs m s: sem_sistate ctx sis rs m -> sem_sfval ctx (sexec_final_sfv i sis) rs m t s -> final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. @@ -640,7 +665,7 @@ Lemma set_sreg_correct ctx dst sv sis (rs rs': regset) m: sem_sistate ctx (set_sreg dst sv sis) rs' m. Proof. intros (PRE&MEM®) NEW OLD. - unfold sem_sistate, set_sreg; simpl. + unfold sem_sistate; simpl. intuition. - rewrite REG in *; congruence. - destruct (Pos.eq_dec dst r); simpl; subst; eauto. @@ -658,7 +683,7 @@ Lemma set_smem_correct ctx sm sis rs m m': sem_sistate ctx (set_smem sm sis) rs m'. Proof. intros (PRE&MEM®) NEW. - unfold sem_sistate, set_smem; simpl. + unfold sem_sistate; simpl. intuition. rewrite MEM in *; congruence. Qed. @@ -813,3 +838,290 @@ Proof. destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). eapply sexec_rec_correct; simpl; eauto. Qed. + + +(* TODO: déplacer les trucs sur equiv_stackframe -> regmap_setres_eq dans BTL! *) +Inductive equiv_stackframe: stackframe -> stackframe -> Prop := + | equiv_stackframe_intro res f sp pc rs1 rs2 + (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) + . + +Inductive equiv_state: state -> state -> Prop := + | State_equiv stack f sp pc rs1 m rs2 + (EQUIV: forall r, rs1#r = rs2#r): + equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) + | Call_equiv stk stk' f args m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Callstate stk f args m) (Callstate stk' f args m) + | Return_equiv stk stk' v m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Returnstate stk v m) (Returnstate stk' v m) + . + +Local Hint Constructors equiv_stackframe equiv_state: core. + +Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf. +Proof. + destruct stf; eauto. +Qed. + +Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. +Proof. + Local Hint Resolve equiv_stackframe_refl: core. + induction stk; simpl; constructor; auto. +Qed. + +Lemma equiv_state_refl s: equiv_state s s. +Proof. + Local Hint Resolve equiv_stack_refl: core. + induction s; simpl; constructor; auto. +Qed. + +Lemma equiv_stackframe_trans stf1 stf2 stf3: + equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3. +Proof. + destruct 1; intros EQ; inv EQ; try econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +Lemma equiv_stack_trans stk1 stk2: + list_forall2 equiv_stackframe stk1 stk2 -> + forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> + list_forall2 equiv_stackframe stk1 stk3. +Proof. + Local Hint Resolve equiv_stackframe_trans: core. + induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. +Qed. + +Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3. +Proof. + Local Hint Resolve equiv_stack_trans: core. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +Lemma regmap_setres_eq (rs rs': regset) res vres: + (forall r, rs # r = rs' # r) -> + forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r. +Proof. + intros RSEQ r. destruct res; simpl; try congruence. + destruct (peq x r). + - subst. repeat (rewrite Regmap.gss). reflexivity. + - repeat (rewrite Regmap.gso); auto. +Qed. + +Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: + sem_sfval ctx sfv rs1 m t s -> + (forall r, rs1#r = rs2#r) -> + exists s', equiv_state s s' /\ sem_sfval ctx sfv rs2 m t s'. +Proof. + Local Hint Resolve equiv_stack_refl equiv_state_refl regmap_setres_eq: core. + Local Hint Constructors sem_sfval: core. + destruct 1; eexists; split; econstructor; eauto. + econstructor; eauto. +Qed. + +Local Hint Resolve sexec_final_svf_exact: core. + +Definition abort_sistate ctx (sis: sistate): Prop := + ~(sis.(si_pre) ctx) + \/ eval_smem ctx sis.(si_smem) = None + \/ exists (r: reg), eval_sval ctx (sis.(si_sreg) r) = None. + +Lemma sem_sistate_exclude_abort ctx sis rs m: + sem_sistate ctx sis rs m -> + abort_sistate ctx sis -> + False. +Proof. + intros SIS ABORT. inv SIS. destruct H0 as (H0 & H0'). + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. +Qed. + +Local Hint Resolve sem_sistate_exclude_abort: core. + +Lemma set_sreg_preserves_abort ctx sv dst sis: + abort_sistate ctx sis -> + abort_sistate ctx (set_sreg dst sv sis). +Proof. + unfold abort_sistate; simpl; intros [PRE|[MEM|REG]]; try tauto. + destruct REG as [r REG]. + destruct (Pos.eq_dec dst r) as [TEST|TEST] eqn: HTEST. + - subst; rewrite REG; tauto. + - right. right. eexists; rewrite HTEST. auto. +Qed. + +Lemma sexec_op_preserves_abort ctx op args dest sis: + abort_sistate ctx sis + -> abort_sistate ctx (sexec_op op args dest sis). +Proof. + intros; eapply set_sreg_preserves_abort; eauto. +Qed. + +Lemma sexec_load_preserves_abort ctx chunk addr args dest sis: + abort_sistate ctx sis + -> abort_sistate ctx (sexec_load TRAP chunk addr args dest sis). +Proof. + intros; eapply set_sreg_preserves_abort; eauto. +Qed. + +Lemma set_smem_preserves_abort ctx sm sis: + abort_sistate ctx sis -> + abort_sistate ctx (set_smem sm sis). +Proof. + unfold abort_sistate; simpl; try tauto. +Qed. + +Lemma sexec_store_preserves_abort ctx chunk addr args src sis: + abort_sistate ctx sis + -> abort_sistate ctx (sexec_store chunk addr args src sis). +Proof. + intros; eapply set_smem_preserves_abort; eauto. +Qed. + +Local Hint Resolve sexec_op_preserves_abort sexec_load_preserves_abort + sexec_store_preserves_abort: core. + +Lemma sexec_exclude_abort ctx ib t s1: forall sis k + (SEXEC: sem_sstate ctx t s1 (sexec_rec ib sis k)) + (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) + (ABORT: abort_sistate ctx sis), + False. +Proof. + induction ib; simpl; intros; eauto. + - (* final *) inversion SEXEC; subst; eauto. + - (* load *) destruct trap; eauto. + inversion SEXEC. + - (* seq *) + eapply IHib1; eauto. + simpl. eauto. + - (* cond *) + inversion SEXEC; subst; eauto. clear SEXEC. + destruct b; eauto. +Qed. + +Lemma set_sreg_abort ctx dst sv sis rs m: + sem_sistate ctx sis rs m -> + (eval_sval ctx sv = None) -> + abort_sistate ctx (set_sreg dst sv sis). +Proof. + intros (PRE&MEM®) NEW. + unfold sem_sistate, abort_sistate; simpl. + right; right. + exists dst; destruct (Pos.eq_dec dst dst); simpl; try congruence. +Qed. + +Lemma sexec_op_abort ctx sis op args dest rs m + (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = None) + (SIS: sem_sistate ctx sis rs m) + : abort_sistate ctx (sexec_op op args dest sis). +Proof. + eapply set_sreg_abort; eauto. + simpl. destruct SIS as (PRE&MEM®). + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. +Qed. + +Lemma sexec_load_TRAP_abort ctx chunk addr args dst sis rs m + (EVAL: forall a, eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a -> Mem.loadv chunk m a = None) + (SIS: sem_sistate ctx sis rs m) + : abort_sistate ctx (sexec_load TRAP chunk addr args dst sis). +Proof. + eapply set_sreg_abort; eauto. + simpl. destruct SIS as (PRE&MEM®). + erewrite eval_list_sval_inj; simpl; auto. + intros; autodestruct; try_simplify_someHyps. +Qed. + +Lemma set_smem_abort ctx sm sis rs m: + sem_sistate ctx sis rs m -> + eval_smem ctx sm = None -> + abort_sistate ctx (set_smem sm sis). +Proof. + intros (PRE&MEM®) NEW. + unfold abort_sistate; simpl. + tauto. +Qed. + +Lemma sexec_store_abort ctx chunk addr args src sis rs m + (EVAL: forall a, eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a -> Mem.storev chunk m a (rs # src) = None) + (SIS: sem_sistate ctx sis rs m) + :abort_sistate ctx (sexec_store chunk addr args src sis). +Proof. + eapply set_smem_abort; eauto. + simpl. destruct SIS as (PRE&MEM®). + erewrite eval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + intros; rewrite REG; autodestruct; try_simplify_someHyps. +Qed. + +Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort: core. + +Lemma sexec_rec_exact ctx ib t s1: forall sis k + (SEXEC: sem_sstate ctx t s1 (sexec_rec ib sis k)) + rs m + (SIS: sem_sistate ctx sis rs m) + (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) + , + match iblock_istep_run (cge ctx) (csp ctx) ib rs m with + | Some (out rs' m' (Some fin)) => + exists s2, final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 + | Some (out rs' m' None) => exists sis', (sem_sstate ctx t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m') + | None => False + end. +Proof. + induction ib; simpl; intros; eauto. + - (* final *) + inv SEXEC. + exploit (sem_sistate_determ ctx sis rs m rs0 m0); eauto. + intros (REG&MEM); subst. + exploit (sem_sfval_equiv rs0 rs); eauto. + intros (s2 & EQUIV & SFV'). + eexists; split; eauto. + - (* Bop *) autodestruct; eauto. + - destruct trap; [| inv SEXEC ]. + repeat autodestruct; eauto. + all: intros; eapply CONT; eauto; + eapply sexec_load_TRAP_abort; eauto; + intros; try_simplify_someHyps. + - repeat autodestruct; eauto. + all: intros; eapply CONT; eauto; + eapply sexec_store_abort; eauto; + intros; try_simplify_someHyps. + - (* Bseq *) + exploit IHib1; eauto. clear sis SEXEC SIS. + { simpl; intros; eapply sexec_exclude_abort; eauto. } + destruct (iblock_istep_run _ _ _ _ _) eqn: ISTEP; auto. + destruct o. + destruct _fin eqn: OFIN; simpl; eauto. + intros (sis1 & SEXEC1 & SIS1). + exploit IHib2; eauto. + - (* Bcond *) + inv SEXEC. + erewrite seval_condition_eq in SEVAL; eauto. + rewrite SEVAL. + destruct b. + + exploit IHib1; eauto. + + exploit IHib2; eauto. +Qed. + + +(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution + (sexec is exact). +*) +Theorem sexec_exact ctx ib t s1: + sem_sstate ctx t s1 (sexec ib) -> + exists s2, iblock_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 + /\ equiv_state s1 s2. +Proof. + intros; exploit sexec_rec_exact; eauto. + { intros sis' SEXEC; inversion SEXEC. } + repeat autodestruct; simpl; try tauto. + - intros D1 D2 ISTEP (s2 & FSTEP & EQSTEP); subst. + eexists; split; eauto. + repeat eexists; eauto. + erewrite iblock_istep_run_equiv; eauto. + - intros D1 D2 ISTEP (sis & SEXEC & _); subst. + inversion SEXEC. +Qed. + -- cgit From 473c17f114f44756c3803c026266417dfe92c242 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 11 May 2021 11:00:47 +0200 Subject: pointeur Justus -> roadmap --- scheduling/BTLroadmap.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 05dcb95e..2f630a72 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -307,6 +307,6 @@ This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. ## Alias analysis in the symbolic simulation -... +A REGARDER [papier pointé par Justus](https://vbpf.github.io/assets/prevail-paper.pdf) -- cgit From 35487b45f778b15118d3cc934622b35429a4c899 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 11 May 2021 16:02:08 +0200 Subject: better autodestruct ? --- scheduling/BTL_SEtheory.v | 16 ---------------- 1 file changed, 16 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 5f7ecb1b..b9a05a8a 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -13,22 +13,6 @@ Require Import RTL BTL OptionMonad. Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) - -Ltac depmatch exp := - match exp with - | context f [match ?expr with | _ => _ end] => ltac: (depmatch expr) - | _ => exp - end. - -Ltac autodestruct := - let EQ := fresh "EQ" in - match goal with - | |- context f [match ?expr with | _ => _ end] => - let t := ltac: (depmatch expr) in - destruct t eqn:EQ; generalize EQ; clear EQ; congruence || trivial - end. - - Record iblock_exec_context := Bctx { cge: BTL.genv; cstk: list stackframe; -- cgit From 9ef1955e96a9bc16395dfe38212107915923260b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 11 May 2021 18:14:13 +0200 Subject: new lemma definition, some preparations in existing proofs --- scheduling/RTLtoBTLproof.v | 117 +++++++++++++++++++++++++-------------------- 1 file changed, 65 insertions(+), 52 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index eb4734b8..d9fec468 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -102,9 +102,8 @@ The simulation diagram for match_states_intro is as follows: >> *) -Inductive match_strong_state: (option iblock) -> RTL.state -> state -> Prop := +Inductive match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst: Prop := | match_strong_state_intro - ib dupmap st f sp rs m st' f' pcB0 pcR0 pcR1 rs0 m0 isfst ib0 (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') (ATpc0: (fn_code f')!pcB0 = Some ib0) @@ -113,13 +112,13 @@ Inductive match_strong_state: (option iblock) -> RTL.state -> state -> Prop := (IS_EXPD: is_expand ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) - : match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) + : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst . Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) - ib st f sp rs m st' f' pcB0 pcR1 rs0 m0 - (MSTRONG: match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0)) + dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst + (MSTRONG: match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) (NGOTO: is_goto ib = false) : match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) | match_states_call @@ -133,23 +132,6 @@ Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := : match_states None (RTL.Returnstate st v m) (Returnstate st' v m) . -Lemma match_strong_state_equiv sp st st' ib f f' rs m rs0 m0 pcB0 pcR1: - match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) -> - match is_goto ib with - | true => - exists ib' succ, - iblock_istep tge sp rs0 m0 (entry ib') rs m (Some (Bgoto succ)) - | false => match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) - end. -Proof. - destruct (is_goto ib) eqn:EQ. - - intros. destruct ib; try destruct fi; try discriminate. - inv H. simpl in BTL_RUN. - rewrite <- iblock_istep_run_equiv in BTL_RUN. - repeat eexists; eauto. - - intros. econstructor; eauto. -Qed. - Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. rewrite <- (Genv.find_symbol_match TRANSL). reflexivity. @@ -252,6 +234,15 @@ Proof. induction ib; simpl; auto; lia. Qed. +Lemma entry_isnt_goto dupmap f pc ib: + match_iblock dupmap (RTL.fn_code f) true pc (entry ib) None -> + is_goto (entry ib) = false. +Proof. + intros. + destruct (entry ib); trivial. + destruct fi; trivial. inv H. inv H4. +Qed. + Lemma list_nth_z_rev_dupmap: forall dupmap ln ln' (pc pc': node) val, list_nth_z ln val = Some pc -> @@ -269,17 +260,29 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. +Lemma match_strong_state_simu + dupmap st st' f f' sp rs m rs0 m0 rs1 m1 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 t + (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t (RTL.State st f sp pcR2 rs1 m1)) + (MSS1 : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst) + (MSS2 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) + (*(MES : measure ib2 < measure ib1)*) + : exists (oib' : option iblock), + (exists s2', step tge (State st' f' sp pcB0 rs0 m0) E0 s2' + /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) s2') + \/ (omeasure oib' < S (measure ib2) /\ t=E0 + /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) (State st' f' sp pcB0 rs0 m0)). +Proof. Admitted. + Lemma opt_simu_intro - sp f f' st st' pcR1 pcB0 ib rs m rs0 m0 t s1' + dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst s1' t (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t s1') - (MSTRONG : match_strong_state (Some ib) (RTL.State st f sp pcR1 rs m) - (State st' f' sp pcB0 rs0 m0)) + (MSTRONG : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) (NGOTO : is_goto ib = false) : exists (oib' : option iblock), (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). Proof. - inv MSTRONG. inv MIB. + inversion MSTRONG; subst. inv MIB. - (* mib_BF *) inv H0; inversion STEP; subst; try_simplify_someHyps; intros. @@ -351,45 +354,54 @@ Proof. inv H4; try_simplify_someHyps. - (* mib_exit *) discriminate. - - (* mib_seq *) admit. - (* inv H. + - (* mib_seq *) + inversion H; subst. + (* Bnop *) - inversion STEP; subst; try_simplify_someHyps. - exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv IS_EXPD; auto; discriminate. + inversion STEP; subst; try_simplify_someHyps; intros. + eapply match_strong_state_simu. + 1,2: do 2 (econstructor; eauto). + econstructor; eauto. (* TODO factorize *) + inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. + (* Bop *) - inversion STEP; subst; try_simplify_someHyps. - exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv IS_EXPD; auto; discriminate. + inversion STEP; subst; try_simplify_someHyps; intros. + eapply match_strong_state_simu. + 1,2: do 2 (econstructor; eauto). + econstructor; eauto. (* TODO factorize *) + inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. - erewrite eval_operation_preserved in H10. - erewrite H10 in BTL_RUN; simpl in BTL_RUN; auto. + erewrite eval_operation_preserved in H11. + erewrite H11 in BTL_RUN; simpl in BTL_RUN; auto. intros; rewrite <- symbols_preserved; trivial. + (* Bload *) - inversion STEP; subst; try_simplify_someHyps. - exists (Some (b2)); right; repeat (split; auto). - econstructor; eauto. inv IS_EXPD; auto; discriminate. + inversion STEP; subst; try_simplify_someHyps; intros. + eapply match_strong_state_simu. + 1,2: do 2 (econstructor; eauto). + econstructor; eauto. (* TODO factorize *) + inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. - erewrite eval_addressing_preserved in H10. - erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. + erewrite eval_addressing_preserved in H11. + erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. intros; rewrite <- symbols_preserved; trivial. + (* Bstore *) - inversion STEP; subst; try_simplify_someHyps. - exists (Some b2); right; repeat (split; auto). - econstructor; eauto. inv IS_EXPD; auto; discriminate. + inversion STEP; subst; try_simplify_someHyps; intros. + eapply match_strong_state_simu. + 1,2: do 2 (econstructor; eauto). + econstructor; eauto. (* TODO factorize *) + inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. - erewrite eval_addressing_preserved in H10. - erewrite H10, H11 in BTL_RUN; simpl in BTL_RUN; auto. + erewrite eval_addressing_preserved in H11. + erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. intros; rewrite <- symbols_preserved; trivial. + (* Absurd case *) - inv IS_EXPD. inv H4. inv H. + inv IS_EXPD; try inv H5; discriminate. + (* Absurd Bcond (Bcond are not allowed in the left part of a Bseq *) - inv IS_EXPD; discriminate.*) + inv IS_EXPD; discriminate. - (* mib_cond *) admit. (* - inversion STEP; subst; try_simplify_someHyps. + inversion STEP; subst; try_simplify_someHyps; intros. intros; rewrite H12 in BTL_RUN. destruct b. * (* Ifso *) + eapply match_strong_state_simu. exists (Some bso); right; repeat (split; eauto). simpl; assert (measure bnot > 0) by apply measure_pos; lia. inv H2; econstructor; eauto. @@ -455,8 +467,9 @@ Proof. eexists; left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. - * econstructor; eauto. + * econstructor. econstructor; eauto. eapply code_expand; eauto. + 3: eapply entry_isnt_goto; eauto. all: erewrite preserv_fnparams; eauto. constructor. + (* External function *) @@ -473,9 +486,9 @@ Proof. apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. eexists; left; eexists; split. + eapply exec_return. - + econstructor; eauto. + + econstructor. econstructor; eauto. eapply code_expand; eauto. - constructor. + constructor. eapply entry_isnt_goto; eauto. Qed. Local Hint Resolve plus_one star_refl: core. -- cgit From fc6b1bead179f0ec628c4af8988c1bda8adebbdf Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 11 May 2021 23:13:21 +0200 Subject: finish proof --- scheduling/RTLtoBTLproof.v | 84 +++++++++++++++++++++++++--------------------- 1 file changed, 45 insertions(+), 39 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index d9fec468..7d039f11 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -116,7 +116,7 @@ Inductive match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib . Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := - | match_states_intro (* TODO: LES DETAILS SONT SANS DOUTE A REVOIR !!! *) + | match_states_intro dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst (MSTRONG: match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) (NGOTO: is_goto ib = false) @@ -261,17 +261,36 @@ Proof. Qed. Lemma match_strong_state_simu - dupmap st st' f f' sp rs m rs0 m0 rs1 m1 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 t - (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t (RTL.State st f sp pcR2 rs1 m1)) + dupmap st st' f f' sp rs m rs0 m0 rs1 m1 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n + (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) E0 (RTL.State st f sp pcR2 rs1 m1)) (MSS1 : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst) (MSS2 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) - (*(MES : measure ib2 < measure ib1)*) + (MES : measure ib2 < n) : exists (oib' : option iblock), (exists s2', step tge (State st' f' sp pcB0 rs0 m0) E0 s2' /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) s2') - \/ (omeasure oib' < S (measure ib2) /\ t=E0 + \/ (omeasure oib' < n /\ E0=E0 /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) (State st' f' sp pcB0 rs0 m0)). -Proof. Admitted. +Proof. + destruct (is_goto ib2) eqn:GT. + destruct ib2; try destruct fi; try discriminate. + - (* Bgoto *) + inv MSS2. inversion MIB; subst; try inv H3. + remember H0 as ODUPLIC; clear HeqODUPLIC. + eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. + apply DMC in H0 as [ib [FNC MI]]; clear DMC. + eexists; left; eexists; split. + + repeat econstructor; eauto. + apply iblock_istep_run_equiv in BTL_RUN; eauto. + + econstructor. econstructor; eauto. + eapply code_expand; eauto. + econstructor. + eapply entry_isnt_goto; eauto. + - (* Others *) + exists (Some ib2); right; split. + simpl; auto. + split; auto. econstructor; eauto. +Qed. Lemma opt_simu_intro dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst s1' t @@ -355,65 +374,52 @@ Proof. - (* mib_exit *) discriminate. - (* mib_seq *) - inversion H; subst. + inversion H; subst; + try (inv IS_EXPD; try inv H5; discriminate; fail); + inversion STEP; subst; try_simplify_someHyps; intros. + (* Bnop *) - inversion STEP; subst; try_simplify_someHyps; intros. eapply match_strong_state_simu. 1,2: do 2 (econstructor; eauto). - econstructor; eauto. (* TODO factorize *) + econstructor; eauto. inv IS_EXPD; eauto. simpl in *; discriminate. - eapply star_right; eauto. + eapply star_right; eauto. lia. + (* Bop *) - inversion STEP; subst; try_simplify_someHyps; intros. eapply match_strong_state_simu. 1,2: do 2 (econstructor; eauto). - econstructor; eauto. (* TODO factorize *) + econstructor; eauto. inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. erewrite eval_operation_preserved in H11. erewrite H11 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. + intros; rewrite <- symbols_preserved; trivial. lia. + (* Bload *) - inversion STEP; subst; try_simplify_someHyps; intros. eapply match_strong_state_simu. 1,2: do 2 (econstructor; eauto). - econstructor; eauto. (* TODO factorize *) + econstructor; eauto. inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. erewrite eval_addressing_preserved in H11. erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. + intros; rewrite <- symbols_preserved; trivial. lia. + (* Bstore *) - inversion STEP; subst; try_simplify_someHyps; intros. eapply match_strong_state_simu. 1,2: do 2 (econstructor; eauto). - econstructor; eauto. (* TODO factorize *) + econstructor; eauto. inv IS_EXPD; eauto. simpl in *; discriminate. eapply star_right; eauto. erewrite eval_addressing_preserved in H11. erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. - + (* Absurd case *) - inv IS_EXPD; try inv H5; discriminate. - + (* Absurd Bcond (Bcond are not allowed in the left part of a Bseq *) - inv IS_EXPD; discriminate. - - (* mib_cond *) admit. (* + intros; rewrite <- symbols_preserved; trivial. lia. + - (* mib_cond *) inversion STEP; subst; try_simplify_someHyps; intros. - intros; rewrite H12 in BTL_RUN. destruct b. - * (* Ifso *) - eapply match_strong_state_simu. - exists (Some bso); right; repeat (split; eauto). - simpl; assert (measure bnot > 0) by apply measure_pos; lia. - inv H2; econstructor; eauto. - 1,3: inv IS_EXPD; auto; discriminate. - all: eapply star_right; eauto. - * (* Ifnot *) - exists (Some bnot); right; repeat (split; eauto). - simpl; assert (measure bso > 0) by apply measure_pos; lia. - inv H2; econstructor; eauto. - 1,3: inv IS_EXPD; auto; discriminate. - all: eapply star_right; eauto.*) -Admitted. + intros; rewrite H12 in BTL_RUN. destruct b; + eapply match_strong_state_simu; eauto. + 1,3: inv H2; econstructor; eauto. + 1,3,5,7: inv IS_EXPD; auto; discriminate. + 1-4: eapply star_right; eauto. + assert (measure bnot > 0) by apply measure_pos; lia. + assert (measure bso > 0) by apply measure_pos; lia. +Qed. (** * Main RTL to BTL simulation theorem -- cgit From 8999ea0a952381a4ba94b5266a268ec350bf3f2d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 12 May 2021 08:51:05 +0200 Subject: ascii simu --- scheduling/RTLtoBTLproof.v | 65 ++++++++++++++++++++++++++++++++++------------ 1 file changed, 48 insertions(+), 17 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7d039f11..93f83450 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -91,18 +91,18 @@ The simulation diagram for match_states_intro is as follows: << RTL state match_states_intro BTL state - [pcR0,rs0,m0] ---------------------------- [pcB0,rs0,m0] + [pcR0,rs0,m0] --------------------------- [pcB0,rs0,m0] | | | | RTL_RUN | *E0 | BTL_RUN | | | MIB | - [pcR1,rs,m] -------------------------------- [ib] + [pcR1,rs1,m1] ------------------------------- [ib] >> *) -Inductive match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst: Prop := +Inductive match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst: Prop := | match_strong_state_intro (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') @@ -110,17 +110,17 @@ Inductive match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib (DUPLIC: dupmap!pcB0 = Some pcR0) (MIB: match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) (IS_EXPD: is_expand ib) - (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs m)) - (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs m) - : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst + (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs1 m1)) + (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs1 m1) + : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst . Inductive match_states: (option iblock) -> RTL.state -> state -> Prop := | match_states_intro - dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst - (MSTRONG: match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) + dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst + (MSTRONG: match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) (NGOTO: is_goto ib = false) - : match_states (Some ib) (RTL.State st f sp pcR1 rs m) (State st' f' sp pcB0 rs0 m0) + : match_states (Some ib) (RTL.State st f sp pcR1 rs1 m1) (State st' f' sp pcB0 rs0 m0) | match_states_call st st' f f' args m (STACKS: list_forall2 match_stackframes st st') @@ -260,17 +260,48 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. +(** * Match strong state property + +Used when executing non-atomic instructions such as Bseq/Bcond(ib1,ib2). +Two possible executions: + +<< + + **ib2 is a Bgoto (left side):** + + RTL state MSS1 BTL state + [pcR1,rs1,m1] -------------------------- [ib1,pcB0,rs0,m0] + | | + | | + | | BTL_STEP + | | + | | + RTL_STEP | *E0 [ib2,pc=(Bgoto succ),rs2,m2] + | / | + | MSS2 / | + | _________________/ | BTL_GOTO + | / | + | / GOAL: match_states | + [pcR2,rs2,m2] ------------------------ [ib?,pc=succ,rs2,m2] + + + **ib2 is any other instruction (right side):** + +See explanations of opt_simu below. + +>> +*) Lemma match_strong_state_simu - dupmap st st' f f' sp rs m rs0 m0 rs1 m1 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n - (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) E0 (RTL.State st f sp pcR2 rs1 m1)) - (MSS1 : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst) - (MSS2 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) + dupmap st st' f f' sp rs2 m2 rs1 m1 rs0 m0 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n + (STEP : RTL.step ge (RTL.State st f sp pcR1 rs1 m1) E0 (RTL.State st f sp pcR2 rs2 m2)) + (MSS1 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst) + (MSS2 : match_strong_state dupmap st st' f f' sp rs2 m2 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) (MES : measure ib2 < n) : exists (oib' : option iblock), (exists s2', step tge (State st' f' sp pcB0 rs0 m0) E0 s2' - /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) s2') + /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) s2') \/ (omeasure oib' < n /\ E0=E0 - /\ match_states oib' (RTL.State st f sp pcR2 rs1 m1) (State st' f' sp pcB0 rs0 m0)). + /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) (State st' f' sp pcB0 rs0 m0)). Proof. destruct (is_goto ib2) eqn:GT. destruct ib2; try destruct fi; try discriminate. @@ -427,7 +458,7 @@ Two possible executions: << - **Last instruction:** + **Last instruction (left side):** RTL state match_states BTL state s1 ------------------------------------ s2 @@ -437,7 +468,7 @@ Two possible executions: s1' ----------------------------------- s2' - **Middle instruction:** + **Middle instruction (right side):** RTL state match_states [oib] BTL state s1 ------------------------------------ s2 -- cgit From 840f0a740f3cddf2d6bbb8d89ae686357ed12e00 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 12 May 2021 10:37:54 +0200 Subject: fix roadmap on "Simulation modulo liveness" --- scheduling/BTLroadmap.md | 85 ++++++++++++++++++++++++++++++++++-------------- 1 file changed, 61 insertions(+), 24 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 2f630a72..dde2090c 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -63,37 +63,74 @@ Branch factorization should also be possible in BTL -> RTL pass. Example: revert Pour faciliter le vérificateur, faire comme le Bgoto: les Bnop en "trop" ne sont autorisés que s'il y a eu une instruction RTL avant. Ajouter aussi un booléen (positionné par l'oracle) sur le Bnop qui indique si le "nop" existe ou pas au niveau RTL. +## Simulation modulo liveness and "Functional" semantics of BTL blocks + +L'approche suivie pour réaliser la simulation modulo liveness dans +RTLpath est compliquée à adapter sur BTL. En effet, un état +symbolique RTLpath correspond à un état symbolique BTL très +particulier: toutes les "feuilles" (les `Sfinal`) sont des `Sgoto` +sauf éventuellement une. Or, dans RTLpath, le traitement de l'info de liveness sur +cette feuille particulière est très adhoc et laborieux (cf. `pre_output_regs`). On +n'a pas envie de généraliser cette usine à gaz. + +On cherche donc une façon plus abstraite de faire... On a l'idée de +coder la "simulation modulo liveness" dans une "simulation +less_undef". Ça peut rendre le cadre du test de simulation plus +simple et plus général. + +L'idée de départ: "Extends BTL with the possibility of destroying a +set of registers at each exit (a destroyed register is simply set to +Vundef)" en prouvant une simulation "less_undef" n'est pas assez +générale! Ça n'autorise pas à introduire de "nouveaux" registres dans +le bloc transformé, juste à donner n'importe quelle valeur aux +registres "hors-liveout". + +**Idée corrigée** à côté de la sémantique "à la RTL" pour BTL, on +définit une nouvelle sémantique (basée sur la même sémantique à grand +pas dans les blocs), où c'est juste "l'entrée dans un bloc" qui change +de sémantique. Intuitivement, cette nouvelle sémantique considère +chaque bloc comme une sorte de fonction paramétrée par les `input_regs` +et appelée uniquement en "tailcall" (via les "goto"). C'est ce qu'on +va appeler la "functional semantics" de BTL (l'autre étant appelée +qqchose comme "CFG semantics" ?). + +Autrement dit, sur l'entrée dans un bloc pour un état (rs,m), on +remplace juste cet état par un nouvel état (rs0,m) où dans rs0, tous les registres +sont initialisés à Vundef sauf ceux du `input_regs` qui sont +initialisés comme une copie de "rs". In fine, on applique le +"iblock_step" sur cet état (rs0,m). + +**NOTE** cette idée de voir les blocs comme des "fonctions" correpond bien à la représentation "SSA" à la Appel/MLIR. +Donc cette sémantique peut servir de base pour un support de formes SSA (partielles ou totales) dans BTL. +Pour l'encodage de SSA, il suffira d'étendre le mécanisme d'initialisation du "rs0" d'un bloc avec un passage de paramètres. + +**IMPLEM PLAN** + +1. définir the "functional semantics" of BTL. +2. prouver la simulation BTL.FunctionalSemantics -> BTL.CFGSemantics (facile normalement: c'est du "less_undef lockstep"). +3. implémenter une passe BTL.CFGSemantics -> BTL.FunctionalSemantics. + en gros: on implémente un vérificateur des infos de liveness. + c'est la correction du "input_regs" qui garantit que la simulation est correct. + La preuve devrait normalement être très similaire à RTLpathLivegenproof. + ## "Basic" symbolic execution / symbolic simulation -IMPLEM PLAN in two steps: +We will implement a "basic" (e.g without rewriting rules) "less_undef" simulation test for BTL.FunctionalSemantics. -1. a very basic version without rewriting rules, and without support for simulation modulo-liveness. +**IMPLEM NOTE** -2. Support a "SVundef" symbolic value, and implement symbolic simulation wrt "less_undef" relation. -This will allow more general rewriting rules and will generalize modulo-liveness (see next section). +- use Continuation Passing Style for an easy recursive generation of "all execution paths". +- pour implementer la "functional semantics", il faut changer + légèrement la sémantique du map qui associe une valeur symbolique à + chaque registre. En RTLpath, un registre "absent" de la map était considéré comme positionné à sa valeur initial. + Là, il faudra considérer qu'il vaut Vundef (il y a une valeur symbolique spéciale à ajouter !) **CURRENT STATUS** -- model of symbolic execution in Continuation-Passing Style, with basic "correctness" thm ("completness" not yet done). - -## Simulation modulo liveness - -Extends BTL with the possibility of destroying a set of registers at each exit (a destroyed register is simply set to Vundef). - -We can the provide a "MarkUnused" pass of BTL -> BTL that insert these destructions at each exit using liveness info ! -(NB: the liveness info could be provided by the RTL -> BTL oracle like on RTLpath). - -In the symbolic execution, destroying a register corresponds to associate it SVundef value. - -NB: this approach through "destroyed registers" and "SVundef" seems more modular (thus simpler) and more general than the current approach on RTLpath. - -The proof of "simulation modulo liveness" of a BTL B1 by a BTL B2 will simply corresponds to compose two forward simulations: B1 --> MarkUnused(B1) --> B2. - -**IMPLEM NOTE:** "MarkUnused" checks the correctness of liveness informations while performing its transformation. -It should computes alive register from block entry, by forward proceeding of "internal instructions" (like in RTLPathLivegen). -When arriving on a final instruction, we marked as "destroyed" all alive registers that are not in the input registers of the exit -Semantically, the destruction of registers happens "in parallel" of the final step of the block. +- model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms. +- next step: high-level specification of "symbolic simulation" + preservation proof w.r.t BTL.FunctionalSemantics +(dans une première étape, on peut laisser le "less_undef" de côté: ce sera facile à ajouter quand le cadre sera en place). ## Port of RTLpath optimizations to BTL @@ -301,7 +338,7 @@ Case-study: support of strength-reduction. ## Support of SSA-optimizations -Extends BTL with "register renamings" at exits (in a sense, "register renaming" extends the feature of "register destroying"). +Extends BTL with "register renamings" at exits. This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. -- cgit From 335a09db280cf71db94c3d54cf55cd0ec518473b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 13 May 2021 14:08:12 +0200 Subject: iblock_istep rec and correct lemmas --- scheduling/RTLtoBTLproof.v | 83 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 83 insertions(+) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 93f83450..cebf867b 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -260,6 +260,88 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. +(* + +TODO Version avec ofin2 quelconque + +(*Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: + forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) + k ofin2 fin2 rs2 m2 + (CONT: match ofin1 with + | None => + exists rem, k = Some rem /\ ofin2=Some fin2 + /\ iblock_istep tge sp rs1 m1 rem rs2 m2 (Some fin2) + | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 + end), + iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 ofin2. +Proof. + induction 1; simpl. + { (* BF *) + destruct k; + intros ? ? ? ? (HRS & HM & HOF); subst. + eapply exec_seq_stop. all: constructor. } + 1-4: (* Bnop, Bop, Bload, Bstore *) + destruct k; intros ? ? ? ? (rem & EQR & EQO & ISTEP); + inversion EQR; subst; clear EQR; + eapply exec_seq_continue; [ econstructor; eauto | assumption]. + - (* Bseq_stop *) + destruct k; intros; apply IHISTEP; eauto. + - (* Bseq_continue *) + destruct ofin; intros. + destruct CONT as [HRS [HM HOF]]; subst. + 2: destruct CONT as [rem [HRS [HM HOF]]]; subst. + all: eapply IHISTEP1; eexists; repeat split; eauto. + - (* Bcond *) + destruct ofin; intros; + econstructor; eauto; + destruct b; eapply IHISTEP; eauto. + Qed.*) + *) + +Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin: + forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin) + k fin2 rs2 m2 + (CONT: match ofin with + | None => + exists rem, k = Some rem + /\ iblock_istep tge sp rs1 m1 rem rs2 m2 (Some fin2) + | Some fin => rs2=rs1 /\ m2=m1 /\ fin2=fin + end), + iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 (Some fin2). +Proof. + induction 1; simpl. + { (* BF *) + destruct k; + intros ? ? ? (HRS & HM & HOF); subst. + eapply exec_seq_stop. all: constructor. } + 1-4: (* Bnop, Bop, Bload, Bstore *) + destruct k; intros ? ? ? (rem & EQR & ISTEP); + inversion EQR; subst; clear EQR; + eapply exec_seq_continue; [ econstructor; eauto | assumption]. + - (* Bseq_stop *) + destruct k; intros; apply IHISTEP; eauto. + - (* Bseq_continue *) + destruct ofin; intros; apply IHISTEP1; + eexists; split; eauto. + - (* Bcond *) + destruct ofin; intros; + econstructor; eauto; + destruct b; eapply IHISTEP; auto. +Qed. + +Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 fin: + iblock_istep tge sp rs0 m0 ib rs1 m1 (Some fin) -> + iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 (Some fin). +Proof. + intros; eapply expand_iblock_istep_rec_correct; eauto; simpl; auto. +Qed. + +(* TODO *) +Lemma expand_iblock_istep_exact sp rs0 m0 rs1 m1: + forall ib fin (ISTEP: iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 (Some fin)), + iblock_istep tge sp rs0 m0 ib rs1 m1 (Some fin). +Admitted. + (** * Match strong state property Used when executing non-atomic instructions such as Bseq/Bcond(ib1,ib2). @@ -291,6 +373,7 @@ See explanations of opt_simu below. >> *) + Lemma match_strong_state_simu dupmap st st' f f' sp rs2 m2 rs1 m1 rs0 m0 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n (STEP : RTL.step ge (RTL.State st f sp pcR1 rs1 m1) E0 (RTL.State st f sp pcR2 rs2 m2)) -- cgit From 07b78e13a57c63d9ed99d3832bf5acdb69d6c6e2 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sun, 16 May 2021 16:48:56 +0200 Subject: some advance --- scheduling/RTLtoBTLproof.v | 215 ++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 191 insertions(+), 24 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index cebf867b..4d785c10 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -260,17 +260,17 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. -(* -TODO Version avec ofin2 quelconque +(*TODO Is this lemma too hard to read? *) -(*Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: +Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) - k ofin2 fin2 rs2 m2 + k ofin2 rs2 m2 (CONT: match ofin1 with | None => - exists rem, k = Some rem /\ ofin2=Some fin2 - /\ iblock_istep tge sp rs1 m1 rem rs2 m2 (Some fin2) + (k = None /\ rs2=rs1 /\ m2=m1 /\ ofin2 = None) + \/ (exists rem, k = Some rem + /\ iblock_istep tge sp rs1 m1 rem rs2 m2 ofin2) | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 end), iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 ofin2. @@ -278,26 +278,31 @@ Proof. induction 1; simpl. { (* BF *) destruct k; - intros ? ? ? ? (HRS & HM & HOF); subst. + intros ? ? ? (HRS & HM & HOF); subst. eapply exec_seq_stop. all: constructor. } + (*destruct k; intros. try inv CONT.*) 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct k; intros ? ? ? ? (rem & EQR & EQO & ISTEP); - inversion EQR; subst; clear EQR; + destruct k; intros; destruct CONT as [[HK [HRS [HM HO]]]|[rem [HR ISTEP]]]; + subst; try (inv HK; fail); try (inv HR; fail); try (econstructor; eauto; fail); + inversion HR; subst; clear HR; eapply exec_seq_continue; [ econstructor; eauto | assumption]. - (* Bseq_stop *) destruct k; intros; apply IHISTEP; eauto. - (* Bseq_continue *) destruct ofin; intros. - destruct CONT as [HRS [HM HOF]]; subst. - 2: destruct CONT as [rem [HRS [HM HOF]]]; subst. - all: eapply IHISTEP1; eexists; repeat split; eauto. + + destruct CONT as [HRS [HM HOF]]; subst. + eapply IHISTEP1; right. eexists; repeat split; eauto. + + destruct CONT as [[HK [HRS [HM HO]]]|[rem [HR ISTEP]]]; subst. + * eapply IHISTEP1; right. eexists; repeat split; eauto. + eapply IHISTEP2; left; simpl; auto. + * eapply IHISTEP1; right. eexists; repeat split; eauto. - (* Bcond *) destruct ofin; intros; econstructor; eauto; destruct b; eapply IHISTEP; eauto. - Qed.*) - *) +Qed. +(* TODO old version Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin: forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin) k fin2 rs2 m2 @@ -327,20 +332,182 @@ Proof. destruct ofin; intros; econstructor; eauto; destruct b; eapply IHISTEP; auto. -Qed. + Qed.*) -Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 fin: - iblock_istep tge sp rs0 m0 ib rs1 m1 (Some fin) -> - iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 (Some fin). +Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: + iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> + iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 ofin. Proof. - intros; eapply expand_iblock_istep_rec_correct; eauto; simpl; auto. + intros; eapply expand_iblock_istep_rec_correct; eauto. + destruct ofin; simpl; auto. Qed. -(* TODO *) -Lemma expand_iblock_istep_exact sp rs0 m0 rs1 m1: - forall ib fin (ISTEP: iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 (Some fin)), - iblock_istep tge sp rs0 m0 ib rs1 m1 (Some fin). -Admitted. +(* TODO Which approach should we use here? *) +Lemma expand_iblock_istep_exact sp: + forall ib ofin rs0 m0 rs1 m1 + (ISTEP: iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 ofin), + iblock_istep tge sp rs0 m0 ib rs1 m1 ofin. +Proof. Admitted. (* + induction ib; intros; + try (inv ISTEP; econstructor; eauto; fail). + - + rewrite iblock_istep_run_equiv in *. + inv ISTEP. simpl. + + + + destruct (iblock_istep_run tge sp (Bseq ib1 ib2) rs0 m0) eqn:?. + * destruct o. symmetry. + inv Heqo. + destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:?; try discriminate. + destruct o, _fin0 eqn:?; simpl in *. + rewrite <- iblock_istep_run_equiv in *. + eapply expand_iblock_istep_rec_correct; eauto. + inv Heqo0. simpl; auto. + rewrite H1. + rewrite <- iblock_istep_run_equiv in *. + eapply expand_iblock_istep_rec_correct; eauto. + simpl; right; eexists; split. eauto. + apply expand_iblock_istep_correct; auto. + * simpl in *. + destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:?; try discriminate. + destruct o; simpl in *. + destruct _fin eqn:?; simpl in *; try discriminate. + inv Heqo. + + inv ISTEP. simpl; repeat autodestruct; intros. + + destruct o; simpl in *. symmetry. + rewrite <- iblock_istep_run_equiv in *. + eapply expand_iblock_istep_rec_correct; eauto. + inv EQ. simpl; auto. + + destruct o; simpl in *. + rewrite H0. + rewrite <- iblock_istep_run_equiv in *. + destruct (iblock_istep_run tge sp ib2 _rs _m) eqn:?. + * eapply IHib2. destruct o. + rewrite <- iblock_istep_run_equiv in Heqo. + eapply expand_iblock_istep_rec_correct; eauto. + * destruct o; simpl in *. symmetry. + rewrite <- iblock_istep_run_equiv in *. + eapply expand_iblock_istep_rec_correct; eauto. + inv EQ; right. eexists; split. eauto. + apply expand_iblock_istep_correct; auto. + * destruct H0. + 2: { + rewrite iblock_istep_run_equiv in *. inv ISTEP. + destruct _fin eqn:?, o eqn:?; simpl in *. + 2:{ inv Heqo0. destruct ib1 + rewrite iblock_istep_run_equiv in *. inv ISTEP. + + - destruct ib1 eqn:?; simpl in *. + + rewrite iblock_istep_run_equiv in *. + simpl in *. auto. + + repeat econstructor. + rewrite iblock_istep_run_equiv in ISTEP. + simpl in *. + rewrite <- iblock_istep_run_equiv in ISTEP. + apply IHib2; auto. + + rewrite iblock_istep_run_equiv in ISTEP. + simpl in *. + destruct (eval_operation _ _ _ _) eqn:?; try discriminate. + repeat econstructor. eauto. simpl in *. + rewrite <- iblock_istep_run_equiv in ISTEP. + apply IHib2; auto. + + rewrite iblock_istep_run_equiv in ISTEP. + simpl in *. + destruct trap eqn:?; try discriminate. + destruct (eval_addressing _ _ _ _) eqn:?; try discriminate. + destruct (Mem.loadv _ _ _) eqn:?; try discriminate. + repeat econstructor; eauto. simpl in *. + rewrite <- iblock_istep_run_equiv in ISTEP. + apply IHib2; auto. + + rewrite iblock_istep_run_equiv in ISTEP. + simpl in *. + destruct (eval_addressing _ _ _ _) eqn:?; try discriminate. + destruct (Mem.storev _ _ _) eqn:?; try discriminate. + repeat econstructor; eauto. simpl in *. + rewrite <- iblock_istep_run_equiv in ISTEP. + apply IHib2; auto. + + destruct ofin. + * econstructor; eauto. simpl in *. + eapply IHib1; eauto. + + + + + eapply IHib1. + rewrite iblock_istep_run_equiv in *; simpl in *. + inv ISTEP. + destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1. + destruct _fin, o. inv ISTEP. + rewrite iblock_istep_run_equiv in *. + + destruct ib1; simpl in *. + + eapply IHib1. inv ISTEP.*) + +Lemma expand_preserves_iblock_istep_run sp ib: + forall rs m, iblock_istep_run tge sp ib rs m = + iblock_istep_run tge sp (expand ib None) rs m. +Proof. Admitted. (* + intros. + destruct (iblock_istep_run tge sp (expand ib None) rs m) eqn:ISTEP. + (*destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP.*) + - destruct o. + rewrite <- iblock_istep_run_equiv in *. + apply expand_iblock_istep_exact; auto. + - generalize dependent ib. + generalize dependent rs. + generalize dependent m. + induction ib; + try (intros; inv ISTEP; simpl; try autodestruct; fail). + + intros. inv ISTEP. simpl.*) + (* TODO Should we keep the None case in this lemma ? + destruct (iblock_istep_run tge sp ib1 rs m) eqn:EQib1; auto. + destruct o, _fin; simpl in *. + discriminate IHib1. + inv H0. + destruct IHib1. + rewrite <- iblock_istep_run_equiv in EQib1. + apply expand_iblock_istep_correct in EQib1. + discriminate IHib1. + apply EQib1 in IHib1. + inv EQib1. + 2: {, + destruct (iblock_istep_run tge sp (expand ib1 None) rs m) eqn:EQib1. + discriminate IHib1. + discriminate IHib1. + + + induction ib; try (simpl; auto; fail). + - intros. simpl. repeat autodestruct; intros. + + destruct o. symmetry. + rewrite <- iblock_istep_run_equiv in *. + simpl in EQ; rewrite EQ. + eapply expand_iblock_istep_rec_correct; eauto. + rewrite EQ; simpl; auto. + + destruct o. simpl in *. inv EQ. + rewrite <- iblock_istep_run_equiv in *. + 2: { intros. inv EQ. + - simpl.*) + +(* TODO +Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst opc: + forall (MIB: match_iblock dupmap cfg isfst pc ib opc) + k fin2 + (CONT: match opc with + | None => + exists rem, k = Some rem + /\ match_iblock dupmap cfg isfst pc rem (Some fin2) + | Some fin => fin2=fin + end), + match_iblock dupmap cfg isfst pc (expand ib k) (Some fin2). +Proof. + induction 1; simpl. + { destruct k. + intros; subst. econstructor. econstructor. + eapply exec_seq_stop. all: constructor. } + + Qed.*) (** * Match strong state property -- cgit From 99e735af1f7726da2903409758bee202cf47c6a4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 17 May 2021 16:02:57 +0200 Subject: Lemmas on ISTEP --- scheduling/BTL.v | 1 + scheduling/RTLtoBTLproof.v | 320 ++++++++++++++++++--------------------------- 2 files changed, 127 insertions(+), 194 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 064297f1..d55a0415 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -819,6 +819,7 @@ Fixpoint expand (ib: iblock) (k: option iblock): iblock := | 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 + | BF fin => fin | ib => match k with | None => ib diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 4d785c10..4d7508f1 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -260,9 +260,6 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. - -(*TODO Is this lemma too hard to read? *) - Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) k ofin2 rs2 m2 @@ -277,9 +274,8 @@ Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: Proof. induction 1; simpl. { (* BF *) - destruct k; - intros ? ? ? (HRS & HM & HOF); subst. - eapply exec_seq_stop. all: constructor. } + intros ? ? ? ? (HRS & HM & HOF); subst. + constructor. } (*destruct k; intros. try inv CONT.*) 1-4: (* Bnop, Bop, Bload, Bstore *) destruct k; intros; destruct CONT as [[HK [HRS [HM HO]]]|[rem [HR ISTEP]]]; @@ -302,38 +298,6 @@ Proof. destruct b; eapply IHISTEP; eauto. Qed. -(* TODO old version -Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin: - forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin) - k fin2 rs2 m2 - (CONT: match ofin with - | None => - exists rem, k = Some rem - /\ iblock_istep tge sp rs1 m1 rem rs2 m2 (Some fin2) - | Some fin => rs2=rs1 /\ m2=m1 /\ fin2=fin - end), - iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 (Some fin2). -Proof. - induction 1; simpl. - { (* BF *) - destruct k; - intros ? ? ? (HRS & HM & HOF); subst. - eapply exec_seq_stop. all: constructor. } - 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct k; intros ? ? ? (rem & EQR & ISTEP); - inversion EQR; subst; clear EQR; - eapply exec_seq_continue; [ econstructor; eauto | assumption]. - - (* Bseq_stop *) - destruct k; intros; apply IHISTEP; eauto. - - (* Bseq_continue *) - destruct ofin; intros; apply IHISTEP1; - eexists; split; eauto. - - (* Bcond *) - destruct ofin; intros; - econstructor; eauto; - destruct b; eapply IHISTEP; auto. - Qed.*) - Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 ofin. @@ -342,172 +306,140 @@ Proof. destruct ofin; simpl; auto. Qed. -(* TODO Which approach should we use here? *) -Lemma expand_iblock_istep_exact sp: - forall ib ofin rs0 m0 rs1 m1 - (ISTEP: iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 ofin), - iblock_istep tge sp rs0 m0 ib rs1 m1 ofin. -Proof. Admitted. (* - induction ib; intros; - try (inv ISTEP; econstructor; eauto; fail). - - - rewrite iblock_istep_run_equiv in *. - inv ISTEP. simpl. - - - - destruct (iblock_istep_run tge sp (Bseq ib1 ib2) rs0 m0) eqn:?. - * destruct o. symmetry. - inv Heqo. - destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:?; try discriminate. - destruct o, _fin0 eqn:?; simpl in *. - rewrite <- iblock_istep_run_equiv in *. - eapply expand_iblock_istep_rec_correct; eauto. - inv Heqo0. simpl; auto. - rewrite H1. - rewrite <- iblock_istep_run_equiv in *. - eapply expand_iblock_istep_rec_correct; eauto. - simpl; right; eexists; split. eauto. - apply expand_iblock_istep_correct; auto. - * simpl in *. - destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:?; try discriminate. - destruct o; simpl in *. - destruct _fin eqn:?; simpl in *; try discriminate. - inv Heqo. - - inv ISTEP. simpl; repeat autodestruct; intros. - + destruct o; simpl in *. symmetry. - rewrite <- iblock_istep_run_equiv in *. - eapply expand_iblock_istep_rec_correct; eauto. - inv EQ. simpl; auto. - + destruct o; simpl in *. - rewrite H0. - rewrite <- iblock_istep_run_equiv in *. - destruct (iblock_istep_run tge sp ib2 _rs _m) eqn:?. - * eapply IHib2. destruct o. - rewrite <- iblock_istep_run_equiv in Heqo. - eapply expand_iblock_istep_rec_correct; eauto. - * destruct o; simpl in *. symmetry. - rewrite <- iblock_istep_run_equiv in *. - eapply expand_iblock_istep_rec_correct; eauto. - inv EQ; right. eexists; split. eauto. - apply expand_iblock_istep_correct; auto. - * destruct H0. - 2: { - rewrite iblock_istep_run_equiv in *. inv ISTEP. - destruct _fin eqn:?, o eqn:?; simpl in *. - 2:{ inv Heqo0. destruct ib1 - rewrite iblock_istep_run_equiv in *. inv ISTEP. - - - destruct ib1 eqn:?; simpl in *. - + rewrite iblock_istep_run_equiv in *. - simpl in *. auto. - + repeat econstructor. - rewrite iblock_istep_run_equiv in ISTEP. - simpl in *. - rewrite <- iblock_istep_run_equiv in ISTEP. - apply IHib2; auto. - + rewrite iblock_istep_run_equiv in ISTEP. - simpl in *. - destruct (eval_operation _ _ _ _) eqn:?; try discriminate. - repeat econstructor. eauto. simpl in *. - rewrite <- iblock_istep_run_equiv in ISTEP. - apply IHib2; auto. - + rewrite iblock_istep_run_equiv in ISTEP. - simpl in *. - destruct trap eqn:?; try discriminate. - destruct (eval_addressing _ _ _ _) eqn:?; try discriminate. - destruct (Mem.loadv _ _ _) eqn:?; try discriminate. - repeat econstructor; eauto. simpl in *. - rewrite <- iblock_istep_run_equiv in ISTEP. - apply IHib2; auto. - + rewrite iblock_istep_run_equiv in ISTEP. - simpl in *. - destruct (eval_addressing _ _ _ _) eqn:?; try discriminate. - destruct (Mem.storev _ _ _) eqn:?; try discriminate. - repeat econstructor; eauto. simpl in *. - rewrite <- iblock_istep_run_equiv in ISTEP. - apply IHib2; auto. - + destruct ofin. - * econstructor; eauto. simpl in *. - eapply IHib1; eauto. - - - - - eapply IHib1. - rewrite iblock_istep_run_equiv in *; simpl in *. - inv ISTEP. - destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1. - destruct _fin, o. inv ISTEP. - rewrite iblock_istep_run_equiv in *. - - destruct ib1; simpl in *. - + eapply IHib1. inv ISTEP.*) +(* TODO useless? *) +Lemma expand_iblock_istep_run_Some_rec sp ib rs0 m0 rs1 m1 ofin1: + forall (ISTEP: iblock_istep_run tge sp ib rs0 m0 = + Some {| _rs := rs1; _m := m1; _fin := ofin1 |}) + k ofin2 rs2 m2 + (CONT: match ofin1 with + | None => + (k = None /\ rs2=rs1 /\ m2=m1 /\ ofin2 = None) + \/ (exists rem, k = Some rem + /\ iblock_istep_run tge sp rem rs1 m1 = + Some {| _rs := rs2; _m := m2; _fin := ofin2 |}) + | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 + end), + iblock_istep_run tge sp (expand ib k) rs0 m0 = + Some {| _rs := rs2; _m := m2; _fin := ofin2 |}. +Proof. + intros. destruct ofin1; + rewrite <- iblock_istep_run_equiv in *. + - destruct CONT as [HRS [HM HO]]; subst. + eapply expand_iblock_istep_rec_correct; eauto. + simpl; auto. + - eapply expand_iblock_istep_rec_correct; eauto. + simpl. destruct CONT as [HL | [rem [HR ISTEP']]]. + left; auto. rewrite <- iblock_istep_run_equiv in ISTEP'. + right; eexists; split; eauto. +Qed. + +Lemma expand_iblock_istep_run_None_rec sp ib: + forall rs0 m0 o k + (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) + (CONT: match o with + | Some (out rs1 m1 ofin) => + exists rem, + k = Some rem /\ ofin = None /\ + iblock_istep_run tge sp rem rs1 m1 = None + | _ => True + end), + iblock_istep_run tge sp (expand ib k) rs0 m0 = None. +Proof. + induction ib; simpl; + try discriminate. + - (* BF *) + intros; destruct o; try discriminate; simpl in *. + inv ISTEP. destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO. + - (* Bnop *) + intros; destruct o; inv ISTEP; destruct k; + destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + - (* Bop *) + intros; destruct o; + destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; destruct k; + simpl; rewrite EVAL; auto; destruct CONT as [rem [HR [HO ISTEP]]]; + inv HR; inv HO; trivial. + - (* Bload *) + intros; destruct o; + destruct (trap) eqn:TRAP; + try destruct (eval_addressing _ _ _ _) eqn:EVAL; + try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; destruct k; + simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; + destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + - (* Bstore *) + intros; destruct o; + destruct (eval_addressing _ _ _ _) eqn:EVAL; + try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; destruct k; + simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; + destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + - (* Bseq *) + intros. + eapply IHib1; eauto. + destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1; try auto. + destruct o0. eexists; split; eauto. simpl in *. + destruct _fin; inv ISTEP. + + destruct CONT as [rem [_ [CONTRA _]]]; inv CONTRA. + + split; auto. eapply IHib2; eauto. + - (* Bcond *) + intros; destruct (eval_condition _ _ _); trivial. + destruct b. + + eapply IHib1; eauto. + + eapply IHib2; eauto. +Qed. + +Lemma expand_preserves_iblock_istep_run_None sp ib: + forall rs m, iblock_istep_run tge sp ib rs m = None + -> iblock_istep_run tge sp (expand ib None) rs m = None. +Proof. + intros; eapply expand_iblock_istep_run_None_rec; eauto. + simpl; auto. +Qed. Lemma expand_preserves_iblock_istep_run sp ib: forall rs m, iblock_istep_run tge sp ib rs m = iblock_istep_run tge sp (expand ib None) rs m. -Proof. Admitted. (* +Proof. intros. - destruct (iblock_istep_run tge sp (expand ib None) rs m) eqn:ISTEP. - (*destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP.*) - - destruct o. + destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP. + - destruct o. symmetry. rewrite <- iblock_istep_run_equiv in *. - apply expand_iblock_istep_exact; auto. - - generalize dependent ib. - generalize dependent rs. - generalize dependent m. - induction ib; - try (intros; inv ISTEP; simpl; try autodestruct; fail). - + intros. inv ISTEP. simpl.*) - (* TODO Should we keep the None case in this lemma ? - destruct (iblock_istep_run tge sp ib1 rs m) eqn:EQib1; auto. - destruct o, _fin; simpl in *. - discriminate IHib1. - inv H0. - destruct IHib1. - rewrite <- iblock_istep_run_equiv in EQib1. - apply expand_iblock_istep_correct in EQib1. - discriminate IHib1. - apply EQib1 in IHib1. - inv EQib1. - 2: {, - destruct (iblock_istep_run tge sp (expand ib1 None) rs m) eqn:EQib1. - discriminate IHib1. - discriminate IHib1. - - - induction ib; try (simpl; auto; fail). - - intros. simpl. repeat autodestruct; intros. - + destruct o. symmetry. - rewrite <- iblock_istep_run_equiv in *. - simpl in EQ; rewrite EQ. - eapply expand_iblock_istep_rec_correct; eauto. - rewrite EQ; simpl; auto. - + destruct o. simpl in *. inv EQ. - rewrite <- iblock_istep_run_equiv in *. - 2: { intros. inv EQ. - - simpl.*) - -(* TODO -Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst opc: - forall (MIB: match_iblock dupmap cfg isfst pc ib opc) - k fin2 - (CONT: match opc with - | None => + apply expand_iblock_istep_correct; auto. + - symmetry. + apply expand_preserves_iblock_istep_run_None; auto. +Qed. + +Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst opc1: + forall (MIB: match_iblock dupmap cfg isfst pc ib opc1) + k opc2 + (CONT: match opc1 with + | Some pc' => exists rem, k = Some rem - /\ match_iblock dupmap cfg isfst pc rem (Some fin2) - | Some fin => fin2=fin + /\ match_iblock dupmap cfg false pc' rem opc2 + | None => opc2=opc1 end), - match_iblock dupmap cfg isfst pc (expand ib k) (Some fin2). -Proof. + match_iblock dupmap cfg isfst pc (expand ib k) opc2. +Proof. Admitted. + (* induction 1; simpl. - { destruct k. - intros; subst. econstructor. econstructor. - eapply exec_seq_stop. all: constructor. } + { admit. (* TODO we should add the seq_None constructor for this case ? *) } + + 1-4: (* Bnop, Bop, Bload, Bstore *) + destruct k; intros ? (rem & HR & MIB); inv HR; + repeat econstructor; eauto. + + { admit. (* What to do if a goto is the left child of a seq ? *) } + + - (* Bseq_continue *) + destruct opc; intros. + + destruct CONT as [rem [HR MIB]]; subst. + eapply IHMIB1. eexists; repeat split; eauto. + + inv CONT. eapply IHMIB1. eexists; repeat split; eauto. + - (* Bcond *) + destruct opc; intros. + + econstructor; eauto. + eapply IHMIB1; eauto. + autodestruct. intros EQ; inv EQ. - Qed.*) + Qed.*) (** * Match strong state property -- cgit From 00ccf386fc75c1fd6681fbab5aa04c523c55ed9d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 17 May 2021 17:35:26 +0200 Subject: mib lemma --- scheduling/RTLtoBTLproof.v | 55 ++++++++++++++++++++++++---------------------- 1 file changed, 29 insertions(+), 26 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 4d7508f1..faf0b76c 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -407,39 +407,42 @@ Proof. apply expand_preserves_iblock_istep_run_None; auto. Qed. -Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst opc1: - forall (MIB: match_iblock dupmap cfg isfst pc ib opc1) - k opc2 +Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst: + forall opc1 + (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2 (CONT: match opc1 with | Some pc' => - exists rem, k = Some rem - /\ match_iblock dupmap cfg false pc' rem opc2 + k = None /\ opc2 = opc1 \/ + (exists rem, k = Some rem + /\ match_iblock dupmap cfg false pc' rem opc2) | None => opc2=opc1 end), match_iblock dupmap cfg isfst pc (expand ib k) opc2. -Proof. Admitted. - (* +Proof. induction 1; simpl. - { admit. (* TODO we should add the seq_None constructor for this case ? *) } - - 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct k; intros ? (rem & HR & MIB); inv HR; - repeat econstructor; eauto. - - { admit. (* What to do if a goto is the left child of a seq ? *) } - - - (* Bseq_continue *) - destruct opc; intros. - + destruct CONT as [rem [HR MIB]]; subst. - eapply IHMIB1. eexists; repeat split; eauto. - + inv CONT. eapply IHMIB1. eexists; repeat split; eauto. - - (* Bcond *) - destruct opc; intros. - + econstructor; eauto. - eapply IHMIB1; eauto. - autodestruct. intros EQ; inv EQ. + { (* BF *) + intros; inv CONT; econstructor; eauto. } + 1-4: (* Bnop *) + destruct k; intros; destruct CONT as [[HK HO] | [rem [HR MIB]]]; + try inv HK; try inv HO; try inv HR; repeat econstructor; eauto. + { (* Bgoto *) + intros; inv CONT; apply mib_exit; auto. } + { (* Bseq *) + intros. eapply IHMIB1. right. eexists; split; eauto. } + { (* Bcond *) + intros. inv H0; + econstructor; eauto; try econstructor. + destruct opc0; econstructor. } +Qed. - Qed.*) +Lemma expand_matchiblock_correct dupmap cfg ib pc isfst opc: + match_iblock dupmap cfg isfst pc ib opc -> + match_iblock dupmap cfg isfst pc (expand ib None) opc. +Proof. + intros. + eapply expand_matchiblock_rec_correct; eauto. + destruct opc; simpl; auto. +Qed. (** * Match strong state property -- cgit From ee558407e59c6794daad70aab2e1e7794535367e Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 17 May 2021 19:12:46 +0200 Subject: finishing RTLtoBTL --- scheduling/BTL.v | 18 --------------- scheduling/RTLtoBTL.v | 2 +- scheduling/RTLtoBTLproof.v | 57 +++++++++++++++++++++++++++------------------- 3 files changed, 35 insertions(+), 42 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index d55a0415..10a000a8 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -833,21 +833,3 @@ Lemma expand_correct ib: forall k, Proof. induction ib; simpl; intros; try autodestruct; auto. Qed. - -Lemma expand_None_correct ib: - is_expand (expand ib None). -Proof. - apply expand_correct; simpl; auto. -Qed. - -Definition expand_code (cfg: code): code := - (PTree.map (fun _ ib => {| entry:=expand ib.(entry) None; input_regs := ib.(input_regs) |}) cfg). - -Lemma expand_code_correct cfg pc ib : - (expand_code cfg)!pc=Some ib -> is_expand (ib.(entry)). -Proof. - unfold expand_code. - rewrite PTree.gmap. - destruct (cfg!pc); simpl; intros; try_simplify_someHyps. - apply expand_None_correct; auto. -Qed. diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index 3617e691..e9319315 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -12,7 +12,7 @@ Local Open Scope error_monad_scope. Definition transf_function (f: RTL.function) : res BTL.function := let (tcte, dupmap) := rtl2btl f in let (tc, te) := tcte in - let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) (expand_code tc) te in + let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) tc te in do u <- verify_function dupmap f' f; OK f'. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index faf0b76c..633e1b8e 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -9,7 +9,6 @@ Require Import Linking. Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); - code_expand: forall pc ib, (fn_code tf)!pc=Some ib -> is_expand ib.(entry); preserv_fnsig: fn_sig tf = RTL.fn_sig f; preserv_fnparams: fn_params tf = RTL.fn_params f; preserv_fnstacksize: fn_stacksize tf = RTL.fn_stacksize f @@ -31,7 +30,6 @@ Lemma verify_function_correct dupmap f f' tt: fn_sig f' = RTL.fn_sig f -> fn_params f' = RTL.fn_params f -> fn_stacksize f' = RTL.fn_stacksize f -> - (forall pc ib, (fn_code f')!pc=Some ib -> is_expand ib.(entry)) -> match_function dupmap f f'. Proof. unfold verify_function; intro VERIF. monadInv VERIF. @@ -46,7 +44,6 @@ Proof. unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. - eapply expand_code_correct; eauto. Qed. Lemma transf_fundef_correct f f': @@ -243,6 +240,18 @@ Proof. destruct fi; trivial. inv H. inv H4. Qed. +Lemma expand_entry_isnt_goto dupmap f pc ib: + match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) None) None -> + is_goto (expand (entry ib) None) = false. +Proof. + destruct (is_goto (expand (entry ib) None))eqn:EQG. + - destruct (expand (entry ib) None); + try destruct fi; try discriminate; trivial. + intros; inv H; inv H4. + - destruct (expand (entry ib) None); + try destruct fi; try discriminate; trivial. +Qed. + Lemma list_nth_z_rev_dupmap: forall dupmap ln ln' (pc pc': node) val, list_nth_z ln val = Some pc -> @@ -498,10 +507,10 @@ Proof. eexists; left; eexists; split. + repeat econstructor; eauto. apply iblock_istep_run_equiv in BTL_RUN; eauto. - + econstructor. econstructor; eauto. - eapply code_expand; eauto. - econstructor. - eapply entry_isnt_goto; eauto. + + econstructor; apply expand_matchiblock_correct in MI. + econstructor; eauto. apply expand_correct; trivial. + econstructor. apply expand_preserves_iblock_istep_run. + eapply expand_entry_isnt_goto; eauto. - (* Others *) exists (Some ib2); right; split. simpl; auto. @@ -566,11 +575,10 @@ Proof. pose symbols_preserved as SYMPRES. eapply eval_builtin_args_preserved; eauto. eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - * econstructor; eauto. - { econstructor; eauto. eapply code_expand; eauto. - apply star_refl. } - inversion MI; subst; try_simplify_someHyps. - inv H3; try_simplify_someHyps. + * econstructor; eauto; apply expand_matchiblock_correct in MI. + { econstructor; eauto. apply expand_correct; trivial. + apply star_refl. apply expand_preserves_iblock_istep_run. } + eapply expand_entry_isnt_goto; eauto. + (* Bjumptable *) exploit list_nth_z_rev_dupmap; eauto. intros (pc'0 & LNZ & DM). @@ -582,11 +590,10 @@ Proof. eexists; eexists; split. eapply iblock_istep_run_equiv in BTL_RUN. eapply BTL_RUN. econstructor; eauto. - * econstructor; eauto. - { econstructor; eauto. eapply code_expand; eauto. - apply star_refl. } - inversion MI; subst; try_simplify_someHyps. - inv H4; try_simplify_someHyps. + * econstructor; eauto; apply expand_matchiblock_correct in MI. + { econstructor; eauto. apply expand_correct; trivial. + apply star_refl. apply expand_preserves_iblock_istep_run. } + eapply expand_entry_isnt_goto; eauto. - (* mib_exit *) discriminate. - (* mib_seq *) @@ -689,11 +696,13 @@ Proof. eexists; left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. - * econstructor. econstructor; eauto. - eapply code_expand; eauto. - 3: eapply entry_isnt_goto; eauto. + * apply expand_matchiblock_correct in MI. + econstructor. econstructor; eauto. + apply expand_correct; trivial. + 3: eapply expand_entry_isnt_goto; eauto. all: erewrite preserv_fnparams; eauto. constructor. + apply expand_preserves_iblock_istep_run. + (* External function *) inv TRANSF. eexists; left; eexists; split. @@ -708,9 +717,11 @@ Proof. apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. eexists; left; eexists; split. + eapply exec_return. - + econstructor. econstructor; eauto. - eapply code_expand; eauto. - constructor. eapply entry_isnt_goto; eauto. + + apply expand_matchiblock_correct in MI. + econstructor. econstructor; eauto. + apply expand_correct; trivial. + constructor. apply expand_preserves_iblock_istep_run. + eapply expand_entry_isnt_goto; eauto. Qed. Local Hint Resolve plus_one star_refl: core. -- cgit From bf443e2f2bf38c30c9b68020c7c43cd7b3e10549 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 17 May 2021 23:19:16 +0200 Subject: preparing compiler passes and ml oracles --- scheduling/BTL.v | 8 ++++---- scheduling/BTLtoRTL.v | 4 +++- scheduling/BTLtoRTLaux.ml | 5 +++++ scheduling/RTLtoBTL.v | 4 +++- scheduling/RTLtoBTLaux.ml | 17 +++++++++++++++++ scheduling/RTLtoBTLproof.v | 11 +++++------ 6 files changed, 37 insertions(+), 12 deletions(-) create mode 100644 scheduling/BTLtoRTLaux.ml create mode 100644 scheduling/RTLtoBTLaux.ml (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 10a000a8..9fdf39be 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -463,8 +463,8 @@ Local Open Scope error_monad_scope. Definition verify_is_copy dupmap n n' := match dupmap!n with - | None => Error(msg "verify_is_copy None") - | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "verify_is_copy invalid map") end + | None => Error(msg "BTL.verify_is_copy None") + | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "BTL.verify_is_copy invalid map") end end. Fixpoint verify_is_copy_list dupmap ln ln' := @@ -472,9 +472,9 @@ Fixpoint verify_is_copy_list dupmap ln ln' := | n::ln => match ln' with | n'::ln' => do _ <- verify_is_copy dupmap n n'; verify_is_copy_list dupmap ln ln' - | nil => Error (msg "verify_is_copy_list: ln' bigger than ln") end + | nil => Error (msg "BTL.verify_is_copy_list: ln' bigger than ln") end | nil => match ln' with - | n :: ln' => Error (msg "verify_is_copy_list: ln bigger than ln'") + | n :: ln' => Error (msg "BTL.verify_is_copy_list: ln bigger than ln'") | nil => OK tt end end. diff --git a/scheduling/BTLtoRTL.v b/scheduling/BTLtoRTL.v index b64fd87a..82ad47b1 100644 --- a/scheduling/BTLtoRTL.v +++ b/scheduling/BTLtoRTL.v @@ -5,7 +5,9 @@ Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking. (** External oracle *) -Parameter btl2rtl: function -> RTL.code * node * (PTree.t node). +Axiom btl2rtl: function -> RTL.code * node * (PTree.t node). + +Extract Constant btl2rtl => "BTLtoRTLaux.btl2rtl". Local Open Scope error_monad_scope. diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml new file mode 100644 index 00000000..3d8d44d0 --- /dev/null +++ b/scheduling/BTLtoRTLaux.ml @@ -0,0 +1,5 @@ +open Maps +open BinNums + +let btl2rtl f = + ((PTree.empty, Coq_xH), PTree.empty) diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index e9319315..b3585157 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -5,7 +5,9 @@ Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking. (** External oracle *) -Parameter rtl2btl: RTL.function -> BTL.code * node * (PTree.t node). +Axiom rtl2btl: RTL.function -> BTL.code * node * (PTree.t node). + +Extract Constant rtl2btl => "RTLtoBTLaux.rtl2btl". Local Open Scope error_monad_scope. diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml new file mode 100644 index 00000000..8ace6e2a --- /dev/null +++ b/scheduling/RTLtoBTLaux.ml @@ -0,0 +1,17 @@ +open RTLpathLivegenaux +open Maps +open RTL +open BinNums +open DebugPrint + +let rtl2btl (f: RTL.coq_function) = + let code = f.fn_code in + let entry = f.fn_entrypoint in + let join_points = get_join_points code entry in + Printf.eprintf "test"; + debug_flag := true; + debug "Join points: "; + print_true_nodes join_points; + debug "\n"; + debug_flag := false; + ((PTree.empty, Coq_xH), PTree.empty) diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 633e1b8e..5cebd33c 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -231,6 +231,7 @@ Proof. induction ib; simpl; auto; lia. Qed. +(* TODO gourdinl remove useless lemma? *) Lemma entry_isnt_goto dupmap f pc ib: match_iblock dupmap (RTL.fn_code f) true pc (entry ib) None -> is_goto (entry ib) = false. @@ -244,12 +245,10 @@ Lemma expand_entry_isnt_goto dupmap f pc ib: match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) None) None -> is_goto (expand (entry ib) None) = false. Proof. - destruct (is_goto (expand (entry ib) None))eqn:EQG. - - destruct (expand (entry ib) None); - try destruct fi; try discriminate; trivial. - intros; inv H; inv H4. - - destruct (expand (entry ib) None); - try destruct fi; try discriminate; trivial. + destruct (is_goto (expand (entry ib) None))eqn:EQG; + destruct (expand (entry ib) None); + try destruct fi; try discriminate; trivial. + intros; inv H; inv H4. Qed. Lemma list_nth_z_rev_dupmap: -- cgit From 39b00a80b58803546c60856b30d452741db79a23 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 18 May 2021 00:27:14 +0200 Subject: status update --- scheduling/BTLroadmap.md | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index dde2090c..bd70e273 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -54,8 +54,12 @@ Branch factorization should also be possible in BTL -> RTL pass. Example: revert **CURRENT STATUS** - verifier: implemented and proved w.r.t match_iblock specification. -- BTL -> RTL: done. -- RTL -> BTL: started. +- Proof: + - BTL -> RTL: done. + - RTL -> BTL: done. +- Oracles: + - BTL -> RTL: TODO. + - RTL -> BTL: started. **TODO** -- cgit From b7940bfaa83db66837be4d3b9d8b352e8ea4e470 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 18 May 2021 08:26:40 +0200 Subject: todos --- scheduling/RTLtoBTLproof.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 5cebd33c..f89ea5cf 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -314,7 +314,7 @@ Proof. destruct ofin; simpl; auto. Qed. -(* TODO useless? *) +(* TODO gourdinl useless? *) Lemma expand_iblock_istep_run_Some_rec sp ib rs0 m0 rs1 m1 ofin1: forall (ISTEP: iblock_istep_run tge sp ib rs0 m0 = Some {| _rs := rs1; _m := m1; _fin := ofin1 |}) -- cgit From 4553ef98a44da4b34263935b5529b206a89d6dd0 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 18 May 2021 12:42:30 +0200 Subject: First draft of the RTL2BTL oracle --- scheduling/RTLtoBTLaux.ml | 109 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 105 insertions(+), 4 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 8ace6e2a..20aa01aa 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -1,17 +1,118 @@ -open RTLpathLivegenaux open Maps open RTL -open BinNums +open BTL +open Registers open DebugPrint +(* TODO gourdinl duplicated parts from RTLpathLivegenaux.ml + This should be accessible everywhere. *) +let get_some = function +| None -> failwith "Got None instead of Some _" +| Some thing -> thing + +let successors_inst = function +| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] +| Icond (_,_,n1,n2,_) -> [n1; n2] +| Ijumptable (_,l) -> l +| Itailcall _ | Ireturn _ -> [] + +let predicted_successor = function +| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> Some n +| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> None +| Icond (_,_,n1,n2,p) -> ( + match p with + | Some true -> Some n1 + | Some false -> Some n2 + | None -> None ) +| Ijumptable _ | Itailcall _ | Ireturn _ -> None + +(* TODO gourdinl Adapted from RTLpathLivegenaux to avoid redundant computations ... *) +let non_predicted_successors i = function + | None -> successors_inst i + | Some ps -> List.filter (fun s -> s != ps) (successors_inst i) + +(* adapted from Linearizeaux.get_join_points *) +let get_join_points code entry = + let reached = ref (PTree.map (fun n i -> false) code) in + let reached_twice = ref (PTree.map (fun n i -> false) code) in + let rec traverse pc = + if get_some @@ PTree.get pc !reached then begin + if not (get_some @@ PTree.get pc !reached_twice) then + reached_twice := PTree.set pc true !reached_twice + end else begin + reached := PTree.set pc true !reached; + traverse_succs (successors_inst @@ get_some @@ PTree.get pc code) + end + and traverse_succs = function + | [] -> () + | [pc] -> traverse pc + | pc :: l -> traverse pc; traverse_succs l + in traverse entry; !reached_twice + +let translate_inst = function + | Inop (_) -> Bnop + | Iop (op,lr,rd,_) -> Bop (op,lr,rd) + | Iload (trap,chk,addr,lr,rd,_) -> Bload (trap,chk,addr,lr,rd) + | Istore (chk,addr,lr,rd,_) -> Bstore (chk,addr,lr,rd) + | Icall (sign,fn,lr,rd,s) -> BF (Bcall (sign,fn,lr,rd,s)) + | Itailcall (sign,fn,lr) -> BF (Btailcall (sign,fn,lr)) + | Ibuiltin (ef,lr,rd,s) -> BF (Bbuiltin (ef,lr,rd,s)) + | Icond (cond,lr,ifso,ifnot,info) -> ( + (* TODO gourdinl remove this *) + assert (info = None); + Bcond (cond,lr,BF (Bgoto(ifso)),BF (Bgoto(ifnot)),info)) + | Ijumptable (arg, tbl) -> BF (Bjumptable (arg,tbl)) + | Ireturn (oreg) -> BF (Breturn (oreg)) + +let translate_function code entry join_points = + let reached = ref (PTree.map (fun n i -> false) code) in + let btl_code = ref PTree.empty in + (* SEPARATE IN A BETTER WAY!! *) + let rec build_btl_tree e = + if (get_some @@ PTree.get e !reached) then () + else ( + reached := PTree.set e true !reached; + let next_nodes = ref [] in + let rec build_btl_block n = + let inst = get_some @@ PTree.get n code in + let psucc = predicted_successor inst in + let succ = + match psucc with + | Some ps -> if get_some @@ PTree.get ps join_points then None else Some ps + | None -> None + in match succ with + | Some s -> ( + match inst with + | Icond (cond,lr,ifso,ifnot,info) -> ( + (* TODO gourdinl remove this *) + assert (s = ifso || s = ifnot); + next_nodes := !next_nodes @ non_predicted_successors inst psucc; + if s = ifso then + Bcond (cond,lr,build_btl_block s,BF (Bgoto (ifnot)),info) + else + Bcond (cond,lr,BF (Bgoto (ifso)),build_btl_block s,info)) + | _ -> Bseq (translate_inst inst, build_btl_block s)) + | None -> ( + next_nodes := !next_nodes @ successors_inst inst; + translate_inst inst) + in + let ib = build_btl_block e in + let succs = !next_nodes in + let ibf = { entry = ib; input_regs = Regset.empty } in + btl_code := PTree.set e ibf !btl_code; + List.iter build_btl_tree succs) + in build_btl_tree entry; !btl_code + let rtl2btl (f: RTL.coq_function) = let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in - Printf.eprintf "test"; + let btl = translate_function code entry join_points in debug_flag := true; + debug"Code: "; + print_code code; debug "Join points: "; print_true_nodes join_points; debug "\n"; debug_flag := false; - ((PTree.empty, Coq_xH), PTree.empty) + ((btl, entry), PTree.empty) -- cgit From c27d87ffe33242840964dd9bd67090409eea79a5 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 18 May 2021 16:55:56 +0200 Subject: oracle simplification, BTL printer, and error msg spec --- scheduling/BTL.v | 82 +++++++++++------------ scheduling/PrintBTL.ml | 107 ++++++++++++++++++++++++++++++ scheduling/RTLtoBTLaux.ml | 161 ++++++++++++++++++++++++++++++---------------- 3 files changed, 253 insertions(+), 97 deletions(-) create mode 100644 scheduling/PrintBTL.ml (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 9fdf39be..3daa40c4 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -529,13 +529,13 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) do u <- verify_is_copy dupmap pc1 pc; if negb isfst then OK None - else Error (msg "verify_block: isfst is true Bgoto") + else Error (msg "BTL.verify_block: isfst is true Bgoto") | Breturn or => match cfg!pc with | Some (Ireturn or') => if option_eq Pos.eq_dec or or' then OK None - else Error (msg "verify_block: different opt reg in Breturn") - | _ => Error (msg "verify_block: incorrect cfg Breturn") + 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 => match cfg!pc with @@ -545,11 +545,11 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (product_eq Pos.eq_dec ident_eq ri ri') then if (list_eq_dec Pos.eq_dec lr lr') then if (Pos.eq_dec r r') then OK None - else Error (msg "verify_block: different r r' in Bcall") - else Error (msg "verify_block: different lr in Bcall") - else Error (msg "verify_block: different ri in Bcall") - else Error (msg "verify_block: different signatures in Bcall") - | _ => Error (msg "verify_block: incorrect cfg Bcall") + else Error (msg "BTL.verify_block: different r r' in Bcall") + else Error (msg "BTL.verify_block: different lr in Bcall") + else Error (msg "BTL.verify_block: different ri in Bcall") + else Error (msg "BTL.verify_block: different signatures in Bcall") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bcall") end | Btailcall s ri lr => match cfg!pc with @@ -557,10 +557,10 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (signature_eq s s') then if (product_eq Pos.eq_dec ident_eq ri ri') then if (list_eq_dec Pos.eq_dec lr lr') then OK None - else Error (msg "verify_block: different lr in Btailcall") - else Error (msg "verify_block: different ri in Btailcall") - else Error (msg "verify_block: different signatures in Btailcall") - | _ => Error (msg "verify_block: incorrect cfg Btailcall") + else Error (msg "BTL.verify_block: different lr in Btailcall") + else Error (msg "BTL.verify_block: different ri in Btailcall") + else Error (msg "BTL.verify_block: different signatures in Btailcall") + | _ => Error (msg "BTL.verify_block: incorrect cfg Btailcall") end | Bbuiltin ef la br pc1 => match cfg!pc with @@ -569,24 +569,24 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (external_function_eq ef ef') then if (list_eq_dec builtin_arg_eq_pos la la') then if (builtin_res_eq_pos br br') then OK None - else Error (msg "verify_block: different brr in Bbuiltin") - else Error (msg "verify_block: different lbar in Bbuiltin") - else Error (msg "verify_block: different ef in Bbuiltin") - | _ => Error (msg "verify_block: incorrect cfg Bbuiltin") + else Error (msg "BTL.verify_block: different brr in Bbuiltin") + else Error (msg "BTL.verify_block: different lbar in Bbuiltin") + else Error (msg "BTL.verify_block: different ef in Bbuiltin") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bbuiltin") end | Bjumptable r ln => match cfg!pc with | Some (Ijumptable r' ln') => do u <- verify_is_copy_list dupmap ln ln'; if (Pos.eq_dec r r') then OK None - else Error (msg "verify_block: different r in Bjumptable") - | _ => Error (msg "verify_block: incorrect cfg Bjumptable") + else Error (msg "BTL.verify_block: different r in Bjumptable") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bjumptable") end end | Bnop => match cfg!pc with | Some (Inop pc') => OK (Some pc') - | _ => Error (msg "verify_block: incorrect cfg Bnop") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") end | Bop op lr r => match cfg!pc with @@ -595,10 +595,10 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (list_eq_dec Pos.eq_dec lr lr') then if (Pos.eq_dec r r') then OK (Some pc') - else Error (msg "verify_block: different r in Bop") - else Error (msg "verify_block: different lr in Bop") - else Error (msg "verify_block: different operations in Bop") - | _ => Error (msg "verify_block: incorrect cfg Bop") + else Error (msg "BTL.verify_block: different r in Bop") + else Error (msg "BTL.verify_block: different lr in Bop") + 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 => match cfg!pc with @@ -609,12 +609,12 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (list_eq_dec Pos.eq_dec lr lr') then if (Pos.eq_dec r r') then OK (Some pc') - else Error (msg "verify_block: different r in Bload") - else Error (msg "verify_block: different lr in Bload") - else Error (msg "verify_block: different addressing in Bload") - else Error (msg "verify_block: different mchunk in Bload") - else Error (msg "verify_block: NOTRAP trapping_mode unsupported in Bload") - | _ => Error (msg "verify_block: incorrect cfg Bload") + else Error (msg "BTL.verify_block: different r in Bload") + else Error (msg "BTL.verify_block: different lr in Bload") + else Error (msg "BTL.verify_block: different addressing in Bload") + else Error (msg "BTL.verify_block: different mchunk in Bload") + 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 => match cfg!pc with @@ -623,18 +623,18 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) if (eq_addressing a a') then if (list_eq_dec Pos.eq_dec lr lr') then if (Pos.eq_dec r r') then OK (Some pc') - else Error (msg "verify_block: different r in Bstore") - else Error (msg "verify_block: different lr in Bstore") - else Error (msg "verify_block: different addressing in Bstore") - else Error (msg "verify_block: different mchunk in Bstore") - | _ => Error (msg "verify_block: incorrect cfg Bstore") + else Error (msg "BTL.verify_block: different r in Bstore") + else Error (msg "BTL.verify_block: different lr in Bstore") + else Error (msg "BTL.verify_block: different addressing in Bstore") + else Error (msg "BTL.verify_block: different mchunk in Bstore") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bstore") end | Bseq b1 b2 => do opc <- verify_block dupmap cfg isfst pc b1; match opc with | Some pc' => verify_block dupmap cfg false pc' b2 - | None => Error (msg "verify_block: None next pc in Bseq (deadcode)") + | None => Error (msg "BTL.verify_block: None next pc in Bseq (deadcode)") end | Bcond c lr bso bnot i => match cfg!pc with @@ -648,11 +648,11 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) | o, None => OK o | Some x, Some x' => if Pos.eq_dec x x' then OK (Some x) - else Error (msg "verify_block: is_join_opt incorrect for Bcond") + else Error (msg "BTL.verify_block: is_join_opt incorrect for Bcond") end - else Error (msg "verify_block: incompatible conditions in Bcond") - else Error (msg "verify_block: different lr in Bcond") - | _ => Error (msg "verify_block: incorrect cfg Bcond") + else Error (msg "BTL.verify_block: incompatible conditions in Bcond") + else Error (msg "BTL.verify_block: different lr in Bcond") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bcond") end end. @@ -753,9 +753,9 @@ Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit := | Some ib => do o <- verify_block dupmap cfg' true pc' ib.(entry); match o with | None => verify_blocks dupmap cfg cfg' l - | _ => Error(msg "verify_blocks.end") + | _ => Error(msg "BTL.verify_blocks.end") end - | _ => Error(msg "verify_blocks.entry") + | _ => Error(msg "BTL.verify_blocks.entry") end end. diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml new file mode 100644 index 00000000..8f61380e --- /dev/null +++ b/scheduling/PrintBTL.ml @@ -0,0 +1,107 @@ +open Printf +open Camlcoq +open Datatypes +open Maps +open AST +open BTL +open PrintAST + +(* Printing of BTL code *) + +let reg pp r = + fprintf pp "x%d" (P.to_int r) + +let rec regs pp = function + | [] -> () + | [r] -> reg pp r + | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl + +let ros pp = function + | Coq_inl r -> reg pp r + | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s) + +let print_succ pp s = + fprintf pp "\tsucc %d\n" (P.to_int s) + +let rec print_iblock pp is_rec pref ib = + fprintf pp "%s" pref; + match ib with + | Bnop -> + fprintf pp "Bnop\n" + | Bop(op, args, res) -> + fprintf pp "Bop: %a = %a\n" + reg res (PrintOp.print_operation reg) (op, args) + | Bload(trap, chunk, addr, args, dst) -> + fprintf pp "Bload: %a = %s[%a]%a\n" + reg dst (name_of_chunk chunk) + (PrintOp.print_addressing reg) (addr, args) + print_trapping_mode trap + | Bstore(chunk, addr, args, src) -> + fprintf pp "Bstore: %s[%a] = %a\n" + (name_of_chunk chunk) + (PrintOp.print_addressing reg) (addr, args) + reg src + | BF (Bcall(sg, fn, args, res, s)) -> + fprintf pp "Bcall: %a = %a(%a)\n" + reg res ros fn regs args; + print_succ pp s + | BF (Btailcall(sg, fn, args)) -> + fprintf pp "Btailcall: %a(%a)\n" + ros fn regs args + | BF (Bbuiltin(ef, args, res, s)) -> + fprintf pp "Bbuiltin: %a = %s(%a)\n" + (print_builtin_res reg) res + (name_of_external ef) + (print_builtin_args reg) args; + print_succ pp s + | Bcond(cond, args, ib1, ib2, info) -> + fprintf pp "Bcond: (%a) (prediction: %s)\n" + (PrintOp.print_condition reg) (cond, args) + (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough"); + let pref' = pref ^ " " in + fprintf pp "%sifso = [\n" pref; + if is_rec then print_iblock pp is_rec pref' ib1 else fprintf pp "...\n"; + fprintf pp "%s]\n" pref; + fprintf pp "%sifnot = [\n" pref; + if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n"; + fprintf pp "%s]\n" pref + | BF (Bjumptable(arg, tbl)) -> + let tbl = Array.of_list tbl in + fprintf pp "Bjumptable: (%a)\n" reg arg; + for i = 0 to Array.length tbl - 1 do + fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i)) + done + | BF (Breturn None) -> + fprintf pp "Breturn\n" + | BF (Breturn (Some arg)) -> + fprintf pp "Breturn: %a\n" reg arg + | BF (Bgoto s) -> + fprintf pp "Bgoto: %d\n" (P.to_int s) + | Bseq (ib1, ib2) -> + if is_rec then ( + print_iblock pp is_rec pref ib1; + print_iblock pp is_rec pref ib2) + else fprintf pp "Bseq...\n" + +let print_btl_code pp btl is_rec = + fprintf pp "\n"; + List.iter (fun (n,ibf) -> + fprintf pp "[BTL block %d]\n" (P.to_int n); + print_iblock pp is_rec " " ibf.entry; + fprintf pp "\n") + (PTree.elements btl); + fprintf pp "\n" + +let print_function pp id f = + fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; + let instrs = List.map (fun (n,i) -> i.entry) (PTree.elements f.fn_code) in + List.iter (print_iblock pp true "") instrs; + fprintf pp "}\n\n" + +let print_globdef pp (id, gd) = + match gd with + | Gfun(Internal f) -> print_function pp id f + | _ -> () + +let print_program pp (prog: BTL.program) = + List.iter (print_globdef pp) prog.prog_defs diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 20aa01aa..3c9b7644 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -3,28 +3,35 @@ open RTL open BTL open Registers open DebugPrint +open PrintRTL +open PrintBTL (* TODO gourdinl duplicated parts from RTLpathLivegenaux.ml This should be accessible everywhere. *) let get_some = function -| None -> failwith "Got None instead of Some _" -| Some thing -> thing + | None -> failwith "Got None instead of Some _" + | Some thing -> thing let successors_inst = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] -| Icond (_,_,n1,n2,_) -> [n1; n2] -| Ijumptable (_,l) -> l -| Itailcall _ | Ireturn _ -> [] + | Inop n + | Iop (_, _, _, n) + | Iload (_, _, _, _, _, n) + | Istore (_, _, _, _, n) + | Icall (_, _, _, _, n) + | Ibuiltin (_, _, _, n) -> + [ n ] + | Icond (_, _, n1, n2, _) -> [ n1; n2 ] + | Ijumptable (_, l) -> l + | Itailcall _ | Ireturn _ -> [] let predicted_successor = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> Some n -| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> None -| Icond (_,_,n1,n2,p) -> ( - match p with - | Some true -> Some n1 - | Some false -> Some n2 - | None -> None ) -| Ijumptable _ | Itailcall _ | Ireturn _ -> None + | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) + -> + Some n + | Icall (_, _, _, _, n) | Ibuiltin (_, _, _, n) -> None + | Icond (_, _, n1, n2, p) -> ( + match p with Some true -> Some n1 | Some false -> Some n2 | None -> None) + | Ijumptable _ | Itailcall _ | Ireturn _ -> None (* TODO gourdinl Adapted from RTLpathLivegenaux to avoid redundant computations ... *) let non_predicted_successors i = function @@ -36,83 +43,125 @@ let get_join_points code entry = let reached = ref (PTree.map (fun n i -> false) code) in let reached_twice = ref (PTree.map (fun n i -> false) code) in let rec traverse pc = - if get_some @@ PTree.get pc !reached then begin + if get_some @@ PTree.get pc !reached then ( if not (get_some @@ PTree.get pc !reached_twice) then - reached_twice := PTree.set pc true !reached_twice - end else begin + reached_twice := PTree.set pc true !reached_twice) + else ( reached := PTree.set pc true !reached; - traverse_succs (successors_inst @@ get_some @@ PTree.get pc code) - end + traverse_succs (successors_inst @@ get_some @@ PTree.get pc code)) and traverse_succs = function | [] -> () - | [pc] -> traverse pc - | pc :: l -> traverse pc; traverse_succs l - in traverse entry; !reached_twice + | [ pc ] -> traverse pc + | pc :: l -> + traverse pc; + traverse_succs l + in + traverse entry; + !reached_twice + +let encaps_final inst osucc = + match inst with + | BF _ | Bcond _ -> inst + | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc))) -let translate_inst = function - | Inop (_) -> Bnop - | Iop (op,lr,rd,_) -> Bop (op,lr,rd) - | Iload (trap,chk,addr,lr,rd,_) -> Bload (trap,chk,addr,lr,rd) - | Istore (chk,addr,lr,rd,_) -> Bstore (chk,addr,lr,rd) - | Icall (sign,fn,lr,rd,s) -> BF (Bcall (sign,fn,lr,rd,s)) - | Itailcall (sign,fn,lr) -> BF (Btailcall (sign,fn,lr)) - | Ibuiltin (ef,lr,rd,s) -> BF (Bbuiltin (ef,lr,rd,s)) - | Icond (cond,lr,ifso,ifnot,info) -> ( +let translate_inst inst is_final = + let osucc = ref None in + let btli = + match inst with + | Inop s -> + osucc := Some s; + Bnop + | Iop (op, lr, rd, s) -> + osucc := Some s; + Bop (op, lr, rd) + | Iload (trap, chk, addr, lr, rd, s) -> + osucc := Some s; + Bload (trap, chk, addr, lr, rd) + | Istore (chk, addr, lr, rd, s) -> + osucc := Some s; + Bstore (chk, addr, lr, rd) + | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s)) + | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr)) + | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s)) + | Icond (cond, lr, ifso, ifnot, info) -> (* TODO gourdinl remove this *) assert (info = None); - Bcond (cond,lr,BF (Bgoto(ifso)),BF (Bgoto(ifnot)),info)) - | Ijumptable (arg, tbl) -> BF (Bjumptable (arg,tbl)) - | Ireturn (oreg) -> BF (Breturn (oreg)) + Bcond (cond, lr, BF (Bgoto ifso), BF (Bgoto ifnot), info) + | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl)) + | Ireturn oreg -> BF (Breturn oreg) + in + if is_final then encaps_final btli !osucc else btli let translate_function code entry join_points = let reached = ref (PTree.map (fun n i -> false) code) in let btl_code = ref PTree.empty in (* SEPARATE IN A BETTER WAY!! *) let rec build_btl_tree e = - if (get_some @@ PTree.get e !reached) then () + if get_some @@ PTree.get e !reached then () else ( reached := PTree.set e true !reached; let next_nodes = ref [] in let rec build_btl_block n = let inst = get_some @@ PTree.get n code in + debug "Inst is : "; + print_instruction stderr (0, inst); + debug "\n"; let psucc = predicted_successor inst in let succ = - match psucc with - | Some ps -> if get_some @@ PTree.get ps join_points then None else Some ps - | None -> None - in match succ with + match psucc with + | Some ps -> + if get_some @@ PTree.get ps join_points then ( + debug "IS JOIN PT\n"; + None) + else Some ps + | None -> None + in + match succ with | Some s -> ( - match inst with - | Icond (cond,lr,ifso,ifnot,info) -> ( - (* TODO gourdinl remove this *) - assert (s = ifso || s = ifnot); - next_nodes := !next_nodes @ non_predicted_successors inst psucc; - if s = ifso then - Bcond (cond,lr,build_btl_block s,BF (Bgoto (ifnot)),info) - else - Bcond (cond,lr,BF (Bgoto (ifso)),build_btl_block s,info)) - | _ -> Bseq (translate_inst inst, build_btl_block s)) - | None -> ( + debug "BLOCK CONTINUE\n"; + match inst with + | Icond (cond, lr, ifso, ifnot, info) -> + (* TODO gourdinl remove this *) + assert (s = ifnot); + next_nodes := !next_nodes @ non_predicted_successors inst psucc; + Bcond (cond, lr, BF (Bgoto ifso), build_btl_block s, info) + | _ -> Bseq (translate_inst inst false, build_btl_block s)) + | None -> + debug "BLOCK END.\n"; next_nodes := !next_nodes @ successors_inst inst; - translate_inst inst) + translate_inst inst true in let ib = build_btl_block e in let succs = !next_nodes in let ibf = { entry = ib; input_regs = Regset.empty } in btl_code := PTree.set e ibf !btl_code; List.iter build_btl_tree succs) - in build_btl_tree entry; !btl_code + in + build_btl_tree entry; + !btl_code + +(*let print_dm dm =*) + (*List.iter (fun (n,ib) ->*) + (*debug "%d:" (P.to_int n);*) + (*print_iblock stderr false ib.entry)*) + (*(PTree.elements dm)*) + -let rtl2btl (f: RTL.coq_function) = + +let rtl2btl (f : RTL.coq_function) = + debug_flag := true; let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in let btl = translate_function code entry join_points in - debug_flag := true; - debug"Code: "; + let dm = PTree.map (fun n i -> n) btl in + debug "RTL Code: "; print_code code; + debug "BTL Code: "; + print_btl_code stderr btl true; debug "Join points: "; print_true_nodes join_points; debug "\n"; + (*print_dm dm;*) debug_flag := false; - ((btl, entry), PTree.empty) + ((btl, entry), dm) -- cgit From af2208a2c7126d4d101fb07c40920e12c9ebbab3 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 18 May 2021 19:15:39 +0200 Subject: first oracle seems ok --- scheduling/PrintBTL.ml | 3 +++ scheduling/RTLtoBTLaux.ml | 23 ++++++++++++----------- 2 files changed, 15 insertions(+), 11 deletions(-) (limited to 'scheduling') diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 8f61380e..781dcaf3 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -83,6 +83,9 @@ let rec print_iblock pp is_rec pref ib = print_iblock pp is_rec pref ib2) else fprintf pp "Bseq...\n" +let print_btl_inst pp ib = + print_iblock pp false " " ib + let print_btl_code pp btl is_rec = fprintf pp "\n"; List.iter (fun (n,ibf) -> diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 3c9b7644..9a61f55e 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -5,6 +5,7 @@ open Registers open DebugPrint open PrintRTL open PrintBTL +open Camlcoq (* TODO gourdinl duplicated parts from RTLpathLivegenaux.ml This should be accessible everywhere. *) @@ -84,8 +85,6 @@ let translate_inst inst is_final = | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr)) | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s)) | Icond (cond, lr, ifso, ifnot, info) -> - (* TODO gourdinl remove this *) - assert (info = None); Bcond (cond, lr, BF (Bgoto ifso), BF (Bgoto ifnot), info) | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl)) | Ireturn oreg -> BF (Breturn oreg) @@ -140,28 +139,30 @@ let translate_function code entry join_points = build_btl_tree entry; !btl_code -(*let print_dm dm =*) - (*List.iter (fun (n,ib) ->*) - (*debug "%d:" (P.to_int n);*) - (*print_iblock stderr false ib.entry)*) - (*(PTree.elements dm)*) - - +let print_dm dm = + List.iter (fun (n,n') -> + debug "%d -> %d" (P.to_int n) (P.to_int n'); + (*print_btl_inst stderr ib.entry;*) + debug "\n" + ) + (PTree.elements dm) let rtl2btl (f : RTL.coq_function) = - debug_flag := true; + (*debug_flag := true;*) let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in let btl = translate_function code entry join_points in let dm = PTree.map (fun n i -> n) btl in + debug "Entry %d\n" (P.to_int entry); debug "RTL Code: "; print_code code; debug "BTL Code: "; print_btl_code stderr btl true; + debug "BTL Dupmap:\n"; + print_dm dm; debug "Join points: "; print_true_nodes join_points; debug "\n"; - (*print_dm dm;*) debug_flag := false; ((btl, entry), dm) -- cgit From ab520acd80f7d39aa14fdda6932accbb2a787ebf Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 19 May 2021 12:47:22 +0200 Subject: Grouping common RTL functions, printer improvement --- scheduling/PrintBTL.ml | 16 ++++++++- scheduling/RTLpathLivegenaux.ml | 49 +++------------------------- scheduling/RTLpathScheduleraux.ml | 1 + scheduling/RTLtoBTLaux.ml | 68 ++++----------------------------------- 4 files changed, 26 insertions(+), 108 deletions(-) (limited to 'scheduling') diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 781dcaf3..b461e3f1 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -23,38 +23,48 @@ let ros pp = function let print_succ pp s = fprintf pp "\tsucc %d\n" (P.to_int s) +let print_pref pp pref = + fprintf pp "%s" pref + let rec print_iblock pp is_rec pref ib = - fprintf pp "%s" pref; match ib with | Bnop -> + print_pref pp pref; fprintf pp "Bnop\n" | Bop(op, args, res) -> + print_pref pp pref; fprintf pp "Bop: %a = %a\n" reg res (PrintOp.print_operation reg) (op, args) | Bload(trap, chunk, addr, args, dst) -> + print_pref pp pref; fprintf pp "Bload: %a = %s[%a]%a\n" reg dst (name_of_chunk chunk) (PrintOp.print_addressing reg) (addr, args) print_trapping_mode trap | Bstore(chunk, addr, args, src) -> + print_pref pp pref; fprintf pp "Bstore: %s[%a] = %a\n" (name_of_chunk chunk) (PrintOp.print_addressing reg) (addr, args) reg src | BF (Bcall(sg, fn, args, res, s)) -> + print_pref pp pref; fprintf pp "Bcall: %a = %a(%a)\n" reg res ros fn regs args; print_succ pp s | BF (Btailcall(sg, fn, args)) -> + print_pref pp pref; fprintf pp "Btailcall: %a(%a)\n" ros fn regs args | BF (Bbuiltin(ef, args, res, s)) -> + print_pref pp pref; fprintf pp "Bbuiltin: %a = %s(%a)\n" (print_builtin_res reg) res (name_of_external ef) (print_builtin_args reg) args; print_succ pp s | Bcond(cond, args, ib1, ib2, info) -> + print_pref pp pref; fprintf pp "Bcond: (%a) (prediction: %s)\n" (PrintOp.print_condition reg) (cond, args) (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough"); @@ -67,15 +77,19 @@ let rec print_iblock pp is_rec pref ib = fprintf pp "%s]\n" pref | BF (Bjumptable(arg, tbl)) -> let tbl = Array.of_list tbl in + print_pref pp pref; fprintf pp "Bjumptable: (%a)\n" reg arg; for i = 0 to Array.length tbl - 1 do fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i)) done | BF (Breturn None) -> + print_pref pp pref; fprintf pp "Breturn\n" | BF (Breturn (Some arg)) -> + print_pref pp pref; fprintf pp "Breturn: %a\n" reg arg | BF (Bgoto s) -> + print_pref pp pref; fprintf pp "Bgoto: %d\n" (P.to_int s) | Bseq (ib1, ib2) -> if is_rec then ( diff --git a/scheduling/RTLpathLivegenaux.ml b/scheduling/RTLpathLivegenaux.ml index 2a20a15d..c9bb94d3 100644 --- a/scheduling/RTLpathLivegenaux.ml +++ b/scheduling/RTLpathLivegenaux.ml @@ -7,31 +7,7 @@ open Datatypes open Kildall open Lattice open DebugPrint - -let get_some = function -| None -> failwith "Got None instead of Some _" -| Some thing -> thing - -let successors_inst = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] -| Icond (_,_,n1,n2,_) -> [n1; n2] -| Ijumptable (_,l) -> l -| Itailcall _ | Ireturn _ -> [] - -let predicted_successor = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> Some n -| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> None -| Icond (_,_,n1,n2,p) -> ( - match p with - | Some true -> Some n1 - | Some false -> Some n2 - | None -> None ) -| Ijumptable _ | Itailcall _ | Ireturn _ -> None - -let non_predicted_successors i = - match predicted_successor i with - | None -> successors_inst i - | Some n -> List.filter (fun n' -> n != n') (successors_inst i) +open AuxTools let rec list_to_regset = function | [] -> Regset.empty @@ -59,24 +35,6 @@ let get_output_reg i = | Iop (_, _, r, _) | Iload (_, _, _, _, r, _) | Icall (_, _, _, r, _) -> Some r | Ibuiltin (_, _, brr, _) -> (match brr with AST.BR r -> Some r | _ -> None) -(* adapted from Linearizeaux.get_join_points *) -let get_join_points code entry = - let reached = ref (PTree.map (fun n i -> false) code) in - let reached_twice = ref (PTree.map (fun n i -> false) code) in - let rec traverse pc = - if get_some @@ PTree.get pc !reached then begin - if not (get_some @@ PTree.get pc !reached_twice) then - reached_twice := PTree.set pc true !reached_twice - end else begin - reached := PTree.set pc true !reached; - traverse_succs (successors_inst @@ get_some @@ PTree.get pc code) - end - and traverse_succs = function - | [] -> () - | [pc] -> traverse pc - | pc :: l -> traverse pc; traverse_succs l - in traverse entry; !reached_twice - (* Does not set the input_regs and liveouts field *) let get_path_map code entry join_points = let visited = ref (PTree.map (fun n i -> false) code) in @@ -92,12 +50,13 @@ let get_path_map code entry join_points = let inst = get_some @@ PTree.get n code in begin psize := !psize + 1; - let successor = match predicted_successor inst with + let psucc = predicted_successor inst in + let successor = match psucc with | None -> None | Some n' -> if get_some @@ PTree.get n' join_points then None else Some n' in match successor with | Some n' -> begin - path_successors := !path_successors @ non_predicted_successors inst; + path_successors := !path_successors @ non_predicted_successors inst psucc; dig_path_rec n' end | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); diff --git a/scheduling/RTLpathScheduleraux.ml b/scheduling/RTLpathScheduleraux.ml index aeed39df..70a0c507 100644 --- a/scheduling/RTLpathScheduleraux.ml +++ b/scheduling/RTLpathScheduleraux.ml @@ -7,6 +7,7 @@ open RTL open Maps open Registers open ExpansionOracle +open AuxTools let config = Machine.config diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 9a61f55e..e14e0ab3 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -3,62 +3,9 @@ open RTL open BTL open Registers open DebugPrint -open PrintRTL open PrintBTL open Camlcoq - -(* TODO gourdinl duplicated parts from RTLpathLivegenaux.ml - This should be accessible everywhere. *) -let get_some = function - | None -> failwith "Got None instead of Some _" - | Some thing -> thing - -let successors_inst = function - | Inop n - | Iop (_, _, _, n) - | Iload (_, _, _, _, _, n) - | Istore (_, _, _, _, n) - | Icall (_, _, _, _, n) - | Ibuiltin (_, _, _, n) -> - [ n ] - | Icond (_, _, n1, n2, _) -> [ n1; n2 ] - | Ijumptable (_, l) -> l - | Itailcall _ | Ireturn _ -> [] - -let predicted_successor = function - | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) - -> - Some n - | Icall (_, _, _, _, n) | Ibuiltin (_, _, _, n) -> None - | Icond (_, _, n1, n2, p) -> ( - match p with Some true -> Some n1 | Some false -> Some n2 | None -> None) - | Ijumptable _ | Itailcall _ | Ireturn _ -> None - -(* TODO gourdinl Adapted from RTLpathLivegenaux to avoid redundant computations ... *) -let non_predicted_successors i = function - | None -> successors_inst i - | Some ps -> List.filter (fun s -> s != ps) (successors_inst i) - -(* adapted from Linearizeaux.get_join_points *) -let get_join_points code entry = - let reached = ref (PTree.map (fun n i -> false) code) in - let reached_twice = ref (PTree.map (fun n i -> false) code) in - let rec traverse pc = - if get_some @@ PTree.get pc !reached then ( - if not (get_some @@ PTree.get pc !reached_twice) then - reached_twice := PTree.set pc true !reached_twice) - else ( - reached := PTree.set pc true !reached; - traverse_succs (successors_inst @@ get_some @@ PTree.get pc code)) - and traverse_succs = function - | [] -> () - | [ pc ] -> traverse pc - | pc :: l -> - traverse pc; - traverse_succs l - in - traverse entry; - !reached_twice +open AuxTools let encaps_final inst osucc = match inst with @@ -94,7 +41,7 @@ let translate_inst inst is_final = let translate_function code entry join_points = let reached = ref (PTree.map (fun n i -> false) code) in let btl_code = ref PTree.empty in - (* SEPARATE IN A BETTER WAY!! *) + (* TODO gourdinl SEPARATE IN A BETTER WAY!! *) let rec build_btl_tree e = if get_some @@ PTree.get e !reached then () else ( @@ -102,9 +49,6 @@ let translate_function code entry join_points = let next_nodes = ref [] in let rec build_btl_block n = let inst = get_some @@ PTree.get n code in - debug "Inst is : "; - print_instruction stderr (0, inst); - debug "\n"; let psucc = predicted_successor inst in let succ = match psucc with @@ -148,7 +92,7 @@ let print_dm dm = (PTree.elements dm) let rtl2btl (f : RTL.coq_function) = - (*debug_flag := true;*) + debug_flag := true; let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in @@ -157,12 +101,12 @@ let rtl2btl (f : RTL.coq_function) = debug "Entry %d\n" (P.to_int entry); debug "RTL Code: "; print_code code; + debug_flag := false; debug "BTL Code: "; print_btl_code stderr btl true; - debug "BTL Dupmap:\n"; - print_dm dm; + debug "Dupmap:\n"; + print_ptree print_pint dm; debug "Join points: "; print_true_nodes join_points; debug "\n"; - debug_flag := false; ((btl, entry), dm) -- cgit 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 +++++++++++++++++++++++---------------------- scheduling/BTL_SEtheory.v | 22 +++---- scheduling/BTLtoRTLaux.ml | 104 ++++++++++++++++++++++++++++++++- scheduling/BTLtoRTLproof.v | 6 +- scheduling/PrintBTL.ml | 110 +++++++++++++++++----------------- scheduling/RTLtoBTLaux.ml | 43 +++++++------- scheduling/RTLtoBTLproof.v | 2 +- 7 files changed, 265 insertions(+), 165 deletions(-) (limited to 'scheduling') 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 diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index b9a05a8a..94d299e7 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -557,21 +557,21 @@ Qed. Definition sexec_final_sfv (i: final) (sis: sistate): sfval := match i with | Bgoto pc => Sgoto pc - | Bcall sig ros args res pc => + | Bcall sig ros args res pc _ => let svos := sum_left_map sis.(si_sreg) ros in let sargs := list_sval_inj (List.map sis.(si_sreg) args) in Scall sig svos sargs res pc - | Btailcall sig ros args => + | Btailcall sig ros args _ => let svos := sum_left_map sis.(si_sreg) ros in let sargs := list_sval_inj (List.map sis.(si_sreg) args) in Stailcall sig svos sargs - | Bbuiltin ef args res pc => + | Bbuiltin ef args res pc _ => let sargs := List.map (builtin_arg_map sis.(si_sreg)) args in Sbuiltin ef sargs res pc - | Breturn or => + | Breturn or _ => let sor := SOME r <- or IN Some (sis.(si_sreg) r) in Sreturn sor - | Bjumptable reg tbl => + | Bjumptable reg tbl _ => let sv := sis.(si_sreg) reg in Sjumptable sv tbl end. @@ -772,15 +772,15 @@ Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := match ib with | BF fin => Sfinal sis (sexec_final_sfv fin sis) (* basic instructions *) - | Bnop => k sis - | Bop op args res => k (sexec_op op args res sis) - | Bload TRAP chunk addr args dst => k (sexec_load TRAP chunk addr args dst sis) - | Bload NOTRAP chunk addr args dst => Sabort (* TODO *) - | Bstore chunk addr args src => k (sexec_store chunk addr args src sis) + | Bnop _ => k sis + | Bop op args res _ => k (sexec_op op args res sis) + | Bload TRAP chunk addr args dst _ => k (sexec_load TRAP chunk addr args dst sis) + | Bload NOTRAP chunk addr args dst _ => Sabort (* TODO *) + | Bstore chunk addr args src _ => k (sexec_store chunk addr args src sis) (* composed instructions *) | Bseq ib1 ib2 => sexec_rec ib1 sis (fun sis2 => sexec_rec ib2 sis2 k) - | Bcond cond args ifso ifnot _ => + | Bcond cond args ifso ifnot _ _ => let args := list_sval_inj (List.map sis.(si_sreg) args) in let ifso := sexec_rec ifso sis k in let ifnot := sexec_rec ifnot sis k in diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 3d8d44d0..dd7ba4c7 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -1,5 +1,103 @@ open Maps -open BinNums +open BTL +open RTL +open Camlcoq +open AuxTools +open DebugPrint +open PrintBTL -let btl2rtl f = - ((PTree.empty, Coq_xH), PTree.empty) +let is_atom = function + | Bseq (_, _) | Bcond (_, _, _, _, _, _) -> false + | _ -> true + +let rec get_nn = function + | Bnop ni + | Bop (_, _, _, ni) + | Bload (_, _, _, _, _, ni) + | Bstore (_, _, _, _, ni) + | Bcond (_, _, _, _, _, ni) + | BF (Breturn (_, ni)) + | BF (Bcall (_, _, _, _, _, ni)) + | BF (Btailcall (_, _, _, ni)) + | BF (Bbuiltin (_, _, _, _, ni)) + | BF (Bjumptable (_, _, ni)) -> + ni + | BF (Bgoto s) -> s + | Bseq (ib1, _) -> get_nn ib1 + +let translate_function code entry = + let reached = ref (PTree.map (fun n i -> false) code) in + let rtl_code = ref PTree.empty in + let rec build_rtl_tree e = + if get_some @@ PTree.get e !reached then () + else ( + reached := PTree.set e true !reached; + let next_nodes = ref [] in + let ib = (get_some @@ PTree.get e code).entry in + let rec translate_btl_block ib k = + print_btl_inst stderr ib; + let rtli = + match ib with + | Bcond (cond, lr, BF (Bgoto s1), BF (Bgoto s2), info, ni) -> + next_nodes := s1 :: s2 :: !next_nodes; + Some (Icond (cond, lr, s1, s2, info), ni) + | Bcond (cond, lr, BF (Bgoto s1), ib2, info, ni) -> + assert (info = Some false); + next_nodes := s1 :: !next_nodes; + translate_btl_block ib2 None; + Some (Icond (cond, lr, s1, get_nn ib2, info), ni) + (* TODO gourdinl remove dynamic check *) + | Bcond (_, _, _, _, _, _) -> + failwith "translate_function: invalid Bcond" + | Bseq (ib1, ib2) -> + (* TODO gourdinl remove dynamic check *) + assert (is_atom ib1); + translate_btl_block ib1 (Some ib2); + translate_btl_block ib2 None; + None + | Bnop ni -> Some (Inop (get_nn (get_some k)), ni) + | Bop (op, lr, rd, ni) -> + Some (Iop (op, lr, rd, get_nn (get_some k)), ni) + | Bload (trap, chk, addr, lr, rd, ni) -> + Some (Iload (trap, chk, addr, lr, rd, get_nn (get_some k)), ni) + | Bstore (chk, addr, lr, rd, ni) -> + Some (Istore (chk, addr, lr, rd, get_nn (get_some k)), ni) + | BF (Bcall (sign, fn, lr, rd, s, ni)) -> + next_nodes := s :: !next_nodes; + Some (Icall (sign, fn, lr, rd, s), ni) + | BF (Btailcall (sign, fn, lr, ni)) -> + Some (Itailcall (sign, fn, lr), ni) + | BF (Bbuiltin (ef, lr, rd, s, ni)) -> + next_nodes := s :: !next_nodes; + Some (Ibuiltin (ef, lr, rd, s), ni) + | BF (Bjumptable (arg, tbl, ni)) -> Some (Ijumptable (arg, tbl), ni) + | BF (Breturn (oreg, ni)) -> Some (Ireturn oreg, ni) + | BF (Bgoto s) -> + next_nodes := s :: !next_nodes; + None + in + match rtli with + | Some (inst, ni) -> rtl_code := PTree.set ni inst !rtl_code + | None -> () + in + translate_btl_block ib None; + let succs = !next_nodes in + List.iter build_rtl_tree succs) + in + build_rtl_tree entry; + !rtl_code + +let btl2rtl (f : BTL.coq_function) = + debug_flag := true; + let code = f.fn_code in + let entry = f.fn_entrypoint in + let rtl = translate_function code entry in + let dm = PTree.map (fun n i -> n) code in + debug "Entry %d\n" (P.to_int entry); + debug "RTL Code: "; + print_code rtl; + debug "Dupmap:\n"; + print_ptree print_pint dm; + debug "\n"; + debug_flag := false; + ((rtl, entry), dm) diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 08a77ae4..6c894b78 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -263,7 +263,7 @@ Proof. eapply plus_trans; eauto. - (* exec_cond *) inv MIB. - rename H10 into JOIN. (* is_join_opt opc1 opc2 opc *) + rename H11 into JOIN. (* is_join_opt opc1 opc2 opc *) destruct b; exploit IHIBIS; eauto. + (* taking ifso *) destruct ofin; simpl. @@ -305,14 +305,14 @@ Proof. erewrite <- preserv_fnstacksize; eauto. econstructor; eauto. - (* call *) - rename H7 into FIND. + rename H8 into FIND. exploit find_function_preserved; eauto. intros (fd' & FIND' & TRANSFU). eexists; split. eapply exec_Icall; eauto. apply function_sig_translated. assumption. repeat (econstructor; eauto). - (* tailcall *) - rename H2 into FIND. + rename H3 into FIND. exploit find_function_preserved; eauto. intros (fd' & FIND' & TRANSFU). eexists; split. eapply exec_Itailcall; eauto. diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index b461e3f1..4b14d28e 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -8,84 +8,80 @@ open PrintAST (* Printing of BTL code *) -let reg pp r = - fprintf pp "x%d" (P.to_int r) +let reg pp r = fprintf pp "x%d" (P.to_int r) let rec regs pp = function | [] -> () - | [r] -> reg pp r - | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl + | [ r ] -> reg pp r + | r1 :: rl -> fprintf pp "%a, %a" reg r1 regs rl let ros pp = function | Coq_inl r -> reg pp r | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s) -let print_succ pp s = - fprintf pp "\tsucc %d\n" (P.to_int s) +let print_succ pp s = fprintf pp "\tsucc %d\n" (P.to_int s) -let print_pref pp pref = - fprintf pp "%s" pref +let print_pref pp pref = fprintf pp "%s" pref let rec print_iblock pp is_rec pref ib = match ib with - | Bnop -> + | Bnop _ -> print_pref pp pref; fprintf pp "Bnop\n" - | Bop(op, args, res) -> + | Bop (op, args, res, _) -> print_pref pp pref; - fprintf pp "Bop: %a = %a\n" - reg res (PrintOp.print_operation reg) (op, args) - | Bload(trap, chunk, addr, args, dst) -> + fprintf pp "Bop: %a = %a\n" reg res + (PrintOp.print_operation reg) + (op, args) + | Bload (trap, chunk, addr, args, dst, _) -> print_pref pp pref; - fprintf pp "Bload: %a = %s[%a]%a\n" - reg dst (name_of_chunk chunk) - (PrintOp.print_addressing reg) (addr, args) - print_trapping_mode trap - | Bstore(chunk, addr, args, src) -> + fprintf pp "Bload: %a = %s[%a]%a\n" reg dst (name_of_chunk chunk) + (PrintOp.print_addressing reg) + (addr, args) print_trapping_mode trap + | Bstore (chunk, addr, args, src, _) -> print_pref pp pref; - fprintf pp "Bstore: %s[%a] = %a\n" - (name_of_chunk chunk) - (PrintOp.print_addressing reg) (addr, args) - reg src - | BF (Bcall(sg, fn, args, res, s)) -> + fprintf pp "Bstore: %s[%a] = %a\n" (name_of_chunk chunk) + (PrintOp.print_addressing reg) + (addr, args) reg src + | BF (Bcall (sg, fn, args, res, s, _)) -> print_pref pp pref; - fprintf pp "Bcall: %a = %a(%a)\n" - reg res ros fn regs args; + fprintf pp "Bcall: %a = %a(%a)\n" reg res ros fn regs args; print_succ pp s - | BF (Btailcall(sg, fn, args)) -> + | BF (Btailcall (sg, fn, args, _)) -> print_pref pp pref; - fprintf pp "Btailcall: %a(%a)\n" - ros fn regs args - | BF (Bbuiltin(ef, args, res, s)) -> + fprintf pp "Btailcall: %a(%a)\n" ros fn regs args + | BF (Bbuiltin (ef, args, res, s, _)) -> print_pref pp pref; - fprintf pp "Bbuiltin: %a = %s(%a)\n" - (print_builtin_res reg) res - (name_of_external ef) - (print_builtin_args reg) args; + fprintf pp "Bbuiltin: %a = %s(%a)\n" (print_builtin_res reg) res + (name_of_external ef) (print_builtin_args reg) args; print_succ pp s - | Bcond(cond, args, ib1, ib2, info) -> + | Bcond (cond, args, ib1, ib2, info, _) -> print_pref pp pref; fprintf pp "Bcond: (%a) (prediction: %s)\n" - (PrintOp.print_condition reg) (cond, args) - (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough"); - let pref' = pref ^ " " in - fprintf pp "%sifso = [\n" pref; - if is_rec then print_iblock pp is_rec pref' ib1 else fprintf pp "...\n"; - fprintf pp "%s]\n" pref; - fprintf pp "%sifnot = [\n" pref; - if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n"; - fprintf pp "%s]\n" pref - | BF (Bjumptable(arg, tbl)) -> + (PrintOp.print_condition reg) + (cond, args) + (match info with + | None -> "none" + | Some true -> "branch" + | Some false -> "fallthrough"); + let pref' = pref ^ " " in + fprintf pp "%sifso = [\n" pref; + if is_rec then print_iblock pp is_rec pref' ib1 else fprintf pp "...\n"; + fprintf pp "%s]\n" pref; + fprintf pp "%sifnot = [\n" pref; + if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n"; + fprintf pp "%s]\n" pref + | BF (Bjumptable (arg, tbl, _)) -> let tbl = Array.of_list tbl in print_pref pp pref; fprintf pp "Bjumptable: (%a)\n" reg arg; for i = 0 to Array.length tbl - 1 do fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i)) done - | BF (Breturn None) -> + | BF (Breturn (None, _)) -> print_pref pp pref; fprintf pp "Breturn\n" - | BF (Breturn (Some arg)) -> + | BF (Breturn (Some arg, _)) -> print_pref pp pref; fprintf pp "Breturn: %a\n" reg arg | BF (Bgoto s) -> @@ -97,28 +93,26 @@ let rec print_iblock pp is_rec pref ib = print_iblock pp is_rec pref ib2) else fprintf pp "Bseq...\n" -let print_btl_inst pp ib = - print_iblock pp false " " ib +let print_btl_inst pp ib = print_iblock pp false " " ib let print_btl_code pp btl is_rec = fprintf pp "\n"; - List.iter (fun (n,ibf) -> - fprintf pp "[BTL block %d]\n" (P.to_int n); - print_iblock pp is_rec " " ibf.entry; - fprintf pp "\n") - (PTree.elements btl); + List.iter + (fun (n, ibf) -> + fprintf pp "[BTL block %d]\n" (P.to_int n); + print_iblock pp is_rec " " ibf.entry; + fprintf pp "\n") + (PTree.elements btl); fprintf pp "\n" let print_function pp id f = fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; - let instrs = List.map (fun (n,i) -> i.entry) (PTree.elements f.fn_code) in + let instrs = List.map (fun (n, i) -> i.entry) (PTree.elements f.fn_code) in List.iter (print_iblock pp true "") instrs; fprintf pp "}\n\n" let print_globdef pp (id, gd) = - match gd with - | Gfun(Internal f) -> print_function pp id f - | _ -> () + match gd with Gfun (Internal f) -> print_function pp id f | _ -> () -let print_program pp (prog: BTL.program) = +let print_program pp (prog : BTL.program) = List.iter (print_globdef pp) prog.prog_defs diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index e14e0ab3..859bbf07 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -7,34 +7,36 @@ open PrintBTL open Camlcoq open AuxTools +let mk_ni n = n + let encaps_final inst osucc = match inst with | BF _ | Bcond _ -> inst | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc))) -let translate_inst inst is_final = +let translate_inst ni inst is_final = let osucc = ref None in let btli = match inst with | Inop s -> osucc := Some s; - Bnop + Bnop ni | Iop (op, lr, rd, s) -> osucc := Some s; - Bop (op, lr, rd) + Bop (op, lr, rd, ni) | Iload (trap, chk, addr, lr, rd, s) -> osucc := Some s; - Bload (trap, chk, addr, lr, rd) + Bload (trap, chk, addr, lr, rd, ni) | Istore (chk, addr, lr, rd, s) -> osucc := Some s; - Bstore (chk, addr, lr, rd) - | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s)) - | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr)) - | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s)) + Bstore (chk, addr, lr, rd, ni) + | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s, ni)) + | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr, ni)) + | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s, ni)) | Icond (cond, lr, ifso, ifnot, info) -> - Bcond (cond, lr, BF (Bgoto ifso), BF (Bgoto ifnot), info) - | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl)) - | Ireturn oreg -> BF (Breturn oreg) + Bcond (cond, lr, BF (Bgoto ifso), BF (Bgoto ifnot), info, ni) + | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl, ni)) + | Ireturn oreg -> BF (Breturn (oreg, ni)) in if is_final then encaps_final btli !osucc else btli @@ -50,6 +52,7 @@ let translate_function code entry join_points = let rec build_btl_block n = let inst = get_some @@ PTree.get n code in let psucc = predicted_successor inst in + let ni = mk_ni n in let succ = match psucc with | Some ps -> @@ -67,12 +70,12 @@ let translate_function code entry join_points = (* TODO gourdinl remove this *) assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; - Bcond (cond, lr, BF (Bgoto ifso), build_btl_block s, info) - | _ -> Bseq (translate_inst inst false, build_btl_block s)) + Bcond (cond, lr, BF (Bgoto ifso), build_btl_block s, info, ni) + | _ -> Bseq (translate_inst ni inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; next_nodes := !next_nodes @ successors_inst inst; - translate_inst inst true + translate_inst ni inst true in let ib = build_btl_block e in let succs = !next_nodes in @@ -84,12 +87,12 @@ let translate_function code entry join_points = !btl_code let print_dm dm = - List.iter (fun (n,n') -> - debug "%d -> %d" (P.to_int n) (P.to_int n'); - (*print_btl_inst stderr ib.entry;*) - debug "\n" - ) - (PTree.elements dm) + List.iter + (fun (n, n') -> + debug "%d -> %d" (P.to_int n) (P.to_int n'); + (*print_btl_inst stderr ib.entry;*) + debug "\n") + (PTree.elements dm) let rtl2btl (f : RTL.coq_function) = debug_flag := true; diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index f89ea5cf..6681d691 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -215,7 +215,7 @@ type defined below. Intuitively, each RTL step corresponds to either Fixpoint measure ib: nat := match ib with | Bseq ib1 ib2 - | Bcond _ _ ib1 ib2 _ => measure ib1 + measure ib2 + | Bcond _ _ ib1 ib2 _ _ => measure ib1 + measure ib2 | ib => 1 end. -- cgit From 2408dd1aecf8d8c3bfd3b24a65f7f57cf602cb10 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 20 May 2021 11:58:40 +0200 Subject: Changing to an opaq record in BTL info, this is a broken commit --- scheduling/BTL.v | 158 +++++++++++++++++++++++---------------------- scheduling/BTL_SEtheory.v | 14 ++-- scheduling/BTLaux.ml | 3 + scheduling/BTLtoRTLproof.v | 16 ++--- scheduling/PrintBTL.ml | 19 +++--- scheduling/RTLtoBTLaux.ml | 44 +++++++------ scheduling/RTLtoBTLproof.v | 12 ++-- 7 files changed, 141 insertions(+), 125 deletions(-) create mode 100644 scheduling/BTLaux.ml (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 44c25b75..ca727f82 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -23,29 +23,32 @@ 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 -}. +(* inst_info is a ghost record to provide instruction information through oracles *) +Parameter inst_info: Set. +Extract Constant inst_info => "BTLaux.inst_info". + +(* block_info is a ghost record to provide block information through oracles *) +Parameter block_info: Set. +Extract Constant block_info => "BTLaux.block_info". (** final instructions (that stops block execution) *) Inductive final: Type := | Bgoto (succ:exit) (** No operation -- just branch to [succ]. *) - | Breturn (res: option reg) (ni: node_info) + | Breturn (res: option reg) (** 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) (ni: node_info) + | Bcall (sig: signature) (fn: reg + ident) (args:list reg) (dest:reg) (succ:exit) (** 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) (ni: node_info) + | Btailcall (sig:signature) (fn: reg + ident) (args: list reg) (** 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) (ni: node_info) + | Bbuiltin (ef:external_function) (args:list (builtin_arg reg)) (dest:builtin_res reg) (succ:exit) (** 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) (ni: node_info) + | Bjumptable (arg:reg) (tbl:list exit) (** [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]. *) . @@ -53,16 +56,16 @@ Inductive final: Type := (* instruction block *) Inductive iblock: Type := (* final instructions that stops block execution *) - | BF (fi: final) + | BF (fi: final) (iinfo: inst_info) (* basic instructions that continues block execution, except when aborting *) - | Bnop (ni: node_info) (* nop instruction *) - | Bop (op:operation) (args:list reg) (dest:reg) (ni: node_info) + | Bnop (iinfo: inst_info) (* nop instruction *) + | Bop (op:operation) (args:list reg) (dest:reg) (iinfo: inst_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) (ni: node_info) + | Bload (trap:trapping_mode) (chunk:memory_chunk) (addr:addressing) (args:list reg) (dest:reg) (iinfo: inst_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) (ni: node_info) + | Bstore (chunk:memory_chunk) (addr:addressing) (args:list reg) (src:reg) (iinfo: inst_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. *) @@ -70,13 +73,13 @@ 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) (ni: node_info) + | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (iinfo: inst_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]. [info] is a ghost field there to provide information relative to branch prediction. *) . -Coercion BF: final >-> iblock. +Coercion BF: final >-> Funclass. (** NB: - a RTL [(Inop pc)] ending a branch of block is encoded by [(Bseq Bnop (Bgoto pc))]. @@ -88,7 +91,8 @@ Coercion BF: final >-> iblock. Record iblock_info := { entry: iblock; - input_regs: Regset.t (* extra liveness information ignored by BTL semantics *) + input_regs: Regset.t; (* extra liveness information ignored by BTL semantics *) + binfo: block_info (* Ghost field used in oracles *) }. Definition code: Type := PTree.t iblock_info. @@ -199,24 +203,24 @@ 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 ni: iblock_istep sp rs m (Bnop ni) rs m None - | exec_op rs m op args res v ni + | exec_final rs m fin iinfo: iblock_istep sp rs m (BF fin iinfo) rs m (Some fin) + | exec_nop rs m iinfo: iblock_istep sp rs m (Bnop iinfo) rs m None + | exec_op rs m op args res v iinfo (EVAL: eval_operation ge sp op rs##args m = Some 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 + : iblock_istep sp rs m (Bop op args res iinfo) (rs#res <- v) m None + | exec_load_TRAP rs m chunk addr args dst a v iinfo (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 ni) (rs#dst <- v) m None + : iblock_istep sp rs m (Bload TRAP chunk addr args dst iinfo) (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' ni + | exec_store rs m chunk addr args src a m' iinfo (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 ni) rs m' None + : iblock_istep sp rs m (Bstore chunk addr args src iinfo) 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) @@ -224,10 +228,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 ni + | exec_cond rs m cond args ifso ifnot b rs' m' ofin iinfo (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 ni) rs' m' ofin + : iblock_istep sp rs m (Bcond cond args ifso ifnot iinfo) rs' m' ofin . Local Hint Constructors iblock_istep: core. @@ -236,7 +240,7 @@ Lemma [iblock_istep_run_equiv] below provides a proof that "relation" [iblock_is *) Fixpoint iblock_istep_run sp ib rs m: option outcome := match ib with - | BF fin => + | BF fin _ => Some {| _rs := rs; _m := m; _fin := Some fin |} (* basic instructions *) | Bnop _ => @@ -261,7 +265,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. @@ -280,32 +284,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' ni: + | exec_Breturn or stk m': sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - final_step stack f sp rs m (Breturn or ni) + final_step stack f sp rs m (Breturn or) E0 (Returnstate stack (regmap_optget or Vundef rs) m') - | exec_Bcall sig ros args res pc' fd ni: + | exec_Bcall sig ros args res pc' fd: find_function ros rs = Some fd -> funsig fd = sig -> - final_step stack f sp rs m (Bcall sig ros args res pc' ni) + final_step stack f sp rs m (Bcall sig ros args res pc') E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) - | exec_Btailcall sig ros args stk m' fd ni: + | exec_Btailcall sig ros args stk m' fd: 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 ni) + final_step stack f sp rs m (Btailcall sig ros args) E0 (Callstate stack fd rs##args m') - | exec_Bbuiltin ef args res pc' vargs t vres m' ni: + | exec_Bbuiltin 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' -> - final_step stack f sp rs m (Bbuiltin ef args res pc' ni) + final_step stack f sp rs m (Bbuiltin ef args res pc') t (State stack f sp pc' (regmap_setres res vres rs) m') - | exec_Bjumptable arg tbl n pc' ni: + | exec_Bjumptable arg tbl n pc': rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - final_step stack f sp rs m (Bjumptable arg tbl ni) + final_step stack f sp rs m (Bjumptable arg tbl) E0 (State stack f sp pc' rs m) . @@ -393,16 +397,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 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: + | 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: list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' -> - match_final_inst dupmap (Bjumptable r ln ni) (Ijumptable r ln') + match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln') . Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := @@ -417,25 +421,25 @@ Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := - if [opc] is [Some pc'], this means that all branches of the block that do not exit, join on [pc']. *) Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> iblock -> (option node) -> Prop := - | mib_BF isfst fi pc i: + | mib_BF isfst fi pc i iinfo: cfg!pc = Some i -> match_final_inst dupmap fi i -> - match_iblock dupmap cfg isfst pc (BF fi) None - | mib_nop isfst pc pc' ni: + match_iblock dupmap cfg isfst pc (BF fi iinfo) None + | mib_nop isfst pc pc' iinfo: cfg!pc = Some (Inop pc') -> - match_iblock dupmap cfg isfst pc (Bnop ni) (Some pc') - | mib_op isfst pc pc' op lr r ni: + match_iblock dupmap cfg isfst pc (Bnop iinfo) (Some pc') + | mib_op isfst pc pc' op lr r iinfo: cfg!pc = Some (Iop op lr r pc') -> - match_iblock dupmap cfg isfst pc (Bop op lr r ni) (Some pc') - | mib_load isfst pc pc' m a lr r ni: + match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc') + | mib_load isfst pc pc' m a lr r iinfo: cfg!pc = Some (Iload TRAP m a lr r pc') -> - 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: + match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r iinfo) (Some pc') + | mib_store isfst pc pc' m a lr r iinfo: cfg!pc = Some (Istore m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bstore m a lr r ni) (Some pc') - | mib_exit pc pc': + match_iblock dupmap cfg isfst pc (Bstore m a lr r iinfo) (Some pc') + | mib_exit pc pc' iinfo: dupmap!pc = (Some pc') -> - match_iblock dupmap cfg false pc' (Bgoto pc) None + match_iblock dupmap cfg false pc' (BF (Bgoto pc) iinfo) None (* NB: on RTL side, we exit the block by a "basic" instruction (or Icond). Thus some step should have been executed before [pc'] in the RTL code *) | mib_seq_Some isfst b1 b2 pc1 pc2 opc: @@ -450,12 +454,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' ni: - cfg!pc = Some (Icond c lr pcso pcnot i') -> + | mib_cond isfst c lr bso bnot pcso pcnot pc opc1 opc2 opc i iinfo: + 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 ni) opc + match_iblock dupmap cfg isfst pc (Bcond c lr bso bnot iinfo) opc . Definition match_cfg dupmap (cfg: code) (cfg':RTL.code): Prop := @@ -528,21 +532,21 @@ Global Opaque builtin_res_eq_pos. Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := match ib with - | BF fi => + | BF fi _ => match fi with | Bgoto pc1 => do u <- verify_is_copy dupmap pc1 pc; 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; @@ -556,7 +560,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 @@ -567,7 +571,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; @@ -579,7 +583,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'; @@ -641,7 +645,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 _ => match cfg!pc with | Some (Icond c' lr' pcso pcnot i') => if (list_eq_dec Pos.eq_dec lr lr') then @@ -793,13 +797,13 @@ Definition verify_function dupmap f f' : res unit := Definition is_goto (ib: iblock): bool := match ib with - | Bgoto _ => true + | BF (Bgoto _) _ => true | _ => false end. Definition is_atom (ib: iblock): bool := match ib with - | Bseq _ _ | Bcond _ _ _ _ _ _ => false + | Bseq _ _ | Bcond _ _ _ _ _ => false | _ => true end. @@ -809,10 +813,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 ni: + | exp_Bcond cond args ib1 ib2 iinfo: is_expand ib1 -> is_expand ib2 -> - is_expand (Bcond cond args ib1 ib2 i ni) + is_expand (Bcond cond args ib1 ib2 iinfo) | exp_others ib: is_atom ib = true -> is_expand ib @@ -822,9 +826,9 @@ 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 ni => - Bcond cond args (expand ib1 k) (expand ib2 k) i ni - | BF fin => fin + | Bcond cond args ib1 ib2 iinfo => + Bcond cond args (expand ib1 k) (expand ib2 k) iinfo + | BF fin iinfo => BF fin iinfo | ib => match k with | None => ib diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 94d299e7..cd69cd37 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -557,21 +557,21 @@ Qed. Definition sexec_final_sfv (i: final) (sis: sistate): sfval := match i with | Bgoto pc => Sgoto pc - | Bcall sig ros args res pc _ => + | Bcall sig ros args res pc => let svos := sum_left_map sis.(si_sreg) ros in let sargs := list_sval_inj (List.map sis.(si_sreg) args) in Scall sig svos sargs res pc - | Btailcall sig ros args _ => + | Btailcall sig ros args => let svos := sum_left_map sis.(si_sreg) ros in let sargs := list_sval_inj (List.map sis.(si_sreg) args) in Stailcall sig svos sargs - | Bbuiltin ef args res pc _ => + | Bbuiltin ef args res pc => let sargs := List.map (builtin_arg_map sis.(si_sreg)) args in Sbuiltin ef sargs res pc - | Breturn or _ => + | Breturn or => let sor := SOME r <- or IN Some (sis.(si_sreg) r) in Sreturn sor - | Bjumptable reg tbl _ => + | Bjumptable reg tbl => let sv := sis.(si_sreg) reg in Sjumptable sv tbl end. @@ -770,7 +770,7 @@ pour représenter l'ensemble des chemins. Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := match ib with - | BF fin => Sfinal sis (sexec_final_sfv fin sis) + | BF fin _ => Sfinal sis (sexec_final_sfv fin sis) (* basic instructions *) | Bnop _ => k sis | Bop op args res _ => k (sexec_op op args res sis) @@ -780,7 +780,7 @@ Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := (* composed instructions *) | Bseq ib1 ib2 => sexec_rec ib1 sis (fun sis2 => sexec_rec ib2 sis2 k) - | Bcond cond args ifso ifnot _ _ => + | Bcond cond args ifso ifnot _ => let args := list_sval_inj (List.map sis.(si_sreg) args) in let ifso := sexec_rec ifso sis k in let ifnot := sexec_rec ifnot sis k in diff --git a/scheduling/BTLaux.ml b/scheduling/BTLaux.ml new file mode 100644 index 00000000..e8e355b1 --- /dev/null +++ b/scheduling/BTLaux.ml @@ -0,0 +1,3 @@ +type inst_info = { mutable inumb : int; mutable pcond : bool option } + +type block_info = { mutable bnumb : int } diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 6c894b78..9b37d8fa 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -221,7 +221,7 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin match ofin with | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => - exists isfst' pc1, match_iblock dupmap (RTL.fn_code f') isfst' pc1 fin None + exists isfst' pc1 iinfo, match_iblock dupmap (RTL.fn_code f') isfst' pc1 (BF fin iinfo) None /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) end. Proof. @@ -255,7 +255,7 @@ Proof. intros (pc1 & EQpc1 & STEP1); inv EQpc1. exploit IHIBIS2; eauto. destruct ofin; simpl. - + intros (ifst2 & pc2 & M2 & STEP2). + + intros (ifst2 & pc2 & iinfo & M2 & STEP2). repeat eexists; eauto. eapply css_plus_trans; eauto. + intros (pc2 & EQpc2 & STEP2); inv EQpc2. @@ -263,12 +263,12 @@ Proof. eapply plus_trans; eauto. - (* exec_cond *) inv MIB. - rename H11 into JOIN. (* is_join_opt opc1 opc2 opc *) + rename H10 into JOIN. (* is_join_opt opc1 opc2 opc *) destruct b; exploit IHIBIS; eauto. + (* taking ifso *) destruct ofin; simpl. * (* ofin is Some final *) - intros (isfst' & pc1 & MI & STAR). + intros (isfst' & pc1 & iinfo' & MI & STAR). repeat eexists; eauto. eapply css_plus_trans; eauto. * (* ofin is None *) @@ -279,7 +279,7 @@ Proof. + (* taking ifnot *) destruct ofin; simpl. * (* ofin is Some final *) - intros (isfst' & pc1 & MI & STAR). + intros (isfst' & pc1 & iinfo' & MI & STAR). repeat eexists; eauto. eapply css_plus_trans; eauto. * (* ofin is None *) @@ -305,14 +305,14 @@ Proof. erewrite <- preserv_fnstacksize; eauto. econstructor; eauto. - (* call *) - rename H8 into FIND. + rename H7 into FIND. exploit find_function_preserved; eauto. intros (fd' & FIND' & TRANSFU). eexists; split. eapply exec_Icall; eauto. apply function_sig_translated. assumption. repeat (econstructor; eauto). - (* tailcall *) - rename H3 into FIND. + rename H2 into FIND. exploit find_function_preserved; eauto. intros (fd' & FIND' & TRANSFU). eexists; split. eapply exec_Itailcall; eauto. @@ -342,7 +342,7 @@ Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fi : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. - simpl. intros (isfst' & pc1 & MI & STAR). clear IBIS MIB. + simpl. intros (isfst' & pc1 & iinfo & MI & STAR). clear IBIS MIB. inv MI. - (* final inst except goto *) exploit final_simu_except_goto; eauto. diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 4b14d28e..23ad91f6 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -5,6 +5,7 @@ open Maps open AST open BTL open PrintAST +open BTLaux (* Printing of BTL code *) @@ -43,24 +44,24 @@ let rec print_iblock pp is_rec pref ib = fprintf pp "Bstore: %s[%a] = %a\n" (name_of_chunk chunk) (PrintOp.print_addressing reg) (addr, args) reg src - | BF (Bcall (sg, fn, args, res, s, _)) -> + | BF (Bcall (sg, fn, args, res, s), _) -> print_pref pp pref; fprintf pp "Bcall: %a = %a(%a)\n" reg res ros fn regs args; print_succ pp s - | BF (Btailcall (sg, fn, args, _)) -> + | BF (Btailcall (sg, fn, args), _) -> print_pref pp pref; fprintf pp "Btailcall: %a(%a)\n" ros fn regs args - | BF (Bbuiltin (ef, args, res, s, _)) -> + | BF (Bbuiltin (ef, args, res, s), _) -> print_pref pp pref; fprintf pp "Bbuiltin: %a = %s(%a)\n" (print_builtin_res reg) res (name_of_external ef) (print_builtin_args reg) args; print_succ pp s - | Bcond (cond, args, ib1, ib2, info, _) -> + | Bcond (cond, args, ib1, ib2, iinfo) -> print_pref pp pref; fprintf pp "Bcond: (%a) (prediction: %s)\n" (PrintOp.print_condition reg) (cond, args) - (match info with + (match iinfo.pcond with | None -> "none" | Some true -> "branch" | Some false -> "fallthrough"); @@ -71,20 +72,20 @@ let rec print_iblock pp is_rec pref ib = fprintf pp "%sifnot = [\n" pref; if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n"; fprintf pp "%s]\n" pref - | BF (Bjumptable (arg, tbl, _)) -> + | BF (Bjumptable (arg, tbl), _) -> let tbl = Array.of_list tbl in print_pref pp pref; fprintf pp "Bjumptable: (%a)\n" reg arg; for i = 0 to Array.length tbl - 1 do fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i)) done - | BF (Breturn (None, _)) -> + | BF (Breturn None, _) -> print_pref pp pref; fprintf pp "Breturn\n" - | BF (Breturn (Some arg, _)) -> + | BF (Breturn (Some arg), _) -> print_pref pp pref; fprintf pp "Breturn: %a\n" reg arg - | BF (Bgoto s) -> + | BF (Bgoto s, _) -> print_pref pp pref; fprintf pp "Bgoto: %d\n" (P.to_int s) | Bseq (ib1, ib2) -> diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 859bbf07..d4fd2643 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -6,37 +6,43 @@ open DebugPrint open PrintBTL open Camlcoq open AuxTools +open BTLaux -let mk_ni n = n +let undef_node = -1 +let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond } +let def_iinfo = { inumb = undef_node; pcond = None } + +let mk_binfo _bnumb = { bnumb = _bnumb } let encaps_final inst osucc = match inst with | BF _ | Bcond _ -> inst - | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc))) + | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo)) -let translate_inst ni inst is_final = +let translate_inst iinfo inst is_final = let osucc = ref None in let btli = match inst with | Inop s -> osucc := Some s; - Bnop ni + Bnop iinfo | Iop (op, lr, rd, s) -> osucc := Some s; - Bop (op, lr, rd, ni) + Bop (op, lr, rd, iinfo) | Iload (trap, chk, addr, lr, rd, s) -> osucc := Some s; - Bload (trap, chk, addr, lr, rd, ni) + Bload (trap, chk, addr, lr, rd, iinfo) | Istore (chk, addr, lr, rd, s) -> osucc := Some s; - Bstore (chk, addr, lr, rd, ni) - | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s, ni)) - | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr, ni)) - | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s, ni)) + Bstore (chk, addr, lr, rd, iinfo) + | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s), iinfo) + | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr), iinfo) + | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s), iinfo) | Icond (cond, lr, ifso, ifnot, info) -> - Bcond (cond, lr, BF (Bgoto ifso), BF (Bgoto ifnot), info, ni) - | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl, ni)) - | Ireturn oreg -> BF (Breturn (oreg, ni)) + iinfo.pcond <- info; + Bcond (cond, lr, BF (Bgoto ifso, def_iinfo), BF (Bgoto ifnot, def_iinfo), iinfo) + | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl), iinfo) + | Ireturn oreg -> BF (Breturn (oreg), iinfo) in if is_final then encaps_final btli !osucc else btli @@ -52,7 +58,7 @@ let translate_function code entry join_points = let rec build_btl_block n = let inst = get_some @@ PTree.get n code in let psucc = predicted_successor inst in - let ni = mk_ni n in + let iinfo = mk_iinfo (p2i n) None in let succ = match psucc with | Some ps -> @@ -70,16 +76,18 @@ let translate_function code entry join_points = (* TODO gourdinl remove this *) assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; - Bcond (cond, lr, BF (Bgoto ifso), build_btl_block s, info, ni) - | _ -> Bseq (translate_inst ni inst false, build_btl_block s)) + iinfo.pcond <- info; + Bcond (cond, lr, BF (Bgoto ifso, def_iinfo), build_btl_block s, iinfo) + | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; next_nodes := !next_nodes @ successors_inst inst; - translate_inst ni inst true + translate_inst iinfo inst true in let ib = build_btl_block e in let succs = !next_nodes in - let ibf = { entry = ib; input_regs = Regset.empty } in + let bi = mk_binfo (p2i e) in + let ibf = { entry = ib; input_regs = Regset.empty; binfo = bi } in btl_code := PTree.set e ibf !btl_code; List.iter build_btl_tree succs) in diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 6681d691..7ad1472b 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -215,7 +215,7 @@ type defined below. Intuitively, each RTL step corresponds to either Fixpoint measure ib: nat := match ib with | Bseq ib1 ib2 - | Bcond _ _ ib1 ib2 _ _ => measure ib1 + measure ib2 + | Bcond _ _ ib1 ib2 _ => measure ib1 + measure ib2 | ib => 1 end. @@ -238,7 +238,7 @@ Lemma entry_isnt_goto dupmap f pc ib: Proof. intros. destruct (entry ib); trivial. - destruct fi; trivial. inv H. inv H4. + destruct fi; trivial. inv H. inv H5. Qed. Lemma expand_entry_isnt_goto dupmap f pc ib: @@ -248,7 +248,7 @@ Proof. destruct (is_goto (expand (entry ib) None))eqn:EQG; destruct (expand (entry ib) None); try destruct fi; try discriminate; trivial. - intros; inv H; inv H4. + intros; inv H; inv H5. Qed. Lemma list_nth_z_rev_dupmap: @@ -499,10 +499,10 @@ Proof. destruct (is_goto ib2) eqn:GT. destruct ib2; try destruct fi; try discriminate. - (* Bgoto *) - inv MSS2. inversion MIB; subst; try inv H3. - remember H0 as ODUPLIC; clear HeqODUPLIC. + inv MSS2. inversion MIB; subst; try inv H4. + remember H2 as ODUPLIC; clear HeqODUPLIC. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - apply DMC in H0 as [ib [FNC MI]]; clear DMC. + apply DMC in H2 as [ib [FNC MI]]; clear DMC. eexists; left; eexists; split. + repeat econstructor; eauto. apply iblock_istep_run_equiv in BTL_RUN; eauto. -- cgit From 8dc70c68f241e1397f2c65981202742fb0ff75a3 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 20 May 2021 15:00:07 +0200 Subject: correction de l'idee de la Functional semantics --- scheduling/BTLroadmap.md | 43 +++++++++++++++++-------------------------- 1 file changed, 17 insertions(+), 26 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index bd70e273..39ea12d0 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -73,40 +73,36 @@ L'approche suivie pour réaliser la simulation modulo liveness dans RTLpath est compliquée à adapter sur BTL. En effet, un état symbolique RTLpath correspond à un état symbolique BTL très particulier: toutes les "feuilles" (les `Sfinal`) sont des `Sgoto` -sauf éventuellement une. Or, dans RTLpath, le traitement de l'info de liveness sur -cette feuille particulière est très adhoc et laborieux (cf. `pre_output_regs`). On -n'a pas envie de généraliser cette usine à gaz. +sauf éventuellement une. Or, dans RTLpath, le traitement de l'info de +liveness sur cette feuille particulière est très adhoc et laborieux +(cf. le traitement de `pre_output_regs` dans RTLpathScheduler, etc). +On n'a pas envie de généraliser cette usine à gaz. On cherche donc une façon plus abstraite de faire... On a l'idée de coder la "simulation modulo liveness" dans une "simulation less_undef". Ça peut rendre le cadre du test de simulation plus simple et plus général. -L'idée de départ: "Extends BTL with the possibility of destroying a -set of registers at each exit (a destroyed register is simply set to -Vundef)" en prouvant une simulation "less_undef" n'est pas assez -générale! Ça n'autorise pas à introduire de "nouveaux" registres dans -le bloc transformé, juste à donner n'importe quelle valeur aux -registres "hors-liveout". - -**Idée corrigée** à côté de la sémantique "à la RTL" pour BTL, on -définit une nouvelle sémantique (basée sur la même sémantique à grand +**Idée** à côté de la sémantique "à la RTL" pour BTL, on pourrait +définir une nouvelle sémantique (basée sur la même sémantique à grand pas dans les blocs), où c'est juste "l'entrée dans un bloc" qui change -de sémantique. Intuitivement, cette nouvelle sémantique considère +de sémantique. Intuitivement, cette nouvelle sémantique considèrerait chaque bloc comme une sorte de fonction paramétrée par les `input_regs` et appelée uniquement en "tailcall" (via les "goto"). C'est ce qu'on va appeler la "functional semantics" de BTL (l'autre étant appelée qqchose comme "CFG semantics" ?). Autrement dit, sur l'entrée dans un bloc pour un état (rs,m), on -remplace juste cet état par un nouvel état (rs0,m) où dans rs0, tous les registres -sont initialisés à Vundef sauf ceux du `input_regs` qui sont -initialisés comme une copie de "rs". In fine, on applique le -"iblock_step" sur cet état (rs0,m). +commence par mettre à Vundef tous les registres qui ne sont pas dans l'`input_regs`. + +**NOTE** cette idée de voir les blocs comme des "fonctions" correpond +bien à la représentation "SSA" à la Appel/MLIR. Donc cette sémantique +peut servir de base pour un support de formes SSA (partielles ou +totales) dans BTL. Pour l'encodage de SSA, il suffira d'étendre le +mécanisme d'initialisation du "rs0" d'un bloc avec un passage de +paramètres. -**NOTE** cette idée de voir les blocs comme des "fonctions" correpond bien à la représentation "SSA" à la Appel/MLIR. -Donc cette sémantique peut servir de base pour un support de formes SSA (partielles ou totales) dans BTL. -Pour l'encodage de SSA, il suffira d'étendre le mécanisme d'initialisation du "rs0" d'un bloc avec un passage de paramètres. +En fait, pour le test d'exécution symbolique, il semble plus simple de mettre les Vundef à la fin de la transition... **IMPLEM PLAN** @@ -125,16 +121,11 @@ We will implement a "basic" (e.g without rewriting rules) "less_undef" simulatio **IMPLEM NOTE** - use Continuation Passing Style for an easy recursive generation of "all execution paths". -- pour implementer la "functional semantics", il faut changer - légèrement la sémantique du map qui associe une valeur symbolique à - chaque registre. En RTLpath, un registre "absent" de la map était considéré comme positionné à sa valeur initial. - Là, il faudra considérer qu'il vaut Vundef (il y a une valeur symbolique spéciale à ajouter !) **CURRENT STATUS** - model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms. -- next step: high-level specification of "symbolic simulation" + preservation proof w.r.t BTL.FunctionalSemantics -(dans une première étape, on peut laisser le "less_undef" de côté: ce sera facile à ajouter quand le cadre sera en place). +- next step: high-level specification of "symbolic simulation" + preservation proof w.r.t BTL.FunctionalSemantics. ## Port of RTLpath optimizations to BTL -- cgit From bc6129876ffc6f0323752908f5de12bb5c5a7c74 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 20 May 2021 17:28:55 +0200 Subject: working oracles (no renumber for now) --- scheduling/BTLaux.ml | 2 +- scheduling/BTLrenumber.ml | 50 +++++++++++++++++++++ scheduling/BTLtoRTLaux.ml | 108 +++++++++++++++++++++++++--------------------- scheduling/PrintBTL.ml | 23 +++++----- scheduling/RTLtoBTLaux.ml | 40 ++++++++--------- 5 files changed, 144 insertions(+), 79 deletions(-) create mode 100644 scheduling/BTLrenumber.ml (limited to 'scheduling') diff --git a/scheduling/BTLaux.ml b/scheduling/BTLaux.ml index e8e355b1..863afdf0 100644 --- a/scheduling/BTLaux.ml +++ b/scheduling/BTLaux.ml @@ -1,3 +1,3 @@ type inst_info = { mutable inumb : int; mutable pcond : bool option } -type block_info = { mutable bnumb : int } +type block_info = { mutable bnumb : int; mutable visited: bool } diff --git a/scheduling/BTLrenumber.ml b/scheduling/BTLrenumber.ml new file mode 100644 index 00000000..dd6baa89 --- /dev/null +++ b/scheduling/BTLrenumber.ml @@ -0,0 +1,50 @@ +(* XXX uncomment +open !BTL +open DebugPrint +open AuxTools +open Maps*) + +(** A quick note on the BTL renumber algorithm + This is a simple first version where we try to reuse the entry numbering from RTL. + In the futur, it would be great to implement a strongly connected components partitionning. + The numbering is done by a postorder traversal. + TODO gourdinl : note perso + - faire un comptage global du nombre d'instructions rtl à générer + - utiliser la structure d'info créée lors de la génération pour tenter de préserver au max + la relation d'ordre, avec une heuristique genre (1 + max des succs) pour l'insertion + - quand il y a un branchement conditionnel, privilégier le parcour du fils gauche en premier + (afin d'avoir de plus petits numéros à gauche) +*) +(* +let rec get_max_rtl_ninsts code entry = +let rec postorder_blk ib = +*) + +(* XXX uncomment +let postorder code entry = + let renumbered_code = ref PTree.empty in + let rec renum_ibf e = + let ibf = get_some @@ PTree.get e code in + if ibf.binfo.visited then () + else ( + ibf.binfo.visited <- true; + let next_nodes = ref [] in + let ib = ibf.entry in + let rec renum_iblock ib = + match ib with + in + let rib = renum_iblock ib in + ibf.entry <- rib; + renumbered_code := PTree.set e ibf !renumbered_code; + let succs = !next_nodes in + List.iter renum_ibf succs) + in + renum_ibf entry + +let renumber (f: BTL.coq_function) = + debug_flag := true; + let code = f.fn_code in + let entry = f.fn_entrypoint in + let renumbered_code = postorder code entry in + debug_flag := false; + renumbered_code*) diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index dd7ba4c7..b4833b2c 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -4,50 +4,56 @@ open RTL open Camlcoq open AuxTools open DebugPrint -open PrintBTL +open BTLaux let is_atom = function - | Bseq (_, _) | Bcond (_, _, _, _, _, _) -> false + | Bseq (_, _) | Bcond (_, _, _, _, _) -> false | _ -> true -let rec get_nn = function - | Bnop ni - | Bop (_, _, _, ni) - | Bload (_, _, _, _, _, ni) - | Bstore (_, _, _, _, ni) - | Bcond (_, _, _, _, _, ni) - | BF (Breturn (_, ni)) - | BF (Bcall (_, _, _, _, _, ni)) - | BF (Btailcall (_, _, _, ni)) - | BF (Bbuiltin (_, _, _, _, ni)) - | BF (Bjumptable (_, _, ni)) -> - ni - | BF (Bgoto s) -> s - | Bseq (ib1, _) -> get_nn ib1 +let get_inumb iinfo = P.of_int iinfo.inumb + +let rec get_ib_num = function + | BF (Bgoto s, _) -> s + | Bnop iinfo + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) + | Bcond (_, _, _, _, iinfo) + | BF (_, iinfo) -> + get_inumb iinfo + | Bseq (ib1, _) -> get_ib_num ib1 let translate_function code entry = - let reached = ref (PTree.map (fun n i -> false) code) in let rtl_code = ref PTree.empty in + let code = + PTree.map + (fun n ibf -> + ibf.binfo.visited <- false; + ibf) + code + in let rec build_rtl_tree e = - if get_some @@ PTree.get e !reached then () + let ibf = get_some @@ PTree.get e code in + if ibf.binfo.visited then () else ( - reached := PTree.set e true !reached; + ibf.binfo.visited <- true; + let ib = ibf.entry in let next_nodes = ref [] in - let ib = (get_some @@ PTree.get e code).entry in let rec translate_btl_block ib k = - print_btl_inst stderr ib; let rtli = match ib with - | Bcond (cond, lr, BF (Bgoto s1), BF (Bgoto s2), info, ni) -> + | Bcond (cond, lr, BF (Bgoto s1, _), BF (Bgoto s2, _), iinfo) -> next_nodes := s1 :: s2 :: !next_nodes; - Some (Icond (cond, lr, s1, s2, info), ni) - | Bcond (cond, lr, BF (Bgoto s1), ib2, info, ni) -> - assert (info = Some false); + Some (Icond (cond, lr, s1, s2, iinfo.pcond), get_inumb iinfo) + | Bcond (cond, lr, BF (Bgoto s1, _), ib2, iinfo) -> + assert (iinfo.pcond = Some false); next_nodes := s1 :: !next_nodes; translate_btl_block ib2 None; - Some (Icond (cond, lr, s1, get_nn ib2, info), ni) + Some + ( Icond (cond, lr, s1, get_ib_num ib2, iinfo.pcond), + get_inumb iinfo ) (* TODO gourdinl remove dynamic check *) - | Bcond (_, _, _, _, _, _) -> + | Bcond (_, _, _, _, _) -> failwith "translate_function: invalid Bcond" | Bseq (ib1, ib2) -> (* TODO gourdinl remove dynamic check *) @@ -55,29 +61,35 @@ let translate_function code entry = translate_btl_block ib1 (Some ib2); translate_btl_block ib2 None; None - | Bnop ni -> Some (Inop (get_nn (get_some k)), ni) - | Bop (op, lr, rd, ni) -> - Some (Iop (op, lr, rd, get_nn (get_some k)), ni) - | Bload (trap, chk, addr, lr, rd, ni) -> - Some (Iload (trap, chk, addr, lr, rd, get_nn (get_some k)), ni) - | Bstore (chk, addr, lr, rd, ni) -> - Some (Istore (chk, addr, lr, rd, get_nn (get_some k)), ni) - | BF (Bcall (sign, fn, lr, rd, s, ni)) -> + | Bnop iinfo -> Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) + | Bop (op, lr, rd, iinfo) -> + Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo) + | Bload (trap, chk, addr, lr, rd, iinfo) -> + Some + ( Iload (trap, chk, addr, lr, rd, get_ib_num (get_some k)), + get_inumb iinfo ) + | Bstore (chk, addr, lr, rd, iinfo) -> + Some + ( Istore (chk, addr, lr, rd, get_ib_num (get_some k)), + get_inumb iinfo ) + | BF (Bcall (sign, fn, lr, rd, s), iinfo) -> next_nodes := s :: !next_nodes; - Some (Icall (sign, fn, lr, rd, s), ni) - | BF (Btailcall (sign, fn, lr, ni)) -> - Some (Itailcall (sign, fn, lr), ni) - | BF (Bbuiltin (ef, lr, rd, s, ni)) -> + Some (Icall (sign, fn, lr, rd, s), get_inumb iinfo) + | BF (Btailcall (sign, fn, lr), iinfo) -> + Some (Itailcall (sign, fn, lr), get_inumb iinfo) + | BF (Bbuiltin (ef, lr, rd, s), iinfo) -> next_nodes := s :: !next_nodes; - Some (Ibuiltin (ef, lr, rd, s), ni) - | BF (Bjumptable (arg, tbl, ni)) -> Some (Ijumptable (arg, tbl), ni) - | BF (Breturn (oreg, ni)) -> Some (Ireturn oreg, ni) - | BF (Bgoto s) -> + Some (Ibuiltin (ef, lr, rd, s), get_inumb iinfo) + | BF (Bjumptable (arg, tbl), iinfo) -> + next_nodes := !next_nodes @ tbl; + Some (Ijumptable (arg, tbl), get_inumb iinfo) + | BF (Breturn oreg, iinfo) -> Some (Ireturn oreg, get_inumb iinfo) + | BF (Bgoto s, iinfo) -> next_nodes := s :: !next_nodes; None in match rtli with - | Some (inst, ni) -> rtl_code := PTree.set ni inst !rtl_code + | Some (inst, inumb) -> rtl_code := PTree.set inumb inst !rtl_code | None -> () in translate_btl_block ib None; @@ -88,16 +100,16 @@ let translate_function code entry = !rtl_code let btl2rtl (f : BTL.coq_function) = - debug_flag := true; + (*debug_flag := true;*) let code = f.fn_code in let entry = f.fn_entrypoint in let rtl = translate_function code entry in let dm = PTree.map (fun n i -> n) code in - debug "Entry %d\n" (P.to_int entry); + debug "Entry %d\n" (p2i entry); debug "RTL Code: "; - print_code rtl; + (*print_code rtl;*) debug "Dupmap:\n"; print_ptree print_pint dm; debug "\n"; - debug_flag := false; + (*debug_flag := false;*) ((rtl, entry), dm) diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 23ad91f6..0ed3981d 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -2,10 +2,10 @@ open Printf open Camlcoq open Datatypes open Maps -open AST open BTL open PrintAST open BTLaux +open DebugPrint (* Printing of BTL code *) @@ -97,15 +97,18 @@ let rec print_iblock pp is_rec pref ib = let print_btl_inst pp ib = print_iblock pp false " " ib let print_btl_code pp btl is_rec = - fprintf pp "\n"; - List.iter - (fun (n, ibf) -> - fprintf pp "[BTL block %d]\n" (P.to_int n); - print_iblock pp is_rec " " ibf.entry; - fprintf pp "\n") - (PTree.elements btl); - fprintf pp "\n" + if !debug_flag then ( + fprintf pp "\n"; + List.iter + (fun (n, ibf) -> + fprintf pp "[BTL block %d]\n" (P.to_int n); + print_iblock pp is_rec " " ibf.entry; + fprintf pp "\n") + (PTree.elements btl); + fprintf pp "\n") + else () +(* TODO gourdinl remove or adapt this? let print_function pp id f = fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; let instrs = List.map (fun (n, i) -> i.entry) (PTree.elements f.fn_code) in @@ -116,4 +119,4 @@ let print_globdef pp (id, gd) = match gd with Gfun (Internal f) -> print_function pp id f | _ -> () let print_program pp (prog : BTL.program) = - List.iter (print_globdef pp) prog.prog_defs + List.iter (print_globdef pp) prog.prog_defs*) diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index d4fd2643..43556150 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -4,15 +4,16 @@ open BTL open Registers open DebugPrint open PrintBTL -open Camlcoq open AuxTools open BTLaux let undef_node = -1 + let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond } + let def_iinfo = { inumb = undef_node; pcond = None } -let mk_binfo _bnumb = { bnumb = _bnumb } +let mk_binfo _bnumb = { bnumb = _bnumb; visited = false } let encaps_final inst osucc = match inst with @@ -40,16 +41,20 @@ let translate_inst iinfo inst is_final = | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s), iinfo) | Icond (cond, lr, ifso, ifnot, info) -> iinfo.pcond <- info; - Bcond (cond, lr, BF (Bgoto ifso, def_iinfo), BF (Bgoto ifnot, def_iinfo), iinfo) + Bcond + ( cond, + lr, + BF (Bgoto ifso, def_iinfo), + BF (Bgoto ifnot, def_iinfo), + iinfo ) | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl), iinfo) - | Ireturn oreg -> BF (Breturn (oreg), iinfo) + | Ireturn oreg -> BF (Breturn oreg, iinfo) in if is_final then encaps_final btli !osucc else btli let translate_function code entry join_points = let reached = ref (PTree.map (fun n i -> false) code) in let btl_code = ref PTree.empty in - (* TODO gourdinl SEPARATE IN A BETTER WAY!! *) let rec build_btl_tree e = if get_some @@ PTree.get e !reached then () else ( @@ -63,21 +68,24 @@ let translate_function code entry join_points = match psucc with | Some ps -> if get_some @@ PTree.get ps join_points then ( - debug "IS JOIN PT\n"; None) else Some ps | None -> None in match succ with | Some s -> ( - debug "BLOCK CONTINUE\n"; match inst with | Icond (cond, lr, ifso, ifnot, info) -> (* TODO gourdinl remove this *) assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; iinfo.pcond <- info; - Bcond (cond, lr, BF (Bgoto ifso, def_iinfo), build_btl_block s, iinfo) + Bcond + ( cond, + lr, + BF (Bgoto ifso, def_iinfo), + build_btl_block s, + iinfo ) | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; @@ -94,25 +102,16 @@ let translate_function code entry join_points = build_btl_tree entry; !btl_code -let print_dm dm = - List.iter - (fun (n, n') -> - debug "%d -> %d" (P.to_int n) (P.to_int n'); - (*print_btl_inst stderr ib.entry;*) - debug "\n") - (PTree.elements dm) - let rtl2btl (f : RTL.coq_function) = - debug_flag := true; + (*debug_flag := true;*) let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in let btl = translate_function code entry join_points in let dm = PTree.map (fun n i -> n) btl in - debug "Entry %d\n" (P.to_int entry); + debug "Entry %d\n" (p2i entry); debug "RTL Code: "; - print_code code; - debug_flag := false; + (*print_code code;*) debug "BTL Code: "; print_btl_code stderr btl true; debug "Dupmap:\n"; @@ -120,4 +119,5 @@ let rtl2btl (f : RTL.coq_function) = debug "Join points: "; print_true_nodes join_points; debug "\n"; + (*debug_flag := false;*) ((btl, entry), dm) -- cgit From d858606e8400e6aab627f4aac5ec33ce9c2c80fe Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 20 May 2021 18:34:46 +0200 Subject: defines fsem (aka functional semantics) of BTL --- scheduling/BTL.v | 81 ++++++++++++++++++++++++++++++++++++++++++---- scheduling/BTL_SEtheory.v | 18 +++++++---- scheduling/BTLtoRTLproof.v | 6 ++-- scheduling/RTLtoBTLproof.v | 6 ++-- 4 files changed, 92 insertions(+), 19 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 10a000a8..cb60fed1 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -83,7 +83,7 @@ Coercion BF: final >-> iblock. Record iblock_info := { entry: iblock; - input_regs: Regset.t (* extra liveness information ignored by BTL semantics *) + input_regs: Regset.t (* extra liveness information for BTL functional semantics *) }. Definition code: Type := PTree.t iblock_info. @@ -161,6 +161,9 @@ Record outcome := out { Section RELSEM. +(* final_step is parametrized by a function to transfer regset on each exit *) +Variable tr_exit: function -> exit -> regset -> regset. + Variable ge: genv. Definition find_function (ros: reg + ident) (rs: regset) : option fundef := @@ -274,7 +277,7 @@ Qed. 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) + (State stack f sp pc (tr_exit f pc rs) m) | exec_Breturn or stk m': sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> @@ -284,7 +287,7 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := find_function ros rs = Some fd -> funsig fd = sig -> final_step stack f sp rs m (Bcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) + E0 (Callstate (Stackframe res f sp pc' (tr_exit f pc' rs) :: stack) fd rs##args m) | exec_Btailcall sig ros args stk m' fd: find_function ros rs = Some fd -> funsig fd = sig -> @@ -296,12 +299,12 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := 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') - t (State stack f sp pc' (regmap_setres res vres rs) m') + t (State stack f sp pc' (tr_exit f pc' (regmap_setres res vres rs)) m') | exec_Bjumptable arg tbl n pc': rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> final_step stack f sp rs m (Bjumptable arg tbl) - E0 (State stack f sp pc' rs m) + E0 (State stack f sp pc' (tr_exit f pc' rs) m) . (** big-step execution of one iblock *) @@ -358,11 +361,75 @@ Inductive final_state: state -> int -> Prop := | final_state_intro: forall r m, final_state (Returnstate nil (Vint r) m) r. -(** The small-step semantics for a program. *) +(** The basic small-step semantics for a BTL program. *) + +(* tid = transfer_identity *) +Definition tid (_:function) (_:exit) (rs:regset) := rs. Definition semantics (p: program) := - Semantics step (initial_state p) final_state (Genv.globalenv p). + Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). + +(** The "functional" small-step semantics for a BTL program. + at each exit, we only transfer register in "input_regs" (i.e. "alive" registers). +*) + +Definition transfer_regs (inputs: list reg) (rs: regset): regset := + init_regs (rs##inputs) inputs. + +Lemma transfer_regs_inputs inputs rs r: + List.In r inputs -> (transfer_regs inputs rs)#r = rs#r. +Proof. + unfold transfer_regs; induction inputs; simpl; intuition subst. + - rewrite Regmap.gss; auto. + - destruct (Pos.eq_dec a r). + + subst; rewrite Regmap.gss; auto. + + rewrite Regmap.gso; auto. +Qed. + +Lemma transfer_regs_dead inputs rs r: + ~List.In r inputs -> (transfer_regs inputs rs)#r = Vundef. +Proof. + unfold transfer_regs; induction inputs; simpl; intuition subst. + - rewrite Regmap.gi; auto. + - rewrite Regmap.gso; auto. +Qed. + +Definition inputs_exit (f:function) (pc: exit) : Regset.t := + match f.(fn_code)!pc with + | None => Regset.empty + | Some ib => ib.(input_regs) + end. + +Definition tr_inputs (f:function) (pc: exit): regset -> regset + := transfer_regs (Regset.elements (inputs_exit f pc)). + +(* TODO: move this elsewhere *) +Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, + SetoidList.InA (fun x y => x = y) x l <-> List.In x l. +Proof. + intros x; split. + - induction 1; simpl; eauto. + - induction l; simpl; intuition. +Qed. + +Lemma tr_inputs_exit f pc rs r: + Regset.In r (inputs_exit f pc) -> (tr_inputs f pc rs)#r = rs#r. +Proof. + intros; eapply transfer_regs_inputs. + rewrite <- SetoidList_InA_eq_equiv. + eapply Regset.elements_1; eauto. +Qed. + +Lemma tr_inputs_dead f pc rs r: + ~(Regset.In r (inputs_exit f pc)) -> (tr_inputs f pc rs)#r = Vundef. +Proof. + intros X; eapply transfer_regs_dead; intuition eauto. + eapply X. eapply Regset.elements_2. + rewrite -> SetoidList_InA_eq_equiv; eauto. +Qed. +Definition fsem (p: program) := + Semantics (step tr_inputs) (initial_state p) final_state (Genv.globalenv p). (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index b9a05a8a..5125ea3c 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -4,6 +4,8 @@ NB: an efficient implementation with hash-consing will be defined in another fil *) +(* TODO: A REVOIR - remplacer les [tid] par [tr_inputs] pour montrer la bisimulation avec [fsem] plutôt. *) + Require Import Coqlib Maps Floats. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. @@ -580,7 +582,7 @@ Local Hint Constructors sem_sfval: core. Lemma sexec_final_svf_correct ctx i sis t rs m s: sem_sistate ctx sis rs m -> - final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> + final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> sem_sfval ctx (sexec_final_sfv i sis) rs m t s. Proof. intros (PRE&MEM®). @@ -604,10 +606,12 @@ Local Hint Resolve seval_builtin_args_exact: core. Lemma sexec_final_svf_exact ctx i sis t rs m s: sem_sistate ctx sis rs m -> sem_sfval ctx (sexec_final_sfv i sis) rs m t s - -> final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. + -> final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. Proof. intros (PRE&MEM®). destruct i; simpl; intros LAST; inv LAST; eauto. + + (* Bgoto *) + econstructor. + (* Breturn *) enough (v=regmap_optget res Vundef rs) as ->; eauto. destruct res; simpl in *; congruence. @@ -621,6 +625,8 @@ Proof. intros; eapply exec_Btailcall; eauto. destruct fn; simpl in * |- *; auto. rewrite REG in * |- ; auto. + + (* Bbuiltin *) + eapply (exec_Bbuiltin tid); eauto. + (* Bjumptable *) eapply exec_Bjumptable; eauto. congruence. @@ -799,7 +805,7 @@ Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin (SIS: sem_sistate ctx sis rs m) (CONT: match ofin with | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx t s (k sis') - | Some fin => final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s + | Some fin => final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s end), sem_sstate ctx t s (sexec_rec ib sis k). Proof. @@ -816,7 +822,7 @@ Qed. (sexec is a correct over-approximation) *) Theorem sexec_correct ctx ib t s: - iblock_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> + iblock_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> sem_sstate ctx t s (sexec ib). Proof. destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). @@ -1049,7 +1055,7 @@ Lemma sexec_rec_exact ctx ib t s1: forall sis k , match iblock_istep_run (cge ctx) (csp ctx) ib rs m with | Some (out rs' m' (Some fin)) => - exists s2, final_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 + exists s2, final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 | Some (out rs' m' None) => exists sis', (sem_sstate ctx t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m') | None => False end. @@ -1095,7 +1101,7 @@ Qed. *) Theorem sexec_exact ctx ib t s1: sem_sstate ctx t s1 (sexec ib) -> - exists s2, iblock_step (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 + exists s2, iblock_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 /\ equiv_state s1 s2. Proof. intros; exploit sexec_rec_exact; eauto. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 08a77ae4..bbb984c4 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -292,7 +292,7 @@ Qed. Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 pc1 fin t s (STACKS : list_forall2 match_stackframes stack stack') (TRANSF : match_function dupmap f f') - (FS : final_step ge stack f sp rs1 m1 fin t s) + (FS : final_step tid ge stack f sp rs1 m1 fin t s) (i : instruction) (ATpc1 : (RTL.fn_code f') ! pc1 = Some i) (MF : match_final_inst dupmap fin i) @@ -338,7 +338,7 @@ Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fi (TRANSF: match_function dupmap f f') (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin)) (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None) - (FS : final_step ge stack f sp rs1 m1 fin t s) + (FS : final_step tid ge stack f sp rs1 m1 fin t s) : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'. Proof. intros; exploit iblock_istep_simulation; eauto. @@ -359,7 +359,7 @@ Proof. Qed. Theorem plus_simulation s1 t s1': - step ge s1 t s1' -> + step tid ge s1 t s1' -> forall s2 (MS: match_states s1 s2), exists s2', plus RTL.step tge s2 t s2' diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 633e1b8e..60edea49 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -492,7 +492,7 @@ Lemma match_strong_state_simu (MSS2 : match_strong_state dupmap st st' f f' sp rs2 m2 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) (MES : measure ib2 < n) : exists (oib' : option iblock), - (exists s2', step tge (State st' f' sp pcB0 rs0 m0) E0 s2' + (exists s2', step tid tge (State st' f' sp pcB0 rs0 m0) E0 s2' /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) s2') \/ (omeasure oib' < n /\ E0=E0 /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) (State st' f' sp pcB0 rs0 m0)). @@ -523,7 +523,7 @@ Lemma opt_simu_intro (MSTRONG : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst) (NGOTO : is_goto ib = false) : exists (oib' : option iblock), - (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') + (exists s2', step tid tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). Proof. inversion MSTRONG; subst. inv MIB. @@ -677,7 +677,7 @@ Theorem opt_simu s1 t s1' oib s2: RTL.step ge s1 t s1' -> match_states oib s1 s2 -> exists (oib' : option iblock), - (exists s2', step tge s2 t s2' /\ match_states oib' s1' s2') + (exists s2', step tid tge s2 t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure oib /\ t=E0 /\ match_states oib' s1' s2) . Proof. -- cgit From 3924865f98f3232e38ef4c1c0464d332de741c12 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 20 May 2021 18:54:03 +0200 Subject: update roadmap --- scheduling/BTLroadmap.md | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 39ea12d0..32d0b365 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -102,7 +102,7 @@ totales) dans BTL. Pour l'encodage de SSA, il suffira d'étendre le mécanisme d'initialisation du "rs0" d'un bloc avec un passage de paramètres. -En fait, pour le test d'exécution symbolique, il semble plus simple de mettre les Vundef à la fin de la transition... +En fait, pour le test d'exécution symbolique, il semble plus simple de mettre les Vundef à la fin de la transition (précédente) plutôt qu'au début (de la suivante). **IMPLEM PLAN** @@ -113,6 +113,9 @@ En fait, pour le test d'exécution symbolique, il semble plus simple de mettre l c'est la correction du "input_regs" qui garantit que la simulation est correct. La preuve devrait normalement être très similaire à RTLpathLivegenproof. +**STATUS** + +1. See [BTL.fsem] (Pour l'instant CFGSemantics => BTL.semantics et FunctionalSemantics => BTL.fsem) ## "Basic" symbolic execution / symbolic simulation @@ -124,8 +127,12 @@ We will implement a "basic" (e.g without rewriting rules) "less_undef" simulatio **CURRENT STATUS** -- model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms. -- next step: high-level specification of "symbolic simulation" + preservation proof w.r.t BTL.FunctionalSemantics. +- model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms wrt BTL.semantics + +**next steps** +- preservation proof w.r.t BTL.fsem. +- high-level specification of "symbolic simulation" +- ... ## Port of RTLpath optimizations to BTL -- cgit From 25a4620c95aaa6b017443da29fcf3d033a44a86f Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 21 May 2021 15:59:31 +0200 Subject: improve fsem --- scheduling/BTL.v | 45 ++++++++++++++++++++++++++++++--------------- 1 file changed, 30 insertions(+), 15 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index cb60fed1..fb6d5cea 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -161,8 +161,16 @@ Record outcome := out { Section RELSEM. -(* final_step is parametrized by a function to transfer regset on each exit *) -Variable tr_exit: function -> exit -> regset -> regset. +(** [step] (and in particular [final_step]) is parametrized by function [tr_exit] to transfer registers on each iblock exit. + + In particular, [tr_exit f pc or rs] computes from [rs] the [rs'] on which the execution from iblock [pc] in [f] will start. + + Here, [or] is an optional register that will be assigned after exiting the iblock, but before entering in [pc]: e.g. the register set by a function call + before entering in return address. + +*) + +Variable tr_exit: function -> exit -> option reg -> regset -> regset. Variable ge: genv. @@ -277,7 +285,7 @@ Qed. 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 (tr_exit f pc rs) m) + (State stack f sp pc (tr_exit f pc None rs) m) | exec_Breturn or stk m': sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> @@ -287,7 +295,7 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := find_function ros rs = Some fd -> funsig fd = sig -> final_step stack f sp rs m (Bcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' (tr_exit f pc' rs) :: stack) fd rs##args m) + E0 (Callstate (Stackframe res f sp pc' (tr_exit f pc' (Some res) rs) :: stack) fd rs##args m) | exec_Btailcall sig ros args stk m' fd: find_function ros rs = Some fd -> funsig fd = sig -> @@ -299,12 +307,15 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := 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') - t (State stack f sp pc' (tr_exit f pc' (regmap_setres res vres rs)) m') + (* TODO: not clear that this is the better choice ! + we coud also do something like [regmap_setres res vres (tr_exit f pc' (reg_builtin_res res) rs)] + *) + t (State stack f sp pc' (tr_exit f pc' None (regmap_setres res vres rs)) m') | exec_Bjumptable arg tbl n pc': rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> final_step stack f sp rs m (Bjumptable arg tbl) - E0 (State stack f sp pc' (tr_exit f pc' rs) m) + E0 (State stack f sp pc' (tr_exit f pc' None rs) m) . (** big-step execution of one iblock *) @@ -364,7 +375,7 @@ Inductive final_state: state -> int -> Prop := (** The basic small-step semantics for a BTL program. *) (* tid = transfer_identity *) -Definition tid (_:function) (_:exit) (rs:regset) := rs. +Definition tid (_:function) (_:exit) (_: option reg) (rs:regset) := rs. Definition semantics (p: program) := Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). @@ -394,14 +405,18 @@ Proof. - rewrite Regmap.gso; auto. Qed. -Definition inputs_exit (f:function) (pc: exit) : Regset.t := +Definition inputs_exit (f:function) (pc: exit) (or:option reg): Regset.t := match f.(fn_code)!pc with | None => Regset.empty - | Some ib => ib.(input_regs) + | Some ib => + match or with + | Some r => Regset.remove r ib.(input_regs) + | None => ib.(input_regs) + end end. -Definition tr_inputs (f:function) (pc: exit): regset -> regset - := transfer_regs (Regset.elements (inputs_exit f pc)). +Definition tr_inputs (f:function) (pc: exit) (or:option reg): regset -> regset + := transfer_regs (Regset.elements (inputs_exit f pc or)). (* TODO: move this elsewhere *) Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, @@ -412,16 +427,16 @@ Proof. - induction l; simpl; intuition. Qed. -Lemma tr_inputs_exit f pc rs r: - Regset.In r (inputs_exit f pc) -> (tr_inputs f pc rs)#r = rs#r. +Lemma tr_inputs_exit f pc or rs r: + Regset.In r (inputs_exit f pc or) -> (tr_inputs f pc or rs)#r = rs#r. Proof. intros; eapply transfer_regs_inputs. rewrite <- SetoidList_InA_eq_equiv. eapply Regset.elements_1; eauto. Qed. -Lemma tr_inputs_dead f pc rs r: - ~(Regset.In r (inputs_exit f pc)) -> (tr_inputs f pc rs)#r = Vundef. +Lemma tr_inputs_dead f pc or rs r: + ~(Regset.In r (inputs_exit f pc or)) -> (tr_inputs f pc or rs)#r = Vundef. Proof. intros X; eapply transfer_regs_dead; intuition eauto. eapply X. eapply Regset.elements_2. -- cgit From 3e6844222d39b8a76bf0af20b9cbb0dfa2e35b5a Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 21 May 2021 21:18:59 +0200 Subject: Now supporting Bnop insertion in conditions --- scheduling/BTL.v | 66 +++++++++++++------- scheduling/BTLtoRTLaux.ml | 17 ++--- scheduling/BTLtoRTLproof.v | 29 ++++++--- scheduling/RTLtoBTLaux.ml | 11 ++-- scheduling/RTLtoBTLproof.v | 151 ++++++++++++++++++--------------------------- 5 files changed, 135 insertions(+), 139 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index ca727f82..05391c58 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -58,7 +58,7 @@ Inductive iblock: Type := (* final instructions that stops block execution *) | BF (fi: final) (iinfo: inst_info) (* basic instructions that continues block execution, except when aborting *) - | Bnop (iinfo: inst_info) (* nop instruction *) + | Bnop (oiinfo: option inst_info) (* nop instruction *) | Bop (op:operation) (args:list reg) (dest:reg) (iinfo: inst_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) (iinfo: inst_info) @@ -204,7 +204,7 @@ 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 iinfo: iblock_istep sp rs m (BF fin iinfo) rs m (Some fin) - | exec_nop rs m iinfo: iblock_istep sp rs m (Bnop iinfo) rs m None +| exec_nop rs m oiinfo: iblock_istep sp rs m (Bnop oiinfo) rs m None | exec_op rs m op args res v iinfo (EVAL: eval_operation ge sp op rs##args m = Some v) : iblock_istep sp rs m (Bop op args res iinfo) (rs#res <- v) m None @@ -425,9 +425,11 @@ 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 iinfo) None - | mib_nop isfst pc pc' iinfo: + | mib_nop_on_rtl isfst pc pc' iinfo: cfg!pc = Some (Inop pc') -> - match_iblock dupmap cfg isfst pc (Bnop iinfo) (Some pc') + match_iblock dupmap cfg isfst pc (Bnop (Some iinfo)) (Some pc') + | mib_nop_skip pc: + match_iblock dupmap cfg false pc (Bnop None) (Some pc) | mib_op isfst pc pc' op lr r iinfo: cfg!pc = Some (Iop op lr r pc') -> match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc') @@ -592,10 +594,16 @@ 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 _ => - match cfg!pc with - | Some (Inop pc') => OK (Some pc') - | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") + | Bnop oiinfo => + match oiinfo with + | Some _ => + match cfg!pc with + | Some (Inop pc') => OK (Some pc') + | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") + end + | None => + if negb isfst then OK (Some pc) + else Error (msg "BTL.verify_block: isfst is true Bnop (on_rtl is false)") end | Bop op lr r _ => match cfg!pc with @@ -712,8 +720,11 @@ Proof. destruct (Pos.eq_dec _ _); try discriminate. subst. inv EQ0. eapply mib_BF; eauto. constructor; assumption. - (* Bnop *) - destruct i; inv H. - constructor; assumption. + destruct oiinfo; simpl in *. + + destruct (cfg!pc) eqn:EQCFG; try discriminate. + destruct i0; inv H. constructor; assumption. + + destruct isfst; try discriminate. inv H. + constructor; assumption. - (* Bop *) destruct i; try discriminate. destruct (eq_operation _ _); try discriminate. @@ -803,10 +814,16 @@ Definition is_goto (ib: iblock): bool := Definition is_atom (ib: iblock): bool := match ib with - | Bseq _ _ | Bcond _ _ _ _ _ => false + | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None => false | _ => true end. +Definition Bnop_dec k: { k=Bnop None} + { k<>(Bnop None) }. +Proof. + destruct k; simpl; try destruct oiinfo; + intuition congruence. +Defined. + (** Is expand property to only support atomic instructions on the left part of a Bseq *) Inductive is_expand: iblock -> Prop := | exp_Bseq ib1 ib2: @@ -814,8 +831,8 @@ Inductive is_expand: iblock -> Prop := is_expand ib2 -> is_expand (Bseq ib1 ib2) | exp_Bcond cond args ib1 ib2 iinfo: - is_expand ib1 -> - is_expand ib2 -> + is_expand ib1 \/ ib1 = Bnop None -> + is_expand ib2 \/ ib2 = Bnop None -> is_expand (Bcond cond args ib1 ib2 iinfo) | exp_others ib: is_atom ib = true -> @@ -823,22 +840,27 @@ Inductive is_expand: iblock -> Prop := . Local Hint Constructors is_expand: core. -Fixpoint expand (ib: iblock) (k: option iblock): iblock := +Fixpoint expand (ib: iblock) (k: iblock): iblock := match ib with - | Bseq ib1 ib2 => expand ib1 (Some (expand ib2 k)) + | Bseq ib1 ib2 => expand ib1 (expand ib2 k) | Bcond cond args ib1 ib2 iinfo => Bcond cond args (expand ib1 k) (expand ib2 k) iinfo | BF fin iinfo => BF fin iinfo - | ib => - match k with - | None => ib - | Some rem => Bseq ib rem - end + | Bnop None => k + | ib => + if Bnop_dec k then ib else Bseq ib k end. Lemma expand_correct ib: forall k, - (match k with Some rem => is_expand rem | None => True end) + (k <> (Bnop None) -> is_expand k) + -> (expand ib k) <> (Bnop None) -> is_expand (expand ib k). Proof. - induction ib; simpl; intros; try autodestruct; auto. + induction ib; simpl; intros; auto. + 1-4: + try destruct oiinfo; + destruct (Bnop_dec k); auto. + constructor. + - destruct (Bnop_dec (expand ib1 k)); auto. + - destruct (Bnop_dec (expand ib2 k)); auto. Qed. diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index b4833b2c..36d3e6a4 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -6,15 +6,11 @@ open AuxTools open DebugPrint open BTLaux -let is_atom = function - | Bseq (_, _) | Bcond (_, _, _, _, _) -> false - | _ -> true - let get_inumb iinfo = P.of_int iinfo.inumb let rec get_ib_num = function | BF (Bgoto s, _) -> s - | Bnop iinfo + | Bnop Some iinfo | Bop (_, _, _, iinfo) | Bload (_, _, _, _, _, iinfo) | Bstore (_, _, _, _, iinfo) @@ -22,6 +18,7 @@ let rec get_ib_num = function | BF (_, iinfo) -> get_inumb iinfo | Bseq (ib1, _) -> get_ib_num ib1 + | Bnop None -> failwith "get_ib_num: Bnop None found" let translate_function code entry = let rtl_code = ref PTree.empty in @@ -45,23 +42,21 @@ let translate_function code entry = | Bcond (cond, lr, BF (Bgoto s1, _), BF (Bgoto s2, _), iinfo) -> next_nodes := s1 :: s2 :: !next_nodes; Some (Icond (cond, lr, s1, s2, iinfo.pcond), get_inumb iinfo) - | Bcond (cond, lr, BF (Bgoto s1, _), ib2, iinfo) -> + | Bcond (cond, lr, BF (Bgoto s1, _), Bnop None, iinfo) -> assert (iinfo.pcond = Some false); next_nodes := s1 :: !next_nodes; - translate_btl_block ib2 None; Some - ( Icond (cond, lr, s1, get_ib_num ib2, iinfo.pcond), + ( Icond (cond, lr, s1, get_ib_num (get_some k), iinfo.pcond), get_inumb iinfo ) (* TODO gourdinl remove dynamic check *) | Bcond (_, _, _, _, _) -> failwith "translate_function: invalid Bcond" | Bseq (ib1, ib2) -> - (* TODO gourdinl remove dynamic check *) - assert (is_atom ib1); translate_btl_block ib1 (Some ib2); translate_btl_block ib2 None; None - | Bnop iinfo -> Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) + | Bnop Some iinfo -> Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) + | Bnop None -> failwith "translate_function: invalid Bnop" | Bop (op, lr, rd, iinfo) -> Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo) | Bload (trap, chk, addr, lr, rd, iinfo) -> diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 9b37d8fa..2e8d960c 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -202,6 +202,17 @@ Proof. eapply plus_trans; eauto. Qed. +Lemma css_E0_trans isfst isfst' s0 s1 s2: + cond_star_step (isfst=false) s0 E0 s1 -> + cond_star_step (false=isfst') s1 E0 s2 -> + cond_star_step (isfst=isfst') s0 E0 s2. +Proof. + intros S1 S2. + inversion S1; subst; eauto. + inversion S2; subst; eauto. + eapply css_plus_trans; eauto. +Qed. + Lemma css_star P s0 s1 t: cond_star_step P s0 t s1 -> star RTL.step tge s0 t s1. @@ -219,7 +230,7 @@ Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin forall pc0 opc isfst (MIB: match_iblock dupmap (RTL.fn_code f') isfst pc0 ib opc), match ofin with - | None => exists pc1,(opc = Some pc1) /\ plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) + | None => exists pc1,(opc = Some pc1) /\ cond_star_step (isfst = false) (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) | Some fin => exists isfst' pc1 iinfo, match_iblock dupmap (RTL.fn_code f') isfst' pc1 (BF fin iinfo) None /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1) @@ -231,19 +242,19 @@ Proof. subst. repeat eexists; eauto. - (* exec_nop *) - inv MIB. exists pc'; split; eauto. + inv MIB; eexists; split; econstructor; eauto. - (* exec_op *) - inv MIB. exists pc'; split; auto. + inv MIB. exists pc'; split; auto; constructor. apply plus_one. eapply exec_Iop; eauto. erewrite <- eval_operation_preserved; eauto. intros; rewrite symbols_preserved; trivial. - (* exec_load *) - inv MIB. exists pc'; split; auto. + inv MIB. exists pc'; split; auto; constructor. apply plus_one. eapply exec_Iload; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. - (* exec_store *) - inv MIB. exists pc'; split; auto. + inv MIB. exists pc'; split; auto; constructor. apply plus_one. eapply exec_Istore; eauto. erewrite <- eval_addressing_preserved; eauto. intros; rewrite symbols_preserved; trivial. @@ -257,10 +268,10 @@ Proof. destruct ofin; simpl. + intros (ifst2 & pc2 & iinfo & M2 & STEP2). repeat eexists; eauto. - eapply css_plus_trans; eauto. + eapply css_E0_trans; eauto. + intros (pc2 & EQpc2 & STEP2); inv EQpc2. eexists; split; auto. - eapply plus_trans; eauto. + eapply css_E0_trans; eauto. - (* exec_cond *) inv MIB. rename H10 into JOIN. (* is_join_opt opc1 opc2 opc *) @@ -275,7 +286,7 @@ Proof. intros (pc1 & OPC & PLUS). inv OPC. inv JOIN; eexists; split; eauto. all: - eapply plus_trans; eauto. + eapply css_plus_trans; eauto. + (* taking ifnot *) destruct ofin; simpl. * (* ofin is Some final *) @@ -286,7 +297,7 @@ Proof. intros (pc1 & OPC & PLUS). subst. inv JOIN; eexists; split; eauto. all: - eapply plus_trans; eauto. + eapply css_plus_trans; eauto. Qed. Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 pc1 fin t s diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 43556150..4be2b180 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -2,9 +2,9 @@ open Maps open RTL open BTL open Registers -open DebugPrint open PrintBTL open AuxTools +open DebugPrint open BTLaux let undef_node = -1 @@ -20,13 +20,13 @@ let encaps_final inst osucc = | BF _ | Bcond _ -> inst | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo)) -let translate_inst iinfo inst is_final = +let translate_inst (iinfo: BTL.inst_info) inst is_final = let osucc = ref None in let btli = match inst with | Inop s -> osucc := Some s; - Bnop iinfo + Bnop (Some iinfo) | Iop (op, lr, rd, s) -> osucc := Some s; Bop (op, lr, rd, iinfo) @@ -80,12 +80,13 @@ let translate_function code entry join_points = assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; iinfo.pcond <- info; + Bseq ( Bcond ( cond, lr, BF (Bgoto ifso, def_iinfo), - build_btl_block s, - iinfo ) + Bnop None, + iinfo ), build_btl_block s) | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7ad1472b..f3e258ae 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -242,11 +242,20 @@ Proof. Qed. Lemma expand_entry_isnt_goto dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) None) None -> - is_goto (expand (entry ib) None) = false. + match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) (Bnop None)) None -> + is_goto (expand (entry ib) (Bnop None)) = false. Proof. - destruct (is_goto (expand (entry ib) None))eqn:EQG; - destruct (expand (entry ib) None); + destruct (is_goto (expand (entry ib) (Bnop None))) eqn:EQG; + destruct (expand (entry ib) (Bnop None)); + try destruct fi; try discriminate; trivial. + intros; inv H; inv H5. +Qed. + +Lemma expand_entry_isnt_bnop dupmap f pc ib: + match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) (Bnop None)) None -> + expand (entry ib) (Bnop None) <> Bnop None. +Proof. + destruct (expand (entry ib) (Bnop None)) eqn:EQG; try destruct fi; try discriminate; trivial. intros; inv H; inv H5. Qed. @@ -268,38 +277,23 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. +Local Hint Constructors iblock_istep: core. Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) k ofin2 rs2 m2 (CONT: match ofin1 with - | None => - (k = None /\ rs2=rs1 /\ m2=m1 /\ ofin2 = None) - \/ (exists rem, k = Some rem - /\ iblock_istep tge sp rs1 m1 rem rs2 m2 ofin2) + | None => iblock_istep tge sp rs1 m1 k rs2 m2 ofin2 | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 end), iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 ofin2. Proof. - induction 1; simpl. - { (* BF *) - intros ? ? ? ? (HRS & HM & HOF); subst. - constructor. } - (*destruct k; intros. try inv CONT.*) - 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct k; intros; destruct CONT as [[HK [HRS [HM HO]]]|[rem [HR ISTEP]]]; - subst; try (inv HK; fail); try (inv HR; fail); try (econstructor; eauto; fail); - inversion HR; subst; clear HR; - eapply exec_seq_continue; [ econstructor; eauto | assumption]. - - (* Bseq_stop *) - destruct k; intros; apply IHISTEP; eauto. - - (* Bseq_continue *) - destruct ofin; intros. - + destruct CONT as [HRS [HM HOF]]; subst. - eapply IHISTEP1; right. eexists; repeat split; eauto. - + destruct CONT as [[HK [HRS [HM HO]]]|[rem [HR ISTEP]]]; subst. - * eapply IHISTEP1; right. eexists; repeat split; eauto. - eapply IHISTEP2; left; simpl; auto. - * eapply IHISTEP1; right. eexists; repeat split; eauto. + induction 1; simpl; intros; intuition subst; eauto. + { (* Bnop *) + autodestruct; intros OI; inv OI; destruct (Bnop_dec k); subst; + try (inv CONT; constructor; fail); repeat econstructor; eauto. } + 1-3: (* Bop, Bload, Bstore *) + intros; destruct (Bnop_dec k); subst; repeat econstructor; eauto; + inv CONT; econstructor; eauto. - (* Bcond *) destruct ofin; intros; econstructor; eauto; @@ -308,47 +302,19 @@ Qed. Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> - iblock_istep tge sp rs0 m0 (expand ib None) rs1 m1 ofin. + iblock_istep tge sp rs0 m0 (expand ib (Bnop None)) rs1 m1 ofin. Proof. intros; eapply expand_iblock_istep_rec_correct; eauto. destruct ofin; simpl; auto. Qed. -(* TODO gourdinl useless? *) -Lemma expand_iblock_istep_run_Some_rec sp ib rs0 m0 rs1 m1 ofin1: - forall (ISTEP: iblock_istep_run tge sp ib rs0 m0 = - Some {| _rs := rs1; _m := m1; _fin := ofin1 |}) - k ofin2 rs2 m2 - (CONT: match ofin1 with - | None => - (k = None /\ rs2=rs1 /\ m2=m1 /\ ofin2 = None) - \/ (exists rem, k = Some rem - /\ iblock_istep_run tge sp rem rs1 m1 = - Some {| _rs := rs2; _m := m2; _fin := ofin2 |}) - | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 - end), - iblock_istep_run tge sp (expand ib k) rs0 m0 = - Some {| _rs := rs2; _m := m2; _fin := ofin2 |}. -Proof. - intros. destruct ofin1; - rewrite <- iblock_istep_run_equiv in *. - - destruct CONT as [HRS [HM HO]]; subst. - eapply expand_iblock_istep_rec_correct; eauto. - simpl; auto. - - eapply expand_iblock_istep_rec_correct; eauto. - simpl. destruct CONT as [HL | [rem [HR ISTEP']]]. - left; auto. rewrite <- iblock_istep_run_equiv in ISTEP'. - right; eexists; split; eauto. -Qed. - Lemma expand_iblock_istep_run_None_rec sp ib: forall rs0 m0 o k (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) (CONT: match o with | Some (out rs1 m1 ofin) => - exists rem, - k = Some rem /\ ofin = None /\ - iblock_istep_run tge sp rem rs1 m1 = None + ofin = None /\ + iblock_istep_run tge sp k rs1 m1 = None | _ => True end), iblock_istep_run tge sp (expand ib k) rs0 m0 = None. @@ -357,35 +323,34 @@ Proof. try discriminate. - (* BF *) intros; destruct o; try discriminate; simpl in *. - inv ISTEP. destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO. + inv ISTEP. destruct CONT as [HO ISTEP]; inv HO. - (* Bnop *) - intros; destruct o; inv ISTEP; destruct k; - destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + intros; destruct o; inv ISTEP; destruct oiinfo, CONT; auto. + destruct (Bnop_dec k); subst; repeat econstructor; eauto. - (* Bop *) - intros; destruct o; - destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; destruct k; - simpl; rewrite EVAL; auto; destruct CONT as [rem [HR [HO ISTEP]]]; - inv HR; inv HO; trivial. + intros; destruct o; destruct (Bnop_dec k); + destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; + simpl; rewrite EVAL; auto; destruct CONT as [HO ISTEP]; + trivial. - (* Bload *) - intros; destruct o; + intros; destruct o; destruct (Bnop_dec k); destruct (trap) eqn:TRAP; try destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; destruct k; + try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + destruct CONT as [HO ISTEP]; trivial. - (* Bstore *) - intros; destruct o; + intros; destruct o; destruct (Bnop_dec k); destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; destruct k; + try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [rem [HR [HO ISTEP]]]; inv HR; inv HO; trivial. + destruct CONT as [HO ISTEP]; inv HO; trivial. - (* Bseq *) intros. eapply IHib1; eauto. destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1; try auto. - destruct o0. eexists; split; eauto. simpl in *. - destruct _fin; inv ISTEP. - + destruct CONT as [rem [_ [CONTRA _]]]; inv CONTRA. + destruct o0; simpl in *. destruct _fin; inv ISTEP. + + destruct CONT as [CONTRA _]; inv CONTRA. + split; auto. eapply IHib2; eauto. - (* Bcond *) intros; destruct (eval_condition _ _ _); trivial. @@ -396,7 +361,7 @@ Qed. Lemma expand_preserves_iblock_istep_run_None sp ib: forall rs m, iblock_istep_run tge sp ib rs m = None - -> iblock_istep_run tge sp (expand ib None) rs m = None. + -> iblock_istep_run tge sp (expand ib (Bnop None)) rs m = None. Proof. intros; eapply expand_iblock_istep_run_None_rec; eauto. simpl; auto. @@ -404,7 +369,7 @@ Qed. Lemma expand_preserves_iblock_istep_run sp ib: forall rs m, iblock_istep_run tge sp ib rs m = - iblock_istep_run tge sp (expand ib None) rs m. + iblock_istep_run tge sp (expand ib (Bnop None)) rs m. Proof. intros. destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP. @@ -415,28 +380,24 @@ Proof. apply expand_preserves_iblock_istep_run_None; auto. Qed. +Local Hint Constructors match_iblock: core. Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst: forall opc1 (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2 (CONT: match opc1 with | Some pc' => - k = None /\ opc2 = opc1 \/ - (exists rem, k = Some rem - /\ match_iblock dupmap cfg false pc' rem opc2) + match_iblock dupmap cfg false pc' k opc2 | None => opc2=opc1 end), match_iblock dupmap cfg isfst pc (expand ib k) opc2. Proof. - induction 1; simpl. + induction 1; simpl; intros; eauto. { (* BF *) - intros; inv CONT; econstructor; eauto. } - 1-4: (* Bnop *) - destruct k; intros; destruct CONT as [[HK HO] | [rem [HR MIB]]]; - try inv HK; try inv HO; try inv HR; repeat econstructor; eauto. + inv CONT; econstructor; eauto. } + 1-4: (* Bnop, Bop, Bload, Bstore *) + destruct (Bnop_dec k); subst; eauto; inv CONT; econstructor; eauto. { (* Bgoto *) intros; inv CONT; apply mib_exit; auto. } - { (* Bseq *) - intros. eapply IHMIB1. right. eexists; split; eauto. } { (* Bcond *) intros. inv H0; econstructor; eauto; try econstructor. @@ -445,7 +406,7 @@ Qed. Lemma expand_matchiblock_correct dupmap cfg ib pc isfst opc: match_iblock dupmap cfg isfst pc ib opc -> - match_iblock dupmap cfg isfst pc (expand ib None) opc. + match_iblock dupmap cfg isfst pc (expand ib (Bnop None)) opc. Proof. intros. eapply expand_matchiblock_rec_correct; eauto. @@ -508,6 +469,7 @@ Proof. apply iblock_istep_run_equiv in BTL_RUN; eauto. + econstructor; apply expand_matchiblock_correct in MI. econstructor; eauto. apply expand_correct; trivial. + intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. econstructor. apply expand_preserves_iblock_istep_run. eapply expand_entry_isnt_goto; eauto. - (* Others *) @@ -575,7 +537,8 @@ Proof. eapply eval_builtin_args_preserved; eauto. eapply external_call_symbols_preserved; eauto. eapply senv_preserved. * econstructor; eauto; apply expand_matchiblock_correct in MI. - { econstructor; eauto. apply expand_correct; trivial. + { econstructor; eauto. apply expand_correct; trivial. + intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. apply star_refl. apply expand_preserves_iblock_istep_run. } eapply expand_entry_isnt_goto; eauto. + (* Bjumptable *) @@ -591,6 +554,7 @@ Proof. eapply BTL_RUN. econstructor; eauto. * econstructor; eauto; apply expand_matchiblock_correct in MI. { econstructor; eauto. apply expand_correct; trivial. + intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. apply star_refl. apply expand_preserves_iblock_istep_run. } eapply expand_entry_isnt_goto; eauto. - (* mib_exit *) @@ -599,7 +563,7 @@ Proof. inversion H; subst; try (inv IS_EXPD; try inv H5; discriminate; fail); inversion STEP; subst; try_simplify_someHyps; intros. - + (* Bnop *) + + (* Bnop is_rtl *) eapply match_strong_state_simu. 1,2: do 2 (econstructor; eauto). econstructor; eauto. @@ -637,7 +601,8 @@ Proof. intros; rewrite H12 in BTL_RUN. destruct b; eapply match_strong_state_simu; eauto. 1,3: inv H2; econstructor; eauto. - 1,3,5,7: inv IS_EXPD; auto; discriminate. + 1,3: inv IS_EXPD; [ inv H3; auto; inv H0 | inv H ]. + 3,5: inv IS_EXPD; [ inv H7; auto; inv H1 | inv H ]. 1-4: eapply star_right; eauto. assert (measure bnot > 0) by apply measure_pos; lia. assert (measure bso > 0) by apply measure_pos; lia. @@ -698,6 +663,7 @@ Proof. * apply expand_matchiblock_correct in MI. econstructor. econstructor; eauto. apply expand_correct; trivial. + intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. 3: eapply expand_entry_isnt_goto; eauto. all: erewrite preserv_fnparams; eauto. constructor. @@ -719,7 +685,8 @@ Proof. + apply expand_matchiblock_correct in MI. econstructor. econstructor; eauto. apply expand_correct; trivial. - constructor. apply expand_preserves_iblock_istep_run. + constructor. congruence. eapply expand_entry_isnt_bnop; eauto. + econstructor. apply expand_preserves_iblock_istep_run. eapply expand_entry_isnt_goto; eauto. Qed. -- cgit From 30e41117b57ab20beb1876e38c26dbddc5a58dfb Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sun, 23 May 2021 20:37:22 +0200 Subject: splitting is_expand property with a weak version for conditions --- scheduling/BTL.v | 81 +++++++++++++++++++++++--------- scheduling/RTLtoBTLproof.v | 114 +++++++++++++++++++++++---------------------- 2 files changed, 117 insertions(+), 78 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 05391c58..0f9ef44f 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -812,7 +812,7 @@ Definition is_goto (ib: iblock): bool := | _ => false end. -Definition is_atom (ib: iblock): bool := +Definition is_RTLatom (ib: iblock): bool := match ib with | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None => false | _ => true @@ -824,43 +824,80 @@ Proof. intuition congruence. Defined. -(** Is expand property to only support atomic instructions on the left part of a Bseq *) +(** Is expand property to only support atomic instructions on the left part of a Bseq: + The tree is right-expanded, and block are normalized to simplify RTL simulation *) Inductive is_expand: iblock -> Prop := | exp_Bseq ib1 ib2: - is_atom ib1 = true -> + is_RTLatom ib1 = true -> is_expand ib2 -> is_expand (Bseq ib1 ib2) - | exp_Bcond cond args ib1 ib2 iinfo: - is_expand ib1 \/ ib1 = Bnop None -> - is_expand ib2 \/ ib2 = Bnop None -> - is_expand (Bcond cond args ib1 ib2 iinfo) + | exp_Bcond cond args ib1 ib2 i: + is_expand ib1 -> + is_expand ib2 -> + is_expand (Bcond cond args ib1 ib2 i) | exp_others ib: - is_atom ib = true -> + is_RTLatom ib = true -> is_expand ib . Local Hint Constructors is_expand: core. -Fixpoint expand (ib: iblock) (k: iblock): iblock := +(** The weak version with more specific hypothesis for conditions *) +Inductive is_wexpand: iblock -> Prop := + | wexp_Bseq ib1 ib2: + is_RTLatom ib1 = true -> + is_wexpand ib2 -> + is_wexpand (Bseq ib1 ib2) + | wexp_Bcond cond args ib1 ib2 iinfo: + (ib1 <> Bnop None -> is_wexpand ib1) -> + (ib2 <> Bnop None -> is_wexpand ib2) -> + is_wexpand (Bcond cond args ib1 ib2 iinfo) + | wexp_others ib: + is_RTLatom ib = true -> + is_wexpand ib + . +Local Hint Constructors is_wexpand: core. + +Fixpoint normRTL (ib: iblock) (k: iblock): iblock := match ib with - | Bseq ib1 ib2 => expand ib1 (expand ib2 k) + | Bseq ib1 ib2 => normRTL ib1 (normRTL ib2 k) | Bcond cond args ib1 ib2 iinfo => - Bcond cond args (expand ib1 k) (expand ib2 k) iinfo + Bcond cond args (normRTL ib1 k) (normRTL ib2 k) iinfo | BF fin iinfo => BF fin iinfo | Bnop None => k | ib => if Bnop_dec k then ib else Bseq ib k end. -Lemma expand_correct ib: forall k, - (k <> (Bnop None) -> is_expand k) - -> (expand ib k) <> (Bnop None) - -> is_expand (expand ib k). +Lemma normRTL_correct ib: forall k, + (k <> (Bnop None) -> is_wexpand k) -> + (normRTL ib k) <> Bnop None -> + is_wexpand (normRTL ib k). Proof. - induction ib; simpl; intros; auto. - 1-4: - try destruct oiinfo; - destruct (Bnop_dec k); auto. - constructor. - - destruct (Bnop_dec (expand ib1 k)); auto. - - destruct (Bnop_dec (expand ib2 k)); auto. + induction ib; simpl; intros; try autodestruct; auto. + intros; destruct (Bnop_dec k); auto. +Qed. + +Lemma is_join_opt_None {A} (opc1 opc2: option A): + is_join_opt opc1 opc2 None -> opc1 = None /\ opc2 = None. +Proof. + intros X. inv X; auto. +Qed. + +Lemma match_iblock_None_not_Bnop dupmap cfg isfst pc ib: + match_iblock dupmap cfg isfst pc ib None -> ib <> Bnop None. +Proof. + intros X; inv X; try congruence. +Qed. +Local Hint Resolve match_iblock_None_not_Bnop: core. + +Lemma is_wexpand_expand dupmap cfg ib: + is_wexpand ib -> + forall isfst pc + (MIB: match_iblock dupmap cfg isfst pc ib None), + is_expand ib. +Proof. + induction 1; simpl; intros; auto; try (inv MIB); eauto. + (* Bcond *) + destruct (is_join_opt_None opc1 opc2); subst; eauto. + econstructor; eauto. Qed. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index f3e258ae..cd6c4dae 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -241,25 +241,34 @@ Proof. destruct fi; trivial. inv H. inv H5. Qed. -Lemma expand_entry_isnt_goto dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) (Bnop None)) None -> - is_goto (expand (entry ib) (Bnop None)) = false. +Lemma normRTL_entry_isnt_goto dupmap f pc ib: + match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None -> + is_goto (normRTL (entry ib) (Bnop None)) = false. Proof. - destruct (is_goto (expand (entry ib) (Bnop None))) eqn:EQG; - destruct (expand (entry ib) (Bnop None)); + destruct (is_goto (normRTL (entry ib) (Bnop None))) eqn:EQG; + destruct (normRTL (entry ib) (Bnop None)); try destruct fi; try discriminate; trivial. intros; inv H; inv H5. Qed. -Lemma expand_entry_isnt_bnop dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (expand (entry ib) (Bnop None)) None -> - expand (entry ib) (Bnop None) <> Bnop None. +Lemma normRTL_entry_isnt_bnop dupmap f pc ib: + match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None -> + normRTL (entry ib) (Bnop None) <> Bnop None. Proof. - destruct (expand (entry ib) (Bnop None)) eqn:EQG; + destruct (normRTL (entry ib) (Bnop None)) eqn:EQG; try destruct fi; try discriminate; trivial. intros; inv H; inv H5. Qed. +Lemma is_expand_normRTL_entry dupmap f ib pc + (MI : match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None): + is_expand (normRTL (entry ib) (Bnop None)). +Proof. + eapply is_wexpand_expand; eauto. + apply normRTL_correct; try congruence. + eapply normRTL_entry_isnt_bnop; eauto. +Qed. + Lemma list_nth_z_rev_dupmap: forall dupmap ln ln' (pc pc': node) val, list_nth_z ln val = Some pc -> @@ -278,14 +287,14 @@ Proof. Qed. Local Hint Constructors iblock_istep: core. -Lemma expand_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: +Lemma normRTL_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) k ofin2 rs2 m2 (CONT: match ofin1 with | None => iblock_istep tge sp rs1 m1 k rs2 m2 ofin2 | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 end), - iblock_istep tge sp rs0 m0 (expand ib k) rs2 m2 ofin2. + iblock_istep tge sp rs0 m0 (normRTL ib k) rs2 m2 ofin2. Proof. induction 1; simpl; intros; intuition subst; eauto. { (* Bnop *) @@ -300,15 +309,15 @@ Proof. destruct b; eapply IHISTEP; eauto. Qed. -Lemma expand_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: +Lemma normRTL_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> - iblock_istep tge sp rs0 m0 (expand ib (Bnop None)) rs1 m1 ofin. + iblock_istep tge sp rs0 m0 (normRTL ib (Bnop None)) rs1 m1 ofin. Proof. - intros; eapply expand_iblock_istep_rec_correct; eauto. + intros; eapply normRTL_iblock_istep_rec_correct; eauto. destruct ofin; simpl; auto. Qed. -Lemma expand_iblock_istep_run_None_rec sp ib: +Lemma normRTL_iblock_istep_run_None_rec sp ib: forall rs0 m0 o k (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) (CONT: match o with @@ -317,7 +326,7 @@ Lemma expand_iblock_istep_run_None_rec sp ib: iblock_istep_run tge sp k rs1 m1 = None | _ => True end), - iblock_istep_run tge sp (expand ib k) rs0 m0 = None. + iblock_istep_run tge sp (normRTL ib k) rs0 m0 = None. Proof. induction ib; simpl; try discriminate. @@ -359,29 +368,29 @@ Proof. + eapply IHib2; eauto. Qed. -Lemma expand_preserves_iblock_istep_run_None sp ib: +Lemma normRTL_preserves_iblock_istep_run_None sp ib: forall rs m, iblock_istep_run tge sp ib rs m = None - -> iblock_istep_run tge sp (expand ib (Bnop None)) rs m = None. + -> iblock_istep_run tge sp (normRTL ib (Bnop None)) rs m = None. Proof. - intros; eapply expand_iblock_istep_run_None_rec; eauto. + intros; eapply normRTL_iblock_istep_run_None_rec; eauto. simpl; auto. Qed. -Lemma expand_preserves_iblock_istep_run sp ib: +Lemma normRTL_preserves_iblock_istep_run sp ib: forall rs m, iblock_istep_run tge sp ib rs m = - iblock_istep_run tge sp (expand ib (Bnop None)) rs m. + iblock_istep_run tge sp (normRTL ib (Bnop None)) rs m. Proof. intros. destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP. - destruct o. symmetry. rewrite <- iblock_istep_run_equiv in *. - apply expand_iblock_istep_correct; auto. + apply normRTL_iblock_istep_correct; auto. - symmetry. - apply expand_preserves_iblock_istep_run_None; auto. + apply normRTL_preserves_iblock_istep_run_None; auto. Qed. Local Hint Constructors match_iblock: core. -Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst: +Lemma normRTL_matchiblock_rec_correct dupmap cfg ib pc isfst: forall opc1 (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2 (CONT: match opc1 with @@ -389,7 +398,7 @@ Lemma expand_matchiblock_rec_correct dupmap cfg ib pc isfst: match_iblock dupmap cfg false pc' k opc2 | None => opc2=opc1 end), - match_iblock dupmap cfg isfst pc (expand ib k) opc2. + match_iblock dupmap cfg isfst pc (normRTL ib k) opc2. Proof. induction 1; simpl; intros; eauto. { (* BF *) @@ -404,12 +413,12 @@ Proof. destruct opc0; econstructor. } Qed. -Lemma expand_matchiblock_correct dupmap cfg ib pc isfst opc: +Lemma normRTL_matchiblock_correct dupmap cfg ib pc isfst opc: match_iblock dupmap cfg isfst pc ib opc -> - match_iblock dupmap cfg isfst pc (expand ib (Bnop None)) opc. + match_iblock dupmap cfg isfst pc (normRTL ib (Bnop None)) opc. Proof. intros. - eapply expand_matchiblock_rec_correct; eauto. + eapply normRTL_matchiblock_rec_correct; eauto. destruct opc; simpl; auto. Qed. @@ -467,11 +476,10 @@ Proof. eexists; left; eexists; split. + repeat econstructor; eauto. apply iblock_istep_run_equiv in BTL_RUN; eauto. - + econstructor; apply expand_matchiblock_correct in MI. - econstructor; eauto. apply expand_correct; trivial. - intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. - econstructor. apply expand_preserves_iblock_istep_run. - eapply expand_entry_isnt_goto; eauto. + + econstructor; apply normRTL_matchiblock_correct in MI. + econstructor; eauto. eapply is_expand_normRTL_entry; eauto. + econstructor. apply normRTL_preserves_iblock_istep_run. + eapply normRTL_entry_isnt_goto; eauto. - (* Others *) exists (Some ib2); right; split. simpl; auto. @@ -536,11 +544,10 @@ Proof. pose symbols_preserved as SYMPRES. eapply eval_builtin_args_preserved; eauto. eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - * econstructor; eauto; apply expand_matchiblock_correct in MI. - { econstructor; eauto. apply expand_correct; trivial. - intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. - apply star_refl. apply expand_preserves_iblock_istep_run. } - eapply expand_entry_isnt_goto; eauto. + * econstructor; eauto; apply normRTL_matchiblock_correct in MI. + { econstructor; eauto. eapply is_expand_normRTL_entry; eauto. + apply star_refl. apply normRTL_preserves_iblock_istep_run. } + eapply normRTL_entry_isnt_goto; eauto. + (* Bjumptable *) exploit list_nth_z_rev_dupmap; eauto. intros (pc'0 & LNZ & DM). @@ -552,11 +559,10 @@ Proof. eexists; eexists; split. eapply iblock_istep_run_equiv in BTL_RUN. eapply BTL_RUN. econstructor; eauto. - * econstructor; eauto; apply expand_matchiblock_correct in MI. - { econstructor; eauto. apply expand_correct; trivial. - intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. - apply star_refl. apply expand_preserves_iblock_istep_run. } - eapply expand_entry_isnt_goto; eauto. + * econstructor; eauto; apply normRTL_matchiblock_correct in MI. + { econstructor; eauto. eapply is_expand_normRTL_entry; eauto. + apply star_refl. apply normRTL_preserves_iblock_istep_run. } + eapply normRTL_entry_isnt_goto; eauto. - (* mib_exit *) discriminate. - (* mib_seq *) @@ -601,8 +607,7 @@ Proof. intros; rewrite H12 in BTL_RUN. destruct b; eapply match_strong_state_simu; eauto. 1,3: inv H2; econstructor; eauto. - 1,3: inv IS_EXPD; [ inv H3; auto; inv H0 | inv H ]. - 3,5: inv IS_EXPD; [ inv H7; auto; inv H1 | inv H ]. + 1,3,5,7: inv IS_EXPD; auto; try discriminate. 1-4: eapply star_right; eauto. assert (measure bnot > 0) by apply measure_pos; lia. assert (measure bso > 0) by apply measure_pos; lia. @@ -660,14 +665,12 @@ Proof. eexists; left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. - * apply expand_matchiblock_correct in MI. + * apply normRTL_matchiblock_correct in MI. econstructor. econstructor; eauto. - apply expand_correct; trivial. - intros CONTRA; congruence. eapply expand_entry_isnt_bnop; eauto. - 3: eapply expand_entry_isnt_goto; eauto. + eapply is_expand_normRTL_entry; eauto. + 3: eapply normRTL_entry_isnt_goto; eauto. all: erewrite preserv_fnparams; eauto. - constructor. - apply expand_preserves_iblock_istep_run. + constructor. apply normRTL_preserves_iblock_istep_run. + (* External function *) inv TRANSF. eexists; left; eexists; split. @@ -682,12 +685,11 @@ Proof. apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. eexists; left; eexists; split. + eapply exec_return. - + apply expand_matchiblock_correct in MI. + + apply normRTL_matchiblock_correct in MI. econstructor. econstructor; eauto. - apply expand_correct; trivial. - constructor. congruence. eapply expand_entry_isnt_bnop; eauto. - econstructor. apply expand_preserves_iblock_istep_run. - eapply expand_entry_isnt_goto; eauto. + eapply is_expand_normRTL_entry; eauto. + econstructor. apply normRTL_preserves_iblock_istep_run. + eapply normRTL_entry_isnt_goto; eauto. Qed. Local Hint Resolve plus_one star_refl: core. -- cgit From 0efe7783c50d72858352fda93d30e0f52792d658 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 24 May 2021 10:07:52 +0200 Subject: Moving common tools, adding liveness input/output information to BTL generation oracle --- scheduling/BTLaux.ml | 8 +++++++- scheduling/BTLrenumber.ml | 2 +- scheduling/BTLtoRTLaux.ml | 2 +- scheduling/PrintBTL.ml | 2 ++ scheduling/RTLpathLivegenaux.ml | 38 +------------------------------------- scheduling/RTLpathScheduleraux.ml | 2 +- scheduling/RTLtoBTLaux.ml | 28 +++++++++++++++++++--------- 7 files changed, 32 insertions(+), 50 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLaux.ml b/scheduling/BTLaux.ml index 863afdf0..e8758f61 100644 --- a/scheduling/BTLaux.ml +++ b/scheduling/BTLaux.ml @@ -1,3 +1,9 @@ +open Registers + type inst_info = { mutable inumb : int; mutable pcond : bool option } -type block_info = { mutable bnumb : int; mutable visited: bool } +type block_info = { + mutable bnumb : int; + mutable visited : bool; + mutable output_regs : Regset.t; +} diff --git a/scheduling/BTLrenumber.ml b/scheduling/BTLrenumber.ml index dd6baa89..f9cacd6a 100644 --- a/scheduling/BTLrenumber.ml +++ b/scheduling/BTLrenumber.ml @@ -1,7 +1,7 @@ (* XXX uncomment open !BTL open DebugPrint -open AuxTools +open RTLcommonaux open Maps*) (** A quick note on the BTL renumber algorithm diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 36d3e6a4..521f6ef2 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -2,7 +2,7 @@ open Maps open BTL open RTL open Camlcoq -open AuxTools +open RTLcommonaux open DebugPrint open BTLaux diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 0ed3981d..502565c2 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -101,6 +101,8 @@ let print_btl_code pp btl is_rec = fprintf pp "\n"; List.iter (fun (n, ibf) -> + fprintf pp "[BTL Liveness]\n"; + print_regset ibf.input_regs; fprintf pp "[BTL block %d]\n" (P.to_int n); print_iblock pp is_rec " " ibf.entry; fprintf pp "\n") diff --git a/scheduling/RTLpathLivegenaux.ml b/scheduling/RTLpathLivegenaux.ml index c9bb94d3..976ddc16 100644 --- a/scheduling/RTLpathLivegenaux.ml +++ b/scheduling/RTLpathLivegenaux.ml @@ -4,10 +4,8 @@ open Registers open Maps open Camlcoq open Datatypes -open Kildall -open Lattice open DebugPrint -open AuxTools +open RTLcommonaux let rec list_to_regset = function | [] -> Regset.empty @@ -77,40 +75,6 @@ let get_path_map code entry join_points = !path_map end -let transfer f pc after = let open Liveness in - match PTree.get pc f.fn_code with - | Some i -> - (match i with - | Inop _ -> after - | Iop (_, args, res, _) -> - reg_list_live args (Regset.remove res after) - | Iload (_, _, _, args, dst, _) -> - reg_list_live args (Regset.remove dst after) - | Istore (_, _, args, src, _) -> - reg_list_live args (Regset.add src after) - | Icall (_, ros, args, res, _) -> - reg_list_live args (reg_sum_live ros (Regset.remove res after)) - | Itailcall (_, ros, args) -> - reg_list_live args (reg_sum_live ros Regset.empty) - | Ibuiltin (_, args, res, _) -> - reg_list_live (AST.params_of_builtin_args args) - (reg_list_dead (AST.params_of_builtin_res res) after) - | Icond (_, args, _, _, _) -> - reg_list_live args after - | Ijumptable (arg, _) -> - Regset.add arg after - | Ireturn optarg -> - reg_option_live optarg Regset.empty) - | None -> Regset.empty - -module RegsetLat = LFSet(Regset) - -module DS = Backward_Dataflow_Solver(RegsetLat)(NodeSetBackward) - -let analyze f = - let liveouts = get_some @@ DS.fixpoint f.fn_code successors_instr (transfer f) in - PTree.map (fun n _ -> let lo = PMap.get n liveouts in transfer f n lo) f.fn_code - (** OLD CODE - If needed to have our own kildall let transfer after = let open Liveness in function diff --git a/scheduling/RTLpathScheduleraux.ml b/scheduling/RTLpathScheduleraux.ml index 70a0c507..3db25d82 100644 --- a/scheduling/RTLpathScheduleraux.ml +++ b/scheduling/RTLpathScheduleraux.ml @@ -7,7 +7,7 @@ open RTL open Maps open Registers open ExpansionOracle -open AuxTools +open RTLcommonaux let config = Machine.config diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 4be2b180..07e7a9d9 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -3,7 +3,7 @@ open RTL open BTL open Registers open PrintBTL -open AuxTools +open RTLcommonaux open DebugPrint open BTLaux @@ -13,7 +13,7 @@ let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond } let def_iinfo = { inumb = undef_node; pcond = None } -let mk_binfo _bnumb = { bnumb = _bnumb; visited = false } +let mk_binfo _bnumb _output_regs = { bnumb = _bnumb; visited = false; output_regs = _output_regs } let encaps_final inst osucc = match inst with @@ -52,7 +52,7 @@ let translate_inst (iinfo: BTL.inst_info) inst is_final = in if is_final then encaps_final btli !osucc else btli -let translate_function code entry join_points = +let translate_function code entry join_points liveness = let reached = ref (PTree.map (fun n i -> false) code) in let btl_code = ref PTree.empty in let rec build_btl_tree e = @@ -60,6 +60,7 @@ let translate_function code entry join_points = else ( reached := PTree.set e true !reached; let next_nodes = ref [] in + let last = ref None in let rec build_btl_block n = let inst = get_some @@ PTree.get n code in let psucc = predicted_successor inst in @@ -90,13 +91,21 @@ let translate_function code entry join_points = | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; + last := Some inst; next_nodes := !next_nodes @ successors_inst inst; translate_inst iinfo inst true in let ib = build_btl_block e in let succs = !next_nodes in - let bi = mk_binfo (p2i e) in - let ibf = { entry = ib; input_regs = Regset.empty; binfo = bi } in + + let inputs = get_some @@ PTree.get e liveness in + let blk_last_successors = successors_inst (get_some @@ !last) in + let list_input_regs = List.map ( + fun n -> get_some @@ PTree.get n liveness + ) blk_last_successors in + let outputs = List.fold_left Regset.union Regset.empty list_input_regs in + let bi = mk_binfo (p2i e) outputs in + let ibf = { entry = ib; input_regs = inputs; binfo = bi } in btl_code := PTree.set e ibf !btl_code; List.iter build_btl_tree succs) in @@ -104,21 +113,22 @@ let translate_function code entry join_points = !btl_code let rtl2btl (f : RTL.coq_function) = - (*debug_flag := true;*) let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in - let btl = translate_function code entry join_points in + let liveness = analyze f in + let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in debug "Entry %d\n" (p2i entry); debug "RTL Code: "; - (*print_code code;*) + print_code code; + (*debug_flag := true;*) debug "BTL Code: "; print_btl_code stderr btl true; + (*debug_flag := false;*) debug "Dupmap:\n"; print_ptree print_pint dm; debug "Join points: "; print_true_nodes join_points; debug "\n"; - (*debug_flag := false;*) ((btl, entry), dm) -- cgit From a6006df63f0d03cc223d13834e81a71651513fbe Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 24 May 2021 17:39:44 +0200 Subject: a draft frontend for prepass --- scheduling/BTLScheduleraux.ml | 221 ++++++++++++++++++++++++++++++++++++++++++ scheduling/BTLaux.ml | 8 +- scheduling/BTLtoRTLaux.ml | 4 +- scheduling/RTLtoBTLaux.ml | 19 ++-- 4 files changed, 231 insertions(+), 21 deletions(-) create mode 100644 scheduling/BTLScheduleraux.ml (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml new file mode 100644 index 00000000..699538ca --- /dev/null +++ b/scheduling/BTLScheduleraux.ml @@ -0,0 +1,221 @@ +open AST +open Maps +open Registers +open BTL +open DebugPrint +open RTLcommonaux +open InstructionScheduler +open PrepassSchedulingOracleDeps + +let use_alias_analysis () = false + +let build_constraints_and_resources (opweights : opweights) insts btl = + let last_reg_reads : int list PTree.t ref = ref PTree.empty + and last_reg_write : (int * int) PTree.t ref = ref PTree.empty + and last_mem_reads : int list ref = ref [] + and last_mem_write : int option ref = ref None + and last_branch : int option ref = ref None + and last_non_pipelined_op : int array = + Array.make opweights.nr_non_pipelined_units (-1) + and latency_constraints : latency_constraint list ref = ref [] + and resources = ref [] in + let add_constraint instr_from instr_to latency = + assert (instr_from <= instr_to); + assert (latency >= 0); + if instr_from = instr_to then + if latency = 0 then () + else + failwith "PrepassSchedulingOracle.get_dependencies: negative self-loop" + else + latency_constraints := + { instr_from; instr_to; latency } :: !latency_constraints + and get_last_reads reg = + match PTree.get reg !last_reg_reads with Some l -> l | None -> [] + in + let add_input_mem i = + if not (use_alias_analysis ()) then ( + (* Read after write *) + (match !last_mem_write with None -> () | Some j -> add_constraint j i 1); + last_mem_reads := i :: !last_mem_reads) + and add_output_mem i = + if not (use_alias_analysis ()) then ( + (* Write after write *) + (match !last_mem_write with None -> () | Some j -> add_constraint j i 1); + (* Write after read *) + List.iter (fun j -> add_constraint j i 0) !last_mem_reads; + last_mem_write := Some i; + last_mem_reads := []) + and add_input_reg i reg = + (* Read after write *) + (match PTree.get reg !last_reg_write with + | None -> () + | Some (j, latency) -> add_constraint j i latency); + last_reg_reads := PTree.set reg (i :: get_last_reads reg) !last_reg_reads + and add_output_reg i latency reg = + (* Write after write *) + (match PTree.get reg !last_reg_write with + | None -> () + | Some (j, _) -> add_constraint j i 1); + (* Write after read *) + List.iter (fun j -> add_constraint j i 0) (get_last_reads reg); + last_reg_write := PTree.set reg (i, latency) !last_reg_write; + last_reg_reads := PTree.remove reg !last_reg_reads + in + let add_input_regs i regs = List.iter (add_input_reg i) regs + and irreversible_action i = + match !last_branch with None -> () | Some j -> add_constraint j i 1 + in + let set_branch i = + irreversible_action i; + last_branch := Some i + and add_non_pipelined_resources i resources = + Array.iter2 + (fun latency last -> + if latency >= 0 && last >= 0 then add_constraint last i latency) + resources last_non_pipelined_op; + Array.iteri + (fun rsc latency -> if latency >= 0 then last_non_pipelined_op.(rsc) <- i) + resources + in + Array.iteri + (fun i inst -> + (* TODO gourdinl liveins for Bcond *) + match inst with + | Bnop _ -> + let rs = Array.map (fun _ -> 0) opweights.pipelined_resource_bounds in + resources := rs :: !resources + | Bop (op, lr, rd, _) -> + add_non_pipelined_resources i + (opweights.non_pipelined_resources_of_op op (List.length lr)); + if Op.is_trapping_op op then irreversible_action i; + add_input_regs i lr; + add_output_reg i (opweights.latency_of_op op (List.length lr)) rd; + let rs = opweights.resources_of_op op (List.length lr) in + resources := rs :: !resources + | Bload (trap, chk, addr, lr, rd, _) -> + if trap = TRAP then irreversible_action i; + add_input_mem i; + add_input_regs i lr; + add_output_reg i + (opweights.latency_of_load trap chk addr (List.length lr)) + rd; + let rs = opweights.resources_of_load trap chk addr (List.length lr) in + resources := rs :: !resources + | Bstore (chk, addr, lr, src, _) -> + irreversible_action i; + add_input_regs i lr; + add_input_reg i src; + add_output_mem i; + let rs = opweights.resources_of_store chk addr (List.length lr) in + resources := rs :: !resources + | Bcond (cond, lr, BF (Bgoto s, _), Bnop _, _) -> + (* TODO gourdinl test with/out this line *) + let live = (get_some @@ PTree.get s btl).input_regs in + add_input_regs i (Regset.elements live); + set_branch i; + add_input_mem i; + add_input_regs i lr; + let rs = opweights.resources_of_cond cond (List.length lr) in + resources := rs :: !resources + | Bcond (_, _, _, _, _) -> + failwith "get_simple_dependencies: invalid Bcond" + | BF (_, _) -> failwith "get_simple_dependencies: BF" + | Bseq (_, _) -> failwith "get_simple_dependencies: Bseq") + insts; + (!latency_constraints, Array.of_list (List.rev !resources)) + +let define_problem (opweights : opweights) ibf btl = + let simple_deps, resources = + build_constraints_and_resources opweights ibf btl + in + { + max_latency = -1; + resource_bounds = opweights.pipelined_resource_bounds; + instruction_usages = resources; + latency_constraints = + (* if (use_alias_analysis ()) + then (get_alias_dependencies seqa) @ simple_deps + else *) + simple_deps; + } + +let zigzag_scheduler problem early_ones = + let nr_instructions = get_nr_instructions problem in + assert (nr_instructions = Array.length early_ones); + match list_scheduler problem with + | Some fwd_schedule -> + let fwd_makespan = fwd_schedule.(Array.length fwd_schedule - 1) in + let constraints' = ref problem.latency_constraints in + Array.iteri + (fun i is_early -> + if is_early then + constraints' := + { + instr_from = i; + instr_to = nr_instructions; + latency = fwd_makespan - fwd_schedule.(i); + } + :: !constraints') + early_ones; + validated_scheduler reverse_list_scheduler + { problem with latency_constraints = !constraints' } + | None -> None + +let prepass_scheduler_by_name name problem early_ones = + match name with + | "zigzag" -> zigzag_scheduler problem early_ones + | _ -> scheduler_by_name name problem + +let schedule_sequence insts btl = + let opweights = OpWeights.get_opweights () in + try + if Array.length insts <= 1 then None + else + let nr_instructions = Array.length insts in + let problem = define_problem opweights insts btl in + match + prepass_scheduler_by_name + !Clflags.option_fprepass_sched + problem + (Array.map + (fun inst -> + match inst with Bcond (_, _, _, _, _) -> true | _ -> false) + insts) + with + | None -> + Printf.printf "no solution in prepass scheduling\n"; + None + | Some solution -> + let positions = Array.init nr_instructions (fun i -> i) in + Array.sort + (fun i j -> + let si = solution.(i) and sj = solution.(j) in + if si < sj then -1 else if si > sj then 1 else i - j) + positions; + Some positions + with Failure s -> + Printf.printf "failure in prepass scheduling: %s\n" s; + None + +let flatten_blk_basics ibf = + let ib = ibf.entry in + let rec traverse_blk ib = + match ib with + | BF (_, _) + | Bcond (_, _, BF (Bgoto _, _), BF (Bgoto _, _), _) -> [] + | Bseq (ib1, ib2) -> + traverse_blk ib1 @ traverse_blk ib2 + | _ -> [ib] + in + Array.of_list (traverse_blk ib) + +let btl_scheduler btl entry = + List.iter (fun (n,ibf) -> + let bseq = flatten_blk_basics ibf in + match schedule_sequence bseq btl with + | Some positions -> + Array.iter (fun p -> debug "%d " p) positions + | None -> () + ) (PTree.elements btl); + (*let seqs = get_sequences seqs in*) + () diff --git a/scheduling/BTLaux.ml b/scheduling/BTLaux.ml index e8758f61..ca34c21c 100644 --- a/scheduling/BTLaux.ml +++ b/scheduling/BTLaux.ml @@ -1,9 +1,3 @@ -open Registers - type inst_info = { mutable inumb : int; mutable pcond : bool option } -type block_info = { - mutable bnumb : int; - mutable visited : bool; - mutable output_regs : Regset.t; -} +type block_info = { mutable bnumb : int; mutable visited : bool } diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 521f6ef2..42c20942 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -63,9 +63,9 @@ let translate_function code entry = Some ( Iload (trap, chk, addr, lr, rd, get_ib_num (get_some k)), get_inumb iinfo ) - | Bstore (chk, addr, lr, rd, iinfo) -> + | Bstore (chk, addr, lr, src, iinfo) -> Some - ( Istore (chk, addr, lr, rd, get_ib_num (get_some k)), + ( Istore (chk, addr, lr, src, get_ib_num (get_some k)), get_inumb iinfo ) | BF (Bcall (sign, fn, lr, rd, s), iinfo) -> next_nodes := s :: !next_nodes; diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 07e7a9d9..e932d766 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -1,11 +1,11 @@ open Maps open RTL open BTL -open Registers open PrintBTL open RTLcommonaux open DebugPrint open BTLaux +open BTLScheduleraux let undef_node = -1 @@ -13,7 +13,7 @@ let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond } let def_iinfo = { inumb = undef_node; pcond = None } -let mk_binfo _bnumb _output_regs = { bnumb = _bnumb; visited = false; output_regs = _output_regs } +let mk_binfo _bnumb = { bnumb = _bnumb; visited = false } let encaps_final inst osucc = match inst with @@ -33,9 +33,9 @@ let translate_inst (iinfo: BTL.inst_info) inst is_final = | Iload (trap, chk, addr, lr, rd, s) -> osucc := Some s; Bload (trap, chk, addr, lr, rd, iinfo) - | Istore (chk, addr, lr, rd, s) -> + | Istore (chk, addr, lr, src, s) -> osucc := Some s; - Bstore (chk, addr, lr, rd, iinfo) + Bstore (chk, addr, lr, src, iinfo) | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s), iinfo) | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr), iinfo) | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s), iinfo) @@ -60,7 +60,6 @@ let translate_function code entry join_points liveness = else ( reached := PTree.set e true !reached; let next_nodes = ref [] in - let last = ref None in let rec build_btl_block n = let inst = get_some @@ PTree.get n code in let psucc = predicted_successor inst in @@ -91,7 +90,6 @@ let translate_function code entry join_points liveness = | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; - last := Some inst; next_nodes := !next_nodes @ successors_inst inst; translate_inst iinfo inst true in @@ -99,12 +97,7 @@ let translate_function code entry join_points liveness = let succs = !next_nodes in let inputs = get_some @@ PTree.get e liveness in - let blk_last_successors = successors_inst (get_some @@ !last) in - let list_input_regs = List.map ( - fun n -> get_some @@ PTree.get n liveness - ) blk_last_successors in - let outputs = List.fold_left Regset.union Regset.empty list_input_regs in - let bi = mk_binfo (p2i e) outputs in + let bi = mk_binfo (p2i e) in let ibf = { entry = ib; input_regs = inputs; binfo = bi } in btl_code := PTree.set e ibf !btl_code; List.iter build_btl_tree succs) @@ -119,6 +112,8 @@ let rtl2btl (f : RTL.coq_function) = let liveness = analyze f in let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in + (* TODO gourdinl move elsewhere *) + btl_scheduler btl entry; debug "Entry %d\n" (p2i entry); debug "RTL Code: "; print_code code; -- cgit From 51ec982b84851ac3526a0cb3f22e41f974b2d592 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 24 May 2021 22:41:35 +0200 Subject: tiny simplifications in RTLtoBTLproof --- scheduling/BTL.v | 92 +------ scheduling/RTLtoBTLproof.v | 599 ++++++++++++++++++++++++++------------------- 2 files changed, 342 insertions(+), 349 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 0f9ef44f..97c54021 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -810,94 +810,4 @@ Definition is_goto (ib: iblock): bool := match ib with | BF (Bgoto _) _ => true | _ => false - end. - -Definition is_RTLatom (ib: iblock): bool := - match ib with - | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None => false - | _ => true - end. - -Definition Bnop_dec k: { k=Bnop None} + { k<>(Bnop None) }. -Proof. - destruct k; simpl; try destruct oiinfo; - intuition congruence. -Defined. - -(** Is expand property to only support atomic instructions on the left part of a Bseq: - The tree is right-expanded, and block are normalized to simplify RTL simulation *) -Inductive is_expand: iblock -> Prop := - | exp_Bseq ib1 ib2: - is_RTLatom ib1 = true -> - is_expand ib2 -> - is_expand (Bseq ib1 ib2) - | exp_Bcond cond args ib1 ib2 i: - is_expand ib1 -> - is_expand ib2 -> - is_expand (Bcond cond args ib1 ib2 i) - | exp_others ib: - is_RTLatom ib = true -> - is_expand ib - . -Local Hint Constructors is_expand: core. - -(** The weak version with more specific hypothesis for conditions *) -Inductive is_wexpand: iblock -> Prop := - | wexp_Bseq ib1 ib2: - is_RTLatom ib1 = true -> - is_wexpand ib2 -> - is_wexpand (Bseq ib1 ib2) - | wexp_Bcond cond args ib1 ib2 iinfo: - (ib1 <> Bnop None -> is_wexpand ib1) -> - (ib2 <> Bnop None -> is_wexpand ib2) -> - is_wexpand (Bcond cond args ib1 ib2 iinfo) - | wexp_others ib: - is_RTLatom ib = true -> - is_wexpand ib - . -Local Hint Constructors is_wexpand: core. - -Fixpoint normRTL (ib: iblock) (k: iblock): iblock := - match ib with - | Bseq ib1 ib2 => normRTL ib1 (normRTL ib2 k) - | Bcond cond args ib1 ib2 iinfo => - Bcond cond args (normRTL ib1 k) (normRTL ib2 k) iinfo - | BF fin iinfo => BF fin iinfo - | Bnop None => k - | ib => - if Bnop_dec k then ib else Bseq ib k - end. - -Lemma normRTL_correct ib: forall k, - (k <> (Bnop None) -> is_wexpand k) -> - (normRTL ib k) <> Bnop None -> - is_wexpand (normRTL ib k). -Proof. - induction ib; simpl; intros; try autodestruct; auto. - intros; destruct (Bnop_dec k); auto. -Qed. - -Lemma is_join_opt_None {A} (opc1 opc2: option A): - is_join_opt opc1 opc2 None -> opc1 = None /\ opc2 = None. -Proof. - intros X. inv X; auto. -Qed. - -Lemma match_iblock_None_not_Bnop dupmap cfg isfst pc ib: - match_iblock dupmap cfg isfst pc ib None -> ib <> Bnop None. -Proof. - intros X; inv X; try congruence. -Qed. -Local Hint Resolve match_iblock_None_not_Bnop: core. - -Lemma is_wexpand_expand dupmap cfg ib: - is_wexpand ib -> - forall isfst pc - (MIB: match_iblock dupmap cfg isfst pc ib None), - is_expand ib. -Proof. - induction 1; simpl; intros; auto; try (inv MIB); eauto. - (* Bcond *) - destruct (is_join_opt_None opc1 opc2); subst; eauto. - econstructor; eauto. -Qed. + end. \ No newline at end of file diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index cd6c4dae..7cd8e47d 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -6,6 +6,269 @@ Require Import Errors Linking RTLtoBTL. Require Import Linking. +(** * Normalization of BTL iblock for simulation of RTL + +Below [normRTL] normalizes the representation of BTL blocks, +in order to represent as sequences of RTL instructions. + +This eases the + +*) + +Definition is_RTLatom (ib: iblock): bool := + match ib with + | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None => false + | _ => true + end. + +Definition is_RTLbasic (ib: iblock): bool := + match ib with + | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None | BF _ _ => false + | _ => true + end. + + +Definition Bnop_dec k: { k=Bnop None} + { k<>(Bnop None) }. +Proof. + destruct k; simpl; try destruct oiinfo; + intuition congruence. +Defined. + +(** The strict [is_normRTL] property specifying the ouput of [normRTL] below *) +Inductive is_normRTL: iblock -> Prop := + | norm_Bseq ib1 ib2: + is_RTLbasic ib1 = true -> + is_normRTL ib2 -> + is_normRTL (Bseq ib1 ib2) + | norm_Bcond cond args ib1 ib2 i: + is_normRTL ib1 -> + is_normRTL ib2 -> + is_normRTL (Bcond cond args ib1 ib2 i) + | norm_others ib: + is_RTLatom ib = true -> + is_normRTL ib + . +Local Hint Constructors is_normRTL: core. + +(** Weaker version with less restrictive hypothesis for conditions *) +Inductive is_wnormRTL: iblock -> Prop := + | wnorm_Bseq ib1 ib2: + is_RTLbasic ib1 = true -> + is_wnormRTL ib2 -> + is_wnormRTL (Bseq ib1 ib2) + | wnorm_Bcond cond args ib1 ib2 iinfo: + (ib1 <> Bnop None -> is_wnormRTL ib1) -> + (ib2 <> Bnop None -> is_wnormRTL ib2) -> + is_wnormRTL (Bcond cond args ib1 ib2 iinfo) + | wnorm_others ib: + is_RTLatom ib = true -> + is_wnormRTL ib + . +Local Hint Constructors is_wnormRTL: core. + +(* NB: [k] is a "continuation" (e.g. semantically [normRTLrec ib k] is like [BSeq ib k]) *) +Fixpoint normRTLrec (ib: iblock) (k: iblock): iblock := + match ib with + | Bseq ib1 ib2 => normRTLrec ib1 (normRTLrec ib2 k) + | Bcond cond args ib1 ib2 iinfo => + Bcond cond args (normRTLrec ib1 k) (normRTLrec ib2 k) iinfo + | BF fin iinfo => BF fin iinfo + | Bnop None => k + | ib => + if Bnop_dec k then ib else Bseq ib k + end. + +Definition normRTL ib := normRTLrec ib (Bnop None). + +Lemma normRTLrec_wcorrect ib: forall k, + (k <> (Bnop None) -> is_wnormRTL k) -> + (normRTLrec ib k) <> Bnop None -> + is_wnormRTL (normRTLrec ib k). +Proof. + induction ib; simpl; intros; repeat autodestruct; auto. +Qed. + +Lemma normRTL_wcorrect ib: + (normRTL ib) <> Bnop None -> + is_wnormRTL (normRTL ib). +Proof. + intros; eapply normRTLrec_wcorrect; eauto. +Qed. + +Lemma is_join_opt_None {A} (opc1 opc2: option A): + is_join_opt opc1 opc2 None -> opc1 = None /\ opc2 = None. +Proof. + intros X. inv X; auto. +Qed. + +Lemma match_iblock_None_not_Bnop dupmap cfg isfst pc ib: + match_iblock dupmap cfg isfst pc ib None -> ib <> Bnop None. +Proof. + intros X; inv X; try congruence. +Qed. +Local Hint Resolve match_iblock_None_not_Bnop: core. + +Lemma is_wnormRTL_normRTL dupmap cfg ib: + is_wnormRTL ib -> + forall isfst pc + (MIB: match_iblock dupmap cfg isfst pc ib None), + is_normRTL ib. +Proof. + induction 1; simpl; intros; auto; try (inv MIB); eauto. + (* Bcond *) + destruct (is_join_opt_None opc1 opc2); subst; eauto. + econstructor; eauto. +Qed. + +Local Hint Constructors iblock_istep: core. +Lemma normRTLrec_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin1: + forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) + k ofin2 rs2 m2 + (CONT: match ofin1 with + | None => iblock_istep tge sp rs1 m1 k rs2 m2 ofin2 + | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 + end), + iblock_istep tge sp rs0 m0 (normRTLrec ib k) rs2 m2 ofin2. +Proof. + induction 1; simpl; intros; intuition subst; eauto. + { (* Bnop *) + autodestruct; intros OI; inv OI; destruct (Bnop_dec k); subst; + try (inv CONT; constructor; fail); repeat econstructor; eauto. } + 1-3: (* Bop, Bload, Bstore *) + intros; destruct (Bnop_dec k); subst; repeat econstructor; eauto; + inv CONT; econstructor; eauto. + - (* Bcond *) + destruct ofin; intros; + econstructor; eauto; + destruct b; eapply IHISTEP; eauto. +Qed. + +Lemma normRTL_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin: + iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> + iblock_istep tge sp rs0 m0 (normRTL ib) rs1 m1 ofin. +Proof. + intros; eapply normRTLrec_iblock_istep_correct; eauto. + destruct ofin; simpl; auto. +Qed. + +Lemma normRTLrec_iblock_istep_run_None tge sp ib: + forall rs0 m0 o k + (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) + (CONT: match o with + | Some (out rs1 m1 ofin) => + ofin = None /\ + iblock_istep_run tge sp k rs1 m1 = None + | _ => True + end), + iblock_istep_run tge sp (normRTLrec ib k) rs0 m0 = None. +Proof. + induction ib; simpl; + try discriminate. + - (* BF *) + intros; destruct o; try discriminate; simpl in *. + inv ISTEP. destruct CONT as [HO ISTEP]; inv HO. + - (* Bnop *) + intros; destruct o; inv ISTEP; destruct oiinfo, CONT; auto. + destruct (Bnop_dec k); subst; repeat econstructor; eauto. + - (* Bop *) + intros; destruct o; destruct (Bnop_dec k); + destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; + simpl; rewrite EVAL; auto; destruct CONT as [HO ISTEP]; + trivial. + - (* Bload *) + intros; destruct o; destruct (Bnop_dec k); + destruct (trap) eqn:TRAP; + try destruct (eval_addressing _ _ _ _) eqn:EVAL; + try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; + simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; + destruct CONT as [HO ISTEP]; trivial. + - (* Bstore *) + intros; destruct o; destruct (Bnop_dec k); + destruct (eval_addressing _ _ _ _) eqn:EVAL; + try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; + simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; + destruct CONT as [HO ISTEP]; inv HO; trivial. + - (* Bseq *) + intros. + eapply IHib1; eauto. + destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1; try auto. + destruct o0; simpl in *. destruct _fin; inv ISTEP. + + destruct CONT as [CONTRA _]; inv CONTRA. + + split; auto. eapply IHib2; eauto. + - (* Bcond *) + intros; destruct (eval_condition _ _ _); trivial. + destruct b. + + eapply IHib1; eauto. + + eapply IHib2; eauto. +Qed. + +Lemma normRTL_preserves_iblock_istep_run_None tge sp ib: + forall rs m, iblock_istep_run tge sp ib rs m = None + -> iblock_istep_run tge sp (normRTL ib) rs m = None. +Proof. + intros; eapply normRTLrec_iblock_istep_run_None; eauto. + simpl; auto. +Qed. + +Lemma normRTL_preserves_iblock_istep_run tge sp ib: + forall rs m, iblock_istep_run tge sp ib rs m = + iblock_istep_run tge sp (normRTL ib) rs m. +Proof. + intros. + destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP. + - destruct o. symmetry. + rewrite <- iblock_istep_run_equiv in *. + apply normRTL_iblock_istep_correct; auto. + - symmetry. + apply normRTL_preserves_iblock_istep_run_None; auto. +Qed. + +Local Hint Constructors match_iblock: core. +Lemma normRTLrec_matchiblock_correct dupmap cfg ib pc isfst: + forall opc1 + (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2 + (CONT: match opc1 with + | Some pc' => + match_iblock dupmap cfg false pc' k opc2 + | None => opc2=opc1 + end), + match_iblock dupmap cfg isfst pc (normRTLrec ib k) opc2. +Proof. + induction 1; simpl; intros; eauto. + { (* BF *) + inv CONT; econstructor; eauto. } + 1-4: (* Bnop, Bop, Bload, Bstore *) + destruct (Bnop_dec k); subst; eauto; inv CONT; econstructor; eauto. + { (* Bgoto *) + intros; inv CONT; apply mib_exit; auto. } + { (* Bcond *) + intros. inv H0; + econstructor; eauto; try econstructor. + destruct opc0; econstructor. } +Qed. + +Lemma normRTL_matchiblock_correct dupmap cfg ib pc isfst opc: + match_iblock dupmap cfg isfst pc ib opc -> + match_iblock dupmap cfg isfst pc (normRTL ib) opc. +Proof. + intros. + eapply normRTLrec_matchiblock_correct; eauto. + destruct opc; simpl; auto. +Qed. + +Lemma is_normRTL_correct dupmap cfg ib pc + (MI : match_iblock dupmap cfg true pc ib None): + is_normRTL (normRTL ib). +Proof. + exploit normRTL_matchiblock_correct; eauto. + intros MI2. + eapply is_wnormRTL_normRTL; eauto. + apply normRTL_wcorrect; try congruence. + inv MI2; discriminate. +Qed. + +(** * Matching relation on functions *) + Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); @@ -106,7 +369,7 @@ Inductive match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 (ATpc0: (fn_code f')!pcB0 = Some ib0) (DUPLIC: dupmap!pcB0 = Some pcR0) (MIB: match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None) - (IS_EXPD: is_expand ib) + (IS_EXPD: is_normRTL ib) (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs1 m1)) (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs1 m1) : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst @@ -231,43 +494,30 @@ Proof. induction ib; simpl; auto; lia. Qed. -(* TODO gourdinl remove useless lemma? *) -Lemma entry_isnt_goto dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (entry ib) None -> - is_goto (entry ib) = false. +Lemma match_iblock_true_isnt_goto dupmap cfg pc ib opc: + match_iblock dupmap cfg true pc ib opc -> + is_goto ib = false. Proof. - intros. - destruct (entry ib); trivial. - destruct fi; trivial. inv H. inv H5. + intros MIB; inversion MIB as [d1 d2 d3 d4 d5 H H0| | | | | | | |]; subst; simpl; try congruence. + inv H0; congruence. Qed. -Lemma normRTL_entry_isnt_goto dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None -> - is_goto (normRTL (entry ib) (Bnop None)) = false. +Local Hint Resolve match_iblock_true_isnt_goto normRTL_preserves_iblock_istep_run star_refl star_right: core. +Local Hint Constructors match_strong_state RTL.step: core. + +(** At entry in a block: we init [match_states] on [normRTL] to normalize the block *) +Lemma match_states_entry dupmap st f sp pc ib rs m st' f' pc' + (STACKS : list_forall2 match_stackframes st st') + (TRANSF : match_function dupmap f f') + (FN : (fn_code f') ! pc' = Some ib) + (MI : match_iblock dupmap (RTL.fn_code f) true pc (entry ib) None) + (DUP : dupmap ! pc' = Some pc): + match_states (Some (normRTL (entry ib))) (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m). Proof. - destruct (is_goto (normRTL (entry ib) (Bnop None))) eqn:EQG; - destruct (normRTL (entry ib) (Bnop None)); - try destruct fi; try discriminate; trivial. - intros; inv H; inv H5. -Qed. - -Lemma normRTL_entry_isnt_bnop dupmap f pc ib: - match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None -> - normRTL (entry ib) (Bnop None) <> Bnop None. -Proof. - destruct (normRTL (entry ib) (Bnop None)) eqn:EQG; - try destruct fi; try discriminate; trivial. - intros; inv H; inv H5. -Qed. - -Lemma is_expand_normRTL_entry dupmap f ib pc - (MI : match_iblock dupmap (RTL.fn_code f) true pc (normRTL (entry ib) (Bnop None)) None): - is_expand (normRTL (entry ib) (Bnop None)). -Proof. - eapply is_wexpand_expand; eauto. - apply normRTL_correct; try congruence. - eapply normRTL_entry_isnt_bnop; eauto. + exploit is_normRTL_correct; eauto. + econstructor; eauto; apply normRTL_matchiblock_correct in MI; eauto. Qed. +Local Hint Resolve match_states_entry: core. Lemma list_nth_z_rev_dupmap: forall dupmap ln ln' (pc pc': node) val, @@ -286,141 +536,6 @@ Proof. intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence. Qed. -Local Hint Constructors iblock_istep: core. -Lemma normRTL_iblock_istep_rec_correct sp ib rs0 m0 rs1 m1 ofin1: - forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1) - k ofin2 rs2 m2 - (CONT: match ofin1 with - | None => iblock_istep tge sp rs1 m1 k rs2 m2 ofin2 - | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1 - end), - iblock_istep tge sp rs0 m0 (normRTL ib k) rs2 m2 ofin2. -Proof. - induction 1; simpl; intros; intuition subst; eauto. - { (* Bnop *) - autodestruct; intros OI; inv OI; destruct (Bnop_dec k); subst; - try (inv CONT; constructor; fail); repeat econstructor; eauto. } - 1-3: (* Bop, Bload, Bstore *) - intros; destruct (Bnop_dec k); subst; repeat econstructor; eauto; - inv CONT; econstructor; eauto. - - (* Bcond *) - destruct ofin; intros; - econstructor; eauto; - destruct b; eapply IHISTEP; eauto. -Qed. - -Lemma normRTL_iblock_istep_correct sp ib rs0 m0 rs1 m1 ofin: - iblock_istep tge sp rs0 m0 ib rs1 m1 ofin -> - iblock_istep tge sp rs0 m0 (normRTL ib (Bnop None)) rs1 m1 ofin. -Proof. - intros; eapply normRTL_iblock_istep_rec_correct; eauto. - destruct ofin; simpl; auto. -Qed. - -Lemma normRTL_iblock_istep_run_None_rec sp ib: - forall rs0 m0 o k - (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) - (CONT: match o with - | Some (out rs1 m1 ofin) => - ofin = None /\ - iblock_istep_run tge sp k rs1 m1 = None - | _ => True - end), - iblock_istep_run tge sp (normRTL ib k) rs0 m0 = None. -Proof. - induction ib; simpl; - try discriminate. - - (* BF *) - intros; destruct o; try discriminate; simpl in *. - inv ISTEP. destruct CONT as [HO ISTEP]; inv HO. - - (* Bnop *) - intros; destruct o; inv ISTEP; destruct oiinfo, CONT; auto. - destruct (Bnop_dec k); subst; repeat econstructor; eauto. - - (* Bop *) - intros; destruct o; destruct (Bnop_dec k); - destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; - simpl; rewrite EVAL; auto; destruct CONT as [HO ISTEP]; - trivial. - - (* Bload *) - intros; destruct o; destruct (Bnop_dec k); - destruct (trap) eqn:TRAP; - try destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; - simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [HO ISTEP]; trivial. - - (* Bstore *) - intros; destruct o; destruct (Bnop_dec k); - destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; - simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [HO ISTEP]; inv HO; trivial. - - (* Bseq *) - intros. - eapply IHib1; eauto. - destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1; try auto. - destruct o0; simpl in *. destruct _fin; inv ISTEP. - + destruct CONT as [CONTRA _]; inv CONTRA. - + split; auto. eapply IHib2; eauto. - - (* Bcond *) - intros; destruct (eval_condition _ _ _); trivial. - destruct b. - + eapply IHib1; eauto. - + eapply IHib2; eauto. -Qed. - -Lemma normRTL_preserves_iblock_istep_run_None sp ib: - forall rs m, iblock_istep_run tge sp ib rs m = None - -> iblock_istep_run tge sp (normRTL ib (Bnop None)) rs m = None. -Proof. - intros; eapply normRTL_iblock_istep_run_None_rec; eauto. - simpl; auto. -Qed. - -Lemma normRTL_preserves_iblock_istep_run sp ib: - forall rs m, iblock_istep_run tge sp ib rs m = - iblock_istep_run tge sp (normRTL ib (Bnop None)) rs m. -Proof. - intros. - destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP. - - destruct o. symmetry. - rewrite <- iblock_istep_run_equiv in *. - apply normRTL_iblock_istep_correct; auto. - - symmetry. - apply normRTL_preserves_iblock_istep_run_None; auto. -Qed. - -Local Hint Constructors match_iblock: core. -Lemma normRTL_matchiblock_rec_correct dupmap cfg ib pc isfst: - forall opc1 - (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2 - (CONT: match opc1 with - | Some pc' => - match_iblock dupmap cfg false pc' k opc2 - | None => opc2=opc1 - end), - match_iblock dupmap cfg isfst pc (normRTL ib k) opc2. -Proof. - induction 1; simpl; intros; eauto. - { (* BF *) - inv CONT; econstructor; eauto. } - 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct (Bnop_dec k); subst; eauto; inv CONT; econstructor; eauto. - { (* Bgoto *) - intros; inv CONT; apply mib_exit; auto. } - { (* Bcond *) - intros. inv H0; - econstructor; eauto; try econstructor. - destruct opc0; econstructor. } -Qed. - -Lemma normRTL_matchiblock_correct dupmap cfg ib pc isfst opc: - match_iblock dupmap cfg isfst pc ib opc -> - match_iblock dupmap cfg isfst pc (normRTL ib (Bnop None)) opc. -Proof. - intros. - eapply normRTL_matchiblock_rec_correct; eauto. - destruct opc; simpl; auto. -Qed. (** * Match strong state property @@ -455,17 +570,20 @@ See explanations of opt_simu below. *) Lemma match_strong_state_simu - dupmap st st' f f' sp rs2 m2 rs1 m1 rs0 m0 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n - (STEP : RTL.step ge (RTL.State st f sp pcR1 rs1 m1) E0 (RTL.State st f sp pcR2 rs2 m2)) + dupmap st st' f f' sp rs2 m2 rs1 m1 rs0 m0 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n t s1' + (EQt: t=E0) + (EQs1': s1'=(RTL.State st f sp pcR2 rs2 m2)) + (STEP : RTL.step ge (RTL.State st f sp pcR1 rs1 m1) t s1') (MSS1 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst) (MSS2 : match_strong_state dupmap st st' f f' sp rs2 m2 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false) (MES : measure ib2 < n) : exists (oib' : option iblock), - (exists s2', step tge (State st' f' sp pcB0 rs0 m0) E0 s2' - /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) s2') - \/ (omeasure oib' < n /\ E0=E0 - /\ match_states oib' (RTL.State st f sp pcR2 rs2 m2) (State st' f' sp pcB0 rs0 m0)). + (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' + /\ match_states oib' s1' s2') + \/ (omeasure oib' < n /\ t=E0 + /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). Proof. + subst. destruct (is_goto ib2) eqn:GT. destruct ib2; try destruct fi; try discriminate. - (* Bgoto *) @@ -473,13 +591,9 @@ Proof. remember H2 as ODUPLIC; clear HeqODUPLIC. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. apply DMC in H2 as [ib [FNC MI]]; clear DMC. - eexists; left; eexists; split. - + repeat econstructor; eauto. - apply iblock_istep_run_equiv in BTL_RUN; eauto. - + econstructor; apply normRTL_matchiblock_correct in MI. - econstructor; eauto. eapply is_expand_normRTL_entry; eauto. - econstructor. apply normRTL_preserves_iblock_istep_run. - eapply normRTL_entry_isnt_goto; eauto. + eexists; left; eexists; split; eauto. + repeat econstructor; eauto. + apply iblock_istep_run_equiv in BTL_RUN; eauto. - (* Others *) exists (Some ib2); right; split. simpl; auto. @@ -495,7 +609,7 @@ Lemma opt_simu_intro (exists s2', step tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2') \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)). Proof. - inversion MSTRONG; subst. inv MIB. + inv MSTRONG; subst. inv MIB. - (* mib_BF *) inv H0; inversion STEP; subst; try_simplify_someHyps; intros. @@ -536,81 +650,60 @@ Proof. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. remember H1 as ODUPLIC; clear HeqODUPLIC. apply DMC in H1 as [ib [FNC MI]]; clear DMC. - eexists; left; eexists; split. - * econstructor; eauto. econstructor. - eexists; eexists; split. - eapply iblock_istep_run_equiv in BTL_RUN. - eapply BTL_RUN. econstructor; eauto. - pose symbols_preserved as SYMPRES. - eapply eval_builtin_args_preserved; eauto. - eapply external_call_symbols_preserved; eauto. eapply senv_preserved. - * econstructor; eauto; apply normRTL_matchiblock_correct in MI. - { econstructor; eauto. eapply is_expand_normRTL_entry; eauto. - apply star_refl. apply normRTL_preserves_iblock_istep_run. } - eapply normRTL_entry_isnt_goto; eauto. + exists (Some (normRTL (entry ib))); left; eexists; split; eauto. + econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. + pose symbols_preserved as SYMPRES. + eapply eval_builtin_args_preserved; eauto. + eapply external_call_symbols_preserved; eauto. eapply senv_preserved. + (* Bjumptable *) exploit list_nth_z_rev_dupmap; eauto. intros (pc'0 & LNZ & DM). eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. remember DM as ODUPLIC; clear HeqODUPLIC. apply DMC in DM as [ib [FNC MI]]; clear DMC. - eexists; left; eexists; split. - * econstructor; eauto. econstructor. - eexists; eexists; split. - eapply iblock_istep_run_equiv in BTL_RUN. - eapply BTL_RUN. econstructor; eauto. - * econstructor; eauto; apply normRTL_matchiblock_correct in MI. - { econstructor; eauto. eapply is_expand_normRTL_entry; eauto. - apply star_refl. apply normRTL_preserves_iblock_istep_run. } - eapply normRTL_entry_isnt_goto; eauto. + exists (Some (normRTL (entry ib))); left; eexists; split; eauto. + econstructor; eauto. econstructor. + eexists; eexists; split. + eapply iblock_istep_run_equiv in BTL_RUN. + eapply BTL_RUN. econstructor; eauto. - (* mib_exit *) discriminate. - (* mib_seq *) - inversion H; subst; - try (inv IS_EXPD; try inv H5; discriminate; fail); - inversion STEP; subst; try_simplify_someHyps; intros. + inv IS_EXPD; try discriminate. + inv H; simpl in *; try congruence; + inv STEP; try_simplify_someHyps; eauto. + (* Bnop is_rtl *) - eapply match_strong_state_simu. - 1,2: do 2 (econstructor; eauto). - econstructor; eauto. - inv IS_EXPD; eauto. simpl in *; discriminate. - eapply star_right; eauto. lia. + intros; eapply match_strong_state_simu; eauto. + (* Bop *) - eapply match_strong_state_simu. - 1,2: do 2 (econstructor; eauto). + intros; eapply match_strong_state_simu; eauto. econstructor; eauto. - inv IS_EXPD; eauto. simpl in *; discriminate. - eapply star_right; eauto. - erewrite eval_operation_preserved in H11. - erewrite H11 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. lia. + erewrite eval_operation_preserved in H12. + erewrite H12 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. + (* Bload *) - eapply match_strong_state_simu. - 1,2: do 2 (econstructor; eauto). + intros; eapply match_strong_state_simu; eauto. econstructor; eauto. - inv IS_EXPD; eauto. simpl in *; discriminate. - eapply star_right; eauto. - erewrite eval_addressing_preserved in H11. - erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. lia. + erewrite eval_addressing_preserved in H12. + erewrite H12, H13 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. + (* Bstore *) - eapply match_strong_state_simu. - 1,2: do 2 (econstructor; eauto). + intros; eapply match_strong_state_simu; eauto. econstructor; eauto. - inv IS_EXPD; eauto. simpl in *; discriminate. - eapply star_right; eauto. - erewrite eval_addressing_preserved in H11. - erewrite H11, H12 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. lia. + erewrite eval_addressing_preserved in H12. + erewrite H12, H13 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. - (* mib_cond *) + inv IS_EXPD; try discriminate. inversion STEP; subst; try_simplify_someHyps; intros. - intros; rewrite H12 in BTL_RUN. destruct b; - eapply match_strong_state_simu; eauto. - 1,3: inv H2; econstructor; eauto. - 1,3,5,7: inv IS_EXPD; auto; try discriminate. - 1-4: eapply star_right; eauto. - assert (measure bnot > 0) by apply measure_pos; lia. - assert (measure bso > 0) by apply measure_pos; lia. + destruct (is_join_opt_None opc1 opc2); eauto. subst. + eapply match_strong_state_simu with (ib1:=Bcond c lr bso bnot iinfo) (ib2:=(if b then bso else bnot)); eauto. + + intros; rewrite H14 in BTL_RUN; destruct b; econstructor; eauto. + + assert (measure (if b then bnot else bso) > 0) by apply measure_pos; destruct b; simpl; lia. + Unshelve. + all: eauto. Qed. (** * Main RTL to BTL simulation theorem @@ -624,9 +717,9 @@ Two possible executions: RTL state match_states BTL state s1 ------------------------------------ s2 | | - STEP | Classical lockstep simu | - | | - s1' ----------------------------------- s2' + STEP | Classical lockstep simu | + | | + s1' ----------------------------------- s2' **Middle instruction (right side):** @@ -634,8 +727,8 @@ Two possible executions: RTL state match_states [oib] BTL state s1 ------------------------------------ s2 | _______/ - STEP | *E0 ___________________/ - | / match_states [oib'] + STEP | *E0 ___________________/ + | / match_states [oib'] s1' ______/ Where omeasure oib' < omeasure oib @@ -662,15 +755,10 @@ Proof. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. apply DMC in ENTRY as DMC'. destruct DMC' as [ib [CENTRY MI]]; clear DMC. - eexists; left; eexists; split. + exists (Some (normRTL (entry ib))); left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. - * apply normRTL_matchiblock_correct in MI. - econstructor. econstructor; eauto. - eapply is_expand_normRTL_entry; eauto. - 3: eapply normRTL_entry_isnt_goto; eauto. - all: erewrite preserv_fnparams; eauto. - constructor. apply normRTL_preserves_iblock_istep_run. + * erewrite preserv_fnparams; eauto. + (* External function *) inv TRANSF. eexists; left; eexists; split. @@ -683,13 +771,8 @@ Proof. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. remember DUPLIC as ODUPLIC; clear HeqODUPLIC. apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. - eexists; left; eexists; split. - + eapply exec_return. - + apply normRTL_matchiblock_correct in MI. - econstructor. econstructor; eauto. - eapply is_expand_normRTL_entry; eauto. - econstructor. apply normRTL_preserves_iblock_istep_run. - eapply normRTL_entry_isnt_goto; eauto. + eexists; left; eexists; split; eauto. + eapply exec_return. Qed. Local Hint Resolve plus_one star_refl: core. -- cgit From 3d5627a374fa88a3c7eaec4c46c44c3327606341 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 25 May 2021 07:24:49 +0200 Subject: simplification of normRTL --- scheduling/RTLtoBTLproof.v | 94 +++++++++++++--------------------------------- 1 file changed, 27 insertions(+), 67 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7cd8e47d..a22be1d5 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -27,13 +27,6 @@ Definition is_RTLbasic (ib: iblock): bool := | _ => true end. - -Definition Bnop_dec k: { k=Bnop None} + { k<>(Bnop None) }. -Proof. - destruct k; simpl; try destruct oiinfo; - intuition congruence. -Defined. - (** The strict [is_normRTL] property specifying the ouput of [normRTL] below *) Inductive is_normRTL: iblock -> Prop := | norm_Bseq ib1 ib2: @@ -50,11 +43,11 @@ Inductive is_normRTL: iblock -> Prop := . Local Hint Constructors is_normRTL: core. -(** Weaker version with less restrictive hypothesis for conditions *) +(** Weaker version allowing for trailing [Bnop None]. *) Inductive is_wnormRTL: iblock -> Prop := | wnorm_Bseq ib1 ib2: is_RTLbasic ib1 = true -> - is_wnormRTL ib2 -> + (ib2 <> Bnop None -> is_wnormRTL ib2) -> is_wnormRTL (Bseq ib1 ib2) | wnorm_Bcond cond args ib1 ib2 iinfo: (ib1 <> Bnop None -> is_wnormRTL ib1) -> @@ -66,7 +59,7 @@ Inductive is_wnormRTL: iblock -> Prop := . Local Hint Constructors is_wnormRTL: core. -(* NB: [k] is a "continuation" (e.g. semantically [normRTLrec ib k] is like [BSeq ib k]) *) +(* NB: [k] is a "continuation" (e.g. semantically [normRTLrec ib k] is like [Bseq ib k]) *) Fixpoint normRTLrec (ib: iblock) (k: iblock): iblock := match ib with | Bseq ib1 ib2 => normRTLrec ib1 (normRTLrec ib2 k) @@ -74,8 +67,7 @@ Fixpoint normRTLrec (ib: iblock) (k: iblock): iblock := Bcond cond args (normRTLrec ib1 k) (normRTLrec ib2 k) iinfo | BF fin iinfo => BF fin iinfo | Bnop None => k - | ib => - if Bnop_dec k then ib else Bseq ib k + | ib => Bseq ib k end. Definition normRTL ib := normRTLrec ib (Bnop None). @@ -130,17 +122,13 @@ Lemma normRTLrec_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin1: end), iblock_istep tge sp rs0 m0 (normRTLrec ib k) rs2 m2 ofin2. Proof. - induction 1; simpl; intros; intuition subst; eauto. - { (* Bnop *) - autodestruct; intros OI; inv OI; destruct (Bnop_dec k); subst; - try (inv CONT; constructor; fail); repeat econstructor; eauto. } + induction 1; simpl; intuition subst; eauto. + { (* Bnop *) autodestruct; eauto. } 1-3: (* Bop, Bload, Bstore *) - intros; destruct (Bnop_dec k); subst; repeat econstructor; eauto; - inv CONT; econstructor; eauto. - - (* Bcond *) - destruct ofin; intros; - econstructor; eauto; - destruct b; eapply IHISTEP; eauto. + intros; repeat econstructor; eauto. + (* Bcond *) + destruct ofin; intuition subst; + destruct b; eapply IHISTEP; eauto. Qed. Lemma normRTL_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin: @@ -152,9 +140,8 @@ Proof. Qed. Lemma normRTLrec_iblock_istep_run_None tge sp ib: - forall rs0 m0 o k - (ISTEP: iblock_istep_run tge sp ib rs0 m0 = o) - (CONT: match o with + forall rs0 m0 k + (CONT: match iblock_istep_run tge sp ib rs0 m0 with | Some (out rs1 m1 ofin) => ofin = None /\ iblock_istep_run tge sp k rs1 m1 = None @@ -162,44 +149,23 @@ Lemma normRTLrec_iblock_istep_run_None tge sp ib: end), iblock_istep_run tge sp (normRTLrec ib k) rs0 m0 = None. Proof. - induction ib; simpl; - try discriminate. - - (* BF *) - intros; destruct o; try discriminate; simpl in *. - inv ISTEP. destruct CONT as [HO ISTEP]; inv HO. + induction ib; simpl; intros; subst; intuition (try discriminate). - (* Bnop *) - intros; destruct o; inv ISTEP; destruct oiinfo, CONT; auto. - destruct (Bnop_dec k); subst; repeat econstructor; eauto. + intros. autodestruct; auto. - (* Bop *) - intros; destruct o; destruct (Bnop_dec k); - destruct (eval_operation _ _ _ _ _) eqn:EVAL; inv ISTEP; - simpl; rewrite EVAL; auto; destruct CONT as [HO ISTEP]; - trivial. + intros; repeat autodestruct; simpl; intuition congruence. - (* Bload *) - intros; destruct o; destruct (Bnop_dec k); - destruct (trap) eqn:TRAP; - try destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.loadv _ _ _) eqn:MEM; inv ISTEP; - simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [HO ISTEP]; trivial. + intros; repeat autodestruct; simpl; intuition congruence. - (* Bstore *) - intros; destruct o; destruct (Bnop_dec k); - destruct (eval_addressing _ _ _ _) eqn:EVAL; - try destruct (Mem.storev _ _ _) eqn:MEM; inv ISTEP; - simpl; try rewrite EVAL; try rewrite MEM; simpl; auto; - destruct CONT as [HO ISTEP]; inv HO; trivial. + intros; repeat autodestruct; simpl; intuition congruence. - (* Bseq *) intros. eapply IHib1; eauto. - destruct (iblock_istep_run tge sp ib1 rs0 m0) eqn:EQib1; try auto. - destruct o0; simpl in *. destruct _fin; inv ISTEP. - + destruct CONT as [CONTRA _]; inv CONTRA. - + split; auto. eapply IHib2; eauto. + autodestruct; simpl in *; destruct o; simpl in *; intuition eauto. + + destruct _fin; intuition eauto. + + destruct _fin; intuition congruence || eauto. - (* Bcond *) - intros; destruct (eval_condition _ _ _); trivial. - destruct b. - + eapply IHib1; eauto. - + eapply IHib2; eauto. + intros; repeat autodestruct; simpl; intuition congruence || eauto. Qed. Lemma normRTL_preserves_iblock_istep_run_None tge sp ib: @@ -207,7 +173,7 @@ Lemma normRTL_preserves_iblock_istep_run_None tge sp ib: -> iblock_istep_run tge sp (normRTL ib) rs m = None. Proof. intros; eapply normRTLrec_iblock_istep_run_None; eauto. - simpl; auto. + rewrite H; simpl; auto. Qed. Lemma normRTL_preserves_iblock_istep_run tge sp ib: @@ -234,17 +200,11 @@ Lemma normRTLrec_matchiblock_correct dupmap cfg ib pc isfst: end), match_iblock dupmap cfg isfst pc (normRTLrec ib k) opc2. Proof. - induction 1; simpl; intros; eauto. - { (* BF *) - inv CONT; econstructor; eauto. } - 1-4: (* Bnop, Bop, Bload, Bstore *) - destruct (Bnop_dec k); subst; eauto; inv CONT; econstructor; eauto. - { (* Bgoto *) - intros; inv CONT; apply mib_exit; auto. } - { (* Bcond *) - intros. inv H0; - econstructor; eauto; try econstructor. - destruct opc0; econstructor. } + induction 1; simpl; intros; subst; eauto. + (* Bcond *) + intros. inv H0; + econstructor; eauto; try econstructor. + destruct opc0; econstructor. Qed. Lemma normRTL_matchiblock_correct dupmap cfg ib pc isfst opc: -- cgit From 6938945d80bf16a6de4986e2815113e938bff6c3 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 25 May 2021 13:36:57 +0200 Subject: starting to experiment SE of fsem --- scheduling/BTL.v | 91 ++++++++++++++++---- scheduling/BTL_SEtheory.v | 213 ++++++++++++++++++++++++++++++++++------------ 2 files changed, 232 insertions(+), 72 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index fb6d5cea..4ae7a6dd 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -159,18 +159,22 @@ Record outcome := out { _fin: option final }. + + + Section RELSEM. + (** [step] (and in particular [final_step]) is parametrized by function [tr_exit] to transfer registers on each iblock exit. - In particular, [tr_exit f pc or rs] computes from [rs] the [rs'] on which the execution from iblock [pc] in [f] will start. + In particular, [tr_exit f lpc or rs] computes from [rs] the [rs'] on which the execution from iblock [pc] in [lpc] in [f] may start. Here, [or] is an optional register that will be assigned after exiting the iblock, but before entering in [pc]: e.g. the register set by a function call before entering in return address. *) -Variable tr_exit: function -> exit -> option reg -> regset -> regset. +Variable tr_exit: function -> list exit -> option reg -> regset -> regset. Variable ge: genv. @@ -282,10 +286,13 @@ Proof. destruct o; try_simplify_someHyps; subst; eauto. Qed. +Import ListNotations. +Local Open Scope list_scope. + 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 (tr_exit f pc None rs) m) + (State stack f sp pc (tr_exit f [pc] None rs) m) | exec_Breturn or stk m': sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> @@ -295,7 +302,7 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := find_function ros rs = Some fd -> funsig fd = sig -> final_step stack f sp rs m (Bcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' (tr_exit f pc' (Some res) rs) :: stack) fd rs##args m) + E0 (Callstate (Stackframe res f sp pc' (tr_exit f [pc'] (Some res) rs) :: stack) fd rs##args m) | exec_Btailcall sig ros args stk m' fd: find_function ros rs = Some fd -> funsig fd = sig -> @@ -308,14 +315,14 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := external_call ef ge vargs m t vres m' -> final_step stack f sp rs m (Bbuiltin ef args res pc') (* TODO: not clear that this is the better choice ! - we coud also do something like [regmap_setres res vres (tr_exit f pc' (reg_builtin_res res) rs)] + we coud also do something like [regmap_setres res vres (tr_exit f [pc'] (reg_builtin_res res) rs)] *) - t (State stack f sp pc' (tr_exit f pc' None (regmap_setres res vres rs)) m') + t (State stack f sp pc' (tr_exit f [pc'] None (regmap_setres res vres rs)) m') | exec_Bjumptable arg tbl n pc': rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> final_step stack f sp rs m (Bjumptable arg tbl) - E0 (State stack f sp pc' (tr_exit f pc' None rs) m) + E0 (State stack f sp pc' (tr_exit f tbl None rs) m) . (** big-step execution of one iblock *) @@ -375,7 +382,7 @@ Inductive final_state: state -> int -> Prop := (** The basic small-step semantics for a BTL program. *) (* tid = transfer_identity *) -Definition tid (_:function) (_:exit) (_: option reg) (rs:regset) := rs. +Definition tid (_:function) (_:list exit) (_: option reg) (rs:regset) := rs. Definition semantics (p: program) := Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). @@ -405,19 +412,51 @@ Proof. - rewrite Regmap.gso; auto. Qed. -Definition inputs_exit (f:function) (pc: exit) (or:option reg): Regset.t := - match f.(fn_code)!pc with - | None => Regset.empty - | Some ib => - match or with - | Some r => Regset.remove r ib.(input_regs) - | None => ib.(input_regs) - end +Fixpoint inputs_exitrec (f:function) (tbl: list exit): Regset.t := + match tbl with + | nil => Regset.empty + | pc::l => let rs:= inputs_exitrec f l in + match f.(fn_code)!pc with + | None => rs + | Some ib => Regset.union rs ib.(input_regs) + end end. -Definition tr_inputs (f:function) (pc: exit) (or:option reg): regset -> regset +Lemma inputs_exitrec_In f tbl r: + Regset.In r (inputs_exitrec f tbl) -> (exists pc, List.In pc tbl /\ exists ib, f.(fn_code)!pc = Some ib /\ Regset.In r ib.(input_regs)). +Proof. + induction tbl as [|pc l]; simpl; intros. + - exploit Regset.empty_1; eauto. intuition. + - destruct ((fn_code f) ! pc) eqn:ATpc. + + exploit Regset.union_1; eauto. + destruct 1 as [X1|X1]. + * destruct IHl as (pc' & H1 & H2); eauto. + * eexists; intuition eauto. + + destruct IHl as (pc' & H1 & H2); eauto. +Qed. + +Lemma inputs_exitrec_notIn f tbl r: + (forall pc ib, List.In pc tbl -> f.(fn_code)!pc = Some ib -> ~Regset.In r ib.(input_regs)) + -> ~Regset.In r (inputs_exitrec f tbl). +Proof. + induction tbl as [|pc l]; simpl; intuition eauto. + - exploit Regset.empty_1; eauto. + - destruct ((fn_code f) ! pc) eqn:ATpc; intuition eauto. + exploit Regset.union_1; intuition eauto. +Qed. + +Definition inputs_exit (f:function) (tbl: list exit) (or: option reg): Regset.t := + let rs := inputs_exitrec f tbl in + match or with + | Some r => Regset.remove r rs + | None => rs + end. + +Definition tr_inputs (f:function) (pc: list exit) (or:option reg): regset -> regset := transfer_regs (Regset.elements (inputs_exit f pc or)). + + (* TODO: move this elsewhere *) Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, SetoidList.InA (fun x y => x = y) x l <-> List.In x l. @@ -443,6 +482,24 @@ Proof. rewrite -> SetoidList_InA_eq_equiv; eauto. Qed. +Definition Regset_In_dec r rs: + { Regset.In r rs } + { ~(Regset.In r rs)}. +Proof. + destruct (Regset.mem r rs) eqn: IsIN. + + left. abstract (eapply Regset.mem_2; auto). + + right. + abstract (intro H; exploit Regset.mem_1; eauto; congruence). +Defined. + +Lemma tr_inputs_ext f pc or rs1 rs2: + (forall r, Regset.In r (inputs_exit f pc or) -> rs1#r = rs2#r) -> + (forall r, (tr_inputs f pc or rs1)#r = (tr_inputs f pc or rs2)#r). +Proof. + intros EQ r. destruct (Regset_In_dec r (inputs_exit f pc or)). + + rewrite! tr_inputs_exit; eauto. + + rewrite! tr_inputs_dead; eauto. +Qed. + Definition fsem (p: program) := Semantics (step tr_inputs) (initial_state p) final_state (Genv.globalenv p). diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 5125ea3c..fb7c650f 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -28,6 +28,7 @@ Record iblock_exec_context := Bctx { (* symbolic value *) Inductive sval := + | Sundef | Sinput (r: reg) | Sop (op:operation) (lsv: list_sval) (sm: smem) | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) @@ -53,6 +54,7 @@ Local Open Scope option_monad_scope. Fixpoint eval_sval ctx (sv: sval): option val := match sv with + | Sundef => Some Vundef | Sinput r => Some ((crs0 ctx)#r) | Sop op l sm => SOME args <- eval_list_sval ctx l IN @@ -383,7 +385,6 @@ Qed. Inductive sfval := | Sgoto (pc: exit) | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:exit) - (* NB: [res] the return register is hard-wired ! Is it restrictive ? *) | Stailcall: signature -> sval + ident -> list_sval -> sfval | Sbuiltin (ef:external_function) (sargs: list (builtin_arg sval)) (res: builtin_res reg) (pc:exit) | Sjumptable (sv: sval) (tbl: list exit) @@ -397,39 +398,60 @@ Definition sfind_function ctx (svos : sval + ident): option fundef := end . -Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := - | exec_Sgoto pc rs m: - sem_sfval ctx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc rs m) +Import ListNotations. +Local Open Scope list_scope. + +(** [sem_sfval] corresponds to [final_step tr_inputs] for symbolic final value. + + A main difference comes from [rsx] which observes the registers saved in the stackframe + of the returned state. + + Intuitively, only liveout registers are stored in the stack + (others will be associated to [Vundef] value). +*) +Inductive sem_sfval ctx (rsx: reg -> option val): sfval -> regset -> mem -> trace -> state -> Prop := + | exec_Sgoto pc rs rs' m: + rs' = tr_inputs ctx.(cf) [pc] None rs -> + (forall r, rsx r = Some rs'#r) -> + sem_sfval ctx rsx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc rs' m) | exec_Sreturn stk osv rs m m' v: (csp ctx) = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> match osv with Some sv => eval_sval ctx sv | None => Some Vundef end = Some v -> - sem_sfval ctx (Sreturn osv) rs m + (forall r, rsx r = Some Vundef) -> + sem_sfval ctx rsx (Sreturn osv) rs m E0 (Returnstate (cstk ctx) v m') - | exec_Scall rs m sig svos lsv args res pc fd: + | exec_Scall rs m sig svos lsv args res pc fd rs': sfind_function ctx svos = Some fd -> funsig fd = sig -> eval_list_sval ctx lsv = Some args -> - sem_sfval ctx (Scall sig svos lsv res pc) rs m - E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc rs :: (cstk ctx)) fd args m) + rs' = tr_inputs ctx.(cf) [pc] (Some res) rs -> + (forall r, rsx r = Some rs'#r) -> + sem_sfval ctx rsx (Scall sig svos lsv res pc) rs m + E0 (Callstate ((Stackframe res (cf ctx) (csp ctx) pc rs')::cstk ctx) fd args m) | exec_Stailcall stk rs m sig svos args fd m' lsv: sfind_function ctx svos = Some fd -> funsig fd = sig -> (csp ctx) = Vptr stk Ptrofs.zero -> Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> eval_list_sval ctx lsv = Some args -> - sem_sfval ctx (Stailcall sig svos lsv) rs m + (forall r, rsx r = Some Vundef) -> + sem_sfval ctx rsx (Stailcall sig svos lsv) rs m E0 (Callstate (cstk ctx) fd args m') - | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: + | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs rs': seval_builtin_args ctx m sargs vargs -> external_call ef (cge ctx) vargs m t vres m' -> - sem_sfval ctx (Sbuiltin ef sargs res pc) rs m - t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres rs) m') - | exec_Sjumptable sv tbl pc' n rs m: + rs' = tr_inputs ctx.(cf) [pc] None (regmap_setres res vres rs) -> + (forall r, rsx r = Some rs'#r) -> + sem_sfval ctx rsx (Sbuiltin ef sargs res pc) rs m + t (State (cstk ctx) (cf ctx) (csp ctx) pc rs' m') + | exec_Sjumptable sv tbl pc' n rs m rs': eval_sval ctx sv = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - sem_sfval ctx (Sjumptable sv tbl) rs m - E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' rs m) + rs' = tr_inputs ctx.(cf) tbl None rs -> + (forall r, rsx r = Some rs'#r) -> + sem_sfval ctx rsx (Sjumptable sv tbl) rs m + E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' rs' m) . (** * Preservation properties *) @@ -536,18 +558,18 @@ End SymbValPreserved. Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg: reg -> sval; si_smem: smem }. (* Predicate on which (rs, m) is a possible final state after evaluating [st] on ((crs0 ctx), (cm0 ctx)) *) -Definition sem_sistate ctx (st: sistate) (rs: regset) (m: mem): Prop := - st.(si_pre) ctx - /\ eval_smem ctx st.(si_smem) = Some m - /\ forall (r:reg), eval_sval ctx (st.(si_sreg) r) = Some (rs#r). +Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop := + sis.(si_pre) ctx + /\ eval_smem ctx sis.(si_smem) = Some m + /\ forall (r:reg), eval_sval ctx (sis.(si_sreg) r) = Some (rs#r). (* Remark that we need to reason modulo "extensionality" wrt Regmap.get about regsets. And, nothing in their representation as (val * PTree.t val) enforces that (forall r, rs1#r = rs2#r) -> rs1 = rs2 *) -Lemma sem_sistate_determ ctx st rs1 m1 rs2 m2: - sem_sistate ctx st rs1 m1 -> - sem_sistate ctx st rs2 m2 -> +Lemma sem_sistate_determ ctx sis rs1 m1 rs2 m2: + sem_sistate ctx sis rs1 m1 -> + sem_sistate ctx sis rs2 m2 -> (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. intros (_&MEM1®1) (_&MEM2®2). @@ -556,6 +578,82 @@ Proof. Qed. (** * Symbolic execution of final step *) +Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := + match inputs with + | nil => fun r => Sundef + | r1::l => fun r => if Pos.eq_dec r r1 then sreg r1 else transfer_sreg l sreg r + end. + +Lemma transfer_sreg_inputs inputs sreg r: + List.In r inputs -> transfer_sreg inputs sreg r = sreg r. +Proof. + induction inputs; simpl; try autodestruct; intuition congruence. +Qed. + +Lemma transfer_sreg_dead inputs sreg r: + ~List.In r inputs -> transfer_sreg inputs sreg r = Sundef. +Proof. + induction inputs; simpl; try autodestruct; intuition congruence. +Qed. + +Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (inputs_exit f tbl or)). + +Definition transfer_sis (f:function) (tbl: list exit) (or:option reg) (sis: sistate) := + {| si_pre := sis.(si_pre); si_sreg := str_inputs f tbl or sis.(si_sreg); si_smem := sis.(si_smem) |}. + +Definition sreg_eval ctx (sis: sistate) (r: reg): option val := + eval_sval ctx (sis.(si_sreg) r). + +Lemma transfer_sis_correct ctx sis rs tbl or r: + (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) -> + sreg_eval ctx (transfer_sis (cf ctx) tbl or sis) r = Some (tr_inputs (cf ctx) tbl or rs) # r. +Admitted. + +Local Hint Resolve transfer_sis_correct: core. + +(* TODO: move the 3 below functions in [BTL] ? + Could be reused for liveness checking ? +*) +Definition reg_builtin_res (res: builtin_res reg): option reg := + match res with + | BR r => Some r + | _ => None + end. + +Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: final): A := + match i with + | Bgoto pc => tr f [pc] None + | Bcall sig ros args res pc => tr f [pc] (Some res) + | Btailcall sig ros args => tr f [] None + | Bbuiltin ef args res pc => tr f [pc] (reg_builtin_res res) + | Breturn or => tr f [] None + | Bjumptable reg tbl => tr f tbl None + end. + +Definition tr_regs: function -> final -> regset -> regset := + poly_tr tr_inputs. + +Definition tr_sis: function -> final -> sistate -> sistate := + poly_tr transfer_sis. + +Lemma tr_sis_regs_correct_aux ctx fin sis rs m: + sem_sistate ctx sis rs m -> + (forall r, sreg_eval ctx (tr_sis (cf ctx) fin sis) r = Some (tr_regs (cf ctx) fin rs) # r). +Proof. + destruct 1 as (_ & _ & REG). + destruct fin; simpl; eauto. +Qed. + +Lemma tr_sis_regs_correct ctx fin sis rs m: + sem_sistate ctx sis rs m -> + sem_sistate ctx (tr_sis (cf ctx) fin sis) (tr_regs (cf ctx) fin rs) m. +Proof. + intros H. + generalize (tr_sis_regs_correct_aux _ fin _ _ _ H). + destruct H as (PRE & MEM & REG). + destruct fin; simpl; econstructor; simpl; intuition eauto. +Qed. + Definition sexec_final_sfv (i: final) (sis: sistate): sfval := match i with | Bgoto pc => Sgoto pc @@ -582,10 +680,10 @@ Local Hint Constructors sem_sfval: core. Lemma sexec_final_svf_correct ctx i sis t rs m s: sem_sistate ctx sis rs m -> - final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> - sem_sfval ctx (sexec_final_sfv i sis) rs m t s. + final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> + sem_sfval ctx (sreg_eval ctx (tr_sis (cf ctx) i sis)) (sexec_final_sfv i sis) rs m t s. Proof. - intros (PRE&MEM®). + intros (PRE&MEM®). destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. + (* Bcall *) intros; eapply exec_Scall; auto. - destruct ros; simpl in * |- *; auto. @@ -597,21 +695,20 @@ Proof. - erewrite eval_list_sval_inj; simpl; auto. + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto. eapply seval_builtin_args_correct; eauto. + admit. + (* Bjumptable *) intros. eapply exec_Sjumptable; eauto. congruence. -Qed. +Admitted. Local Hint Constructors final_step: core. Local Hint Resolve seval_builtin_args_exact: core. Lemma sexec_final_svf_exact ctx i sis t rs m s: sem_sistate ctx sis rs m -> - sem_sfval ctx (sexec_final_sfv i sis) rs m t s - -> final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. + sem_sfval ctx (sreg_eval ctx (tr_sis (cf ctx) i sis)) (sexec_final_sfv i sis) rs m t s + -> final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. Proof. intros (PRE&MEM®). destruct i; simpl; intros LAST; inv LAST; eauto. - + (* Bgoto *) - econstructor. + (* Breturn *) enough (v=regmap_optget res Vundef rs) as ->; eauto. destruct res; simpl in *; congruence. @@ -625,8 +722,6 @@ Proof. intros; eapply exec_Btailcall; eauto. destruct fn; simpl in * |- *; auto. rewrite REG in * |- ; auto. - + (* Bbuiltin *) - eapply (exec_Bbuiltin tid); eauto. + (* Bjumptable *) eapply exec_Bjumptable; eauto. congruence. @@ -752,9 +847,10 @@ Inductive sstate := (* transition (t,s) produced by a sstate in initial context ctx *) Inductive sem_sstate ctx t s: sstate -> Prop := - | sem_Sfinal sis sfv rs m - (SIS: sem_sistate ctx sis rs m) - (SFV: sem_sfval ctx sfv rs m t s) + | sem_Sfinal sis sfv rs rs' rsx m + (SIS: sem_sistate ctx sis rs' m) + (EXT: forall r, rsx r = Some (rs'#r)) + (SFV: sem_sfval ctx rsx sfv rs m t s) : sem_sstate ctx t s (Sfinal sis sfv) | sem_Scond b cond args sm ifso ifnot (SEVAL: seval_condition ctx cond args sm = Some b) @@ -774,9 +870,9 @@ pour représenter l'ensemble des chemins. *) -Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := +Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := match ib with - | BF fin => Sfinal sis (sexec_final_sfv fin sis) + | BF fin => Sfinal (tr_sis f fin sis) (sexec_final_sfv fin sis) (* basic instructions *) | Bnop => k sis | Bop op args res => k (sexec_op op args res sis) @@ -785,19 +881,19 @@ Fixpoint sexec_rec ib sis (k: sistate -> sstate): sstate := | Bstore chunk addr args src => k (sexec_store chunk addr args src sis) (* composed instructions *) | Bseq ib1 ib2 => - sexec_rec ib1 sis (fun sis2 => sexec_rec ib2 sis2 k) + sexec_rec f ib1 sis (fun sis2 => sexec_rec f ib2 sis2 k) | Bcond cond args ifso ifnot _ => let args := list_sval_inj (List.map sis.(si_sreg) args) in - let ifso := sexec_rec ifso sis k in - let ifnot := sexec_rec ifnot sis k in + let ifso := sexec_rec f ifso sis k in + let ifnot := sexec_rec f ifnot sis k in Scond cond args sis.(si_smem) ifso ifnot end . -Definition sexec ib := sexec_rec ib sis_init (fun _ => Sabort). +Definition sexec f ib := sexec_rec f ib sis_init (fun _ => Sabort). Local Hint Constructors sem_sstate: core. -Local Hint Resolve sexec_op_correct sexec_final_svf_correct +Local Hint Resolve sexec_op_correct sexec_final_svf_correct tr_sis_regs_correct_aux tr_sis_regs_correct sexec_load_TRAP_correct sexec_store_correct sis_init_correct: core. Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin @@ -805,16 +901,16 @@ Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin (SIS: sem_sistate ctx sis rs m) (CONT: match ofin with | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx t s (k sis') - | Some fin => final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s + | Some fin => final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s end), - sem_sstate ctx t s (sexec_rec ib sis k). + sem_sstate ctx t s (sexec_rec (cf ctx) ib sis k). Proof. induction ISTEP; simpl; try autodestruct; eauto. (* condition *) all: intros; eapply sem_Scond; eauto; [ erewrite seval_condition_eq; eauto | - replace (if b then sexec_rec ifso sis k else sexec_rec ifnot sis k) with (sexec_rec (if b then ifso else ifnot) sis k); + replace (if b then sexec_rec (cf ctx) ifso sis k else sexec_rec (cf ctx) ifnot sis k) with (sexec_rec (cf ctx) (if b then ifso else ifnot) sis k); try autodestruct; eauto ]. Qed. @@ -822,8 +918,8 @@ Qed. (sexec is a correct over-approximation) *) Theorem sexec_correct ctx ib t s: - iblock_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> - sem_sstate ctx t s (sexec ib). + iblock_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> + sem_sstate ctx t s (sexec (cf ctx) ib). Proof. destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). eapply sexec_rec_correct; simpl; eauto. @@ -901,6 +997,9 @@ Proof. - repeat (rewrite Regmap.gso); auto. Qed. +Local Hint Resolve tr_inputs_ext: core. + +(* TODO Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: sem_sfval ctx sfv rs1 m t s -> (forall r, rs1#r = rs2#r) -> @@ -911,6 +1010,7 @@ Proof. destruct 1; eexists; split; econstructor; eauto. econstructor; eauto. Qed. +*) Local Hint Resolve sexec_final_svf_exact: core. @@ -973,13 +1073,14 @@ Local Hint Resolve sexec_op_preserves_abort sexec_load_preserves_abort sexec_store_preserves_abort: core. Lemma sexec_exclude_abort ctx ib t s1: forall sis k - (SEXEC: sem_sstate ctx t s1 (sexec_rec ib sis k)) + (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) (ABORT: abort_sistate ctx sis), False. Proof. induction ib; simpl; intros; eauto. - (* final *) inversion SEXEC; subst; eauto. + admit. - (* load *) destruct trap; eauto. inversion SEXEC. - (* seq *) @@ -988,7 +1089,7 @@ Proof. - (* cond *) inversion SEXEC; subst; eauto. clear SEXEC. destruct b; eauto. -Qed. +Admitted. Lemma set_sreg_abort ctx dst sv sis rs m: sem_sistate ctx sis rs m -> @@ -1048,14 +1149,14 @@ Qed. Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort: core. Lemma sexec_rec_exact ctx ib t s1: forall sis k - (SEXEC: sem_sstate ctx t s1 (sexec_rec ib sis k)) + (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) rs m (SIS: sem_sistate ctx sis rs m) (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) , match iblock_istep_run (cge ctx) (csp ctx) ib rs m with | Some (out rs' m' (Some fin)) => - exists s2, final_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 + exists s2, final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 | Some (out rs' m' None) => exists sis', (sem_sstate ctx t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m') | None => False end. @@ -1063,11 +1164,13 @@ Proof. induction ib; simpl; intros; eauto. - (* final *) inv SEXEC. + admit. + (* TODO exploit (sem_sistate_determ ctx sis rs m rs0 m0); eauto. intros (REG&MEM); subst. exploit (sem_sfval_equiv rs0 rs); eauto. intros (s2 & EQUIV & SFV'). - eexists; split; eauto. + eexists; split; eauto. *) - (* Bop *) autodestruct; eauto. - destruct trap; [| inv SEXEC ]. repeat autodestruct; eauto. @@ -1093,15 +1196,15 @@ Proof. destruct b. + exploit IHib1; eauto. + exploit IHib2; eauto. -Qed. +Admitted. (* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution (sexec is exact). *) Theorem sexec_exact ctx ib t s1: - sem_sstate ctx t s1 (sexec ib) -> - exists s2, iblock_step tid (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 + sem_sstate ctx t s1 (sexec (cf ctx) ib) -> + exists s2, iblock_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 /\ equiv_state s1 s2. Proof. intros; exploit sexec_rec_exact; eauto. -- cgit From 85aa6d418e077a9f492f800b3f61fb5ead705e4d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 25 May 2021 18:25:22 +0200 Subject: fix Builtin semantics --- scheduling/BTL.v | 34 ++++++++++++++++++++++++---------- scheduling/BTL_SEtheory.v | 30 +++--------------------------- 2 files changed, 27 insertions(+), 37 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 4ae7a6dd..954b4cb4 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -18,6 +18,8 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad. +Import ListNotations. + (** * Abstract syntax *) Definition exit := node. (* we may generalize this with register renamings at exit, @@ -159,12 +161,15 @@ Record outcome := out { _fin: option final }. - - +(* follows RTL semantics to set the result of builtin *) +Definition reg_builtin_res (res: builtin_res reg): option reg := + match res with + | BR r => Some r + | _ => None + end. Section RELSEM. - (** [step] (and in particular [final_step]) is parametrized by function [tr_exit] to transfer registers on each iblock exit. In particular, [tr_exit f lpc or rs] computes from [rs] the [rs'] on which the execution from iblock [pc] in [lpc] in [f] may start. @@ -286,7 +291,6 @@ Proof. destruct o; try_simplify_someHyps; subst; eauto. Qed. -Import ListNotations. Local Open Scope list_scope. Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := @@ -314,10 +318,7 @@ Inductive final_step stack f sp rs m: final -> trace -> state -> Prop := 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') - (* TODO: not clear that this is the better choice ! - we coud also do something like [regmap_setres res vres (tr_exit f [pc'] (reg_builtin_res res) rs)] - *) - t (State stack f sp pc' (tr_exit f [pc'] None (regmap_setres res vres rs)) m') + t (State stack f sp pc' (regmap_setres res vres (tr_exit f [pc'] (reg_builtin_res res) rs)) m') | exec_Bjumptable arg tbl n pc': rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> @@ -455,8 +456,6 @@ Definition inputs_exit (f:function) (tbl: list exit) (or: option reg): Regset.t Definition tr_inputs (f:function) (pc: list exit) (or:option reg): regset -> regset := transfer_regs (Regset.elements (inputs_exit f pc or)). - - (* TODO: move this elsewhere *) Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, SetoidList.InA (fun x y => x = y) x l <-> List.In x l. @@ -503,6 +502,21 @@ Qed. Definition fsem (p: program) := Semantics (step tr_inputs) (initial_state p) final_state (Genv.globalenv p). +Local Open Scope list_scope. + +Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: final): A := + match i with + | Bgoto pc => tr f [pc] None + | Bcall sig ros args res pc => tr f [pc] (Some res) + | Btailcall sig ros args => tr f [] None + | Bbuiltin ef args res pc => tr f [pc] (reg_builtin_res res) + | Breturn or => tr f [] None + | Bjumptable reg tbl => tr f tbl None + end. + +Definition tr_regs: function -> final -> regset -> regset := + poly_tr tr_inputs. + (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) (** Matching BTL and RTL code diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index fb7c650f..d682f776 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -441,10 +441,10 @@ Inductive sem_sfval ctx (rsx: reg -> option val): sfval -> regset -> mem -> trac | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs rs': seval_builtin_args ctx m sargs vargs -> external_call ef (cge ctx) vargs m t vres m' -> - rs' = tr_inputs ctx.(cf) [pc] None (regmap_setres res vres rs) -> + rs' = tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs -> (forall r, rsx r = Some rs'#r) -> sem_sfval ctx rsx (Sbuiltin ef sargs res pc) rs m - t (State (cstk ctx) (cf ctx) (csp ctx) pc rs' m') + t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres rs') m') | exec_Sjumptable sv tbl pc' n rs m rs': eval_sval ctx sv = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> @@ -611,28 +611,6 @@ Admitted. Local Hint Resolve transfer_sis_correct: core. -(* TODO: move the 3 below functions in [BTL] ? - Could be reused for liveness checking ? -*) -Definition reg_builtin_res (res: builtin_res reg): option reg := - match res with - | BR r => Some r - | _ => None - end. - -Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: final): A := - match i with - | Bgoto pc => tr f [pc] None - | Bcall sig ros args res pc => tr f [pc] (Some res) - | Btailcall sig ros args => tr f [] None - | Bbuiltin ef args res pc => tr f [pc] (reg_builtin_res res) - | Breturn or => tr f [] None - | Bjumptable reg tbl => tr f tbl None - end. - -Definition tr_regs: function -> final -> regset -> regset := - poly_tr tr_inputs. - Definition tr_sis: function -> final -> sistate -> sistate := poly_tr transfer_sis. @@ -695,9 +673,8 @@ Proof. - erewrite eval_list_sval_inj; simpl; auto. + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto. eapply seval_builtin_args_correct; eauto. - admit. + (* Bjumptable *) intros. eapply exec_Sjumptable; eauto. congruence. -Admitted. +Qed. Local Hint Constructors final_step: core. Local Hint Resolve seval_builtin_args_exact: core. @@ -727,7 +704,6 @@ Proof. congruence. Qed. - (** * symbolic execution of basic instructions *) Definition sis_init : sistate := {| si_pre:= fun _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. -- cgit From 42e887969f126635cb438fcf8b6f8969555b7eb7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 26 May 2021 13:01:54 +0200 Subject: avancement roadmap --- scheduling/BTLroadmap.md | 39 ++++++++++++++++++++++++++++----------- 1 file changed, 28 insertions(+), 11 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 32d0b365..6042803d 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -3,8 +3,8 @@ BTL aims to be an IR dedicated to defensive certification of middle-end optimizations (before register allocation). It provides a CFG of "loop-free" blocks, where each block is run in one step emitting at most a single observational event. The "local" optimizations (i.e. preserving "locally" the semantics of such blocks) would be checked by symbolic execution with rewriting rules. -The only info required from oracles: a "dupmap" mapping block entries. -In contrast, the "global" optimizations would require some invariants annotations at block entry (provided by oracles). +The main info required from oracles: a "dupmap" mapping block entries (and maybe strategies for controlling path explosion during symbolic execution -- see below). +Moreover, the "global" optimizations would require some invariants annotations at block entry (provided by oracles). Examples of optimizations that we aim to support: @@ -16,6 +16,9 @@ Examples of optimizations that we aim to support: We sketch below the various kinds of supported optimizations in development order... +Remark that we may port most of the current optimizations from RTL to BTL (even maybe, register allocation). +However, RTL will probably remain useful for "fine-tuned" duplication that crosses block boundaries (e.g. Duplicate module). + ## Introduction En gros la syntaxe d'un bloc BTL est définie par: @@ -83,7 +86,7 @@ coder la "simulation modulo liveness" dans une "simulation less_undef". Ça peut rendre le cadre du test de simulation plus simple et plus général. -**Idée** à côté de la sémantique "à la RTL" pour BTL, on pourrait +**Idée_Informelle** à côté de la sémantique "à la RTL" pour BTL, on pourrait définir une nouvelle sémantique (basée sur la même sémantique à grand pas dans les blocs), où c'est juste "l'entrée dans un bloc" qui change de sémantique. Intuitivement, cette nouvelle sémantique considèrerait @@ -102,11 +105,21 @@ totales) dans BTL. Pour l'encodage de SSA, il suffira d'étendre le mécanisme d'initialisation du "rs0" d'un bloc avec un passage de paramètres. -En fait, pour le test d'exécution symbolique, il semble plus simple de mettre les Vundef à la fin de la transition (précédente) plutôt qu'au début (de la suivante). +**REM** pour le test d'exécution symbolique, il semble plus adapté de +mettre les Vundef juste à la fin de la transition (précédente) plutôt +qu'au début (de la suivante). On pourrait aussi faire les deux (dans le détail, ça ne ferait pas exactement deux fois la même chose: +par exemple, quand on sort sur un call ou un builtin, le résultat du call/builtin n'est jamais dans le liveout puisqu'il va être écrasé de toute façon pendant la transition, +mais il peut être -- ou pas -- dans le livein de la transition suivante dans ce même bloc). + +Avoir une initialisation à Vundef en début de bloc permettrait de +faire une analyse des expressions mal initialisées - par exemple à +partir du bloc d'entrée de la fonction. Ça semble complémentaire de +l'analyse de "liveout". Mais utilisable uniquement dans le cadre d'une +combinaison "exécution symbolique"+"value analysis" (donc pas tout de suite). **IMPLEM PLAN** -1. définir the "functional semantics" of BTL. +1. définir the "functional semantics" of BTL (dans un premier temps, avec uniquement Vundef en sortie de bloc). Rem: en pratique, la "functional semantics" est définie pour être une "concrétisation" de la sémantique symbolique. Les définitions des deux sémantiques (symbolique et "functional") se font donc en simultanée. 2. prouver la simulation BTL.FunctionalSemantics -> BTL.CFGSemantics (facile normalement: c'est du "less_undef lockstep"). 3. implémenter une passe BTL.CFGSemantics -> BTL.FunctionalSemantics. en gros: on implémente un vérificateur des infos de liveness. @@ -119,11 +132,13 @@ En fait, pour le test d'exécution symbolique, il semble plus simple de mettre l ## "Basic" symbolic execution / symbolic simulation -We will implement a "basic" (e.g without rewriting rules) "less_undef" simulation test for BTL.FunctionalSemantics. +We will implement a "basic" (e.g without rewriting rules) simulation test for BTL.FunctionalSemantics. +Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter de tester l'égalité des registres de chaque côté: les registres hors liveout seront égaux à Vundef de chaque côté. **IMPLEM NOTE** - use Continuation Passing Style for an easy recursive generation of "all execution paths". +- dans l'etat concret, prévoir un booleen qui indique la valeur symbolique par defaut des registres: deux cas, soit "Input" (valeur initiale du registre), soit "Sundef" (valeur Vundef). **CURRENT STATUS** @@ -334,9 +349,13 @@ Au final, on obtient deux BDD identiques sur cet exemple (sans explosion combina ## Invariants at block entry -Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. +Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. Case-study: support of strength-reduction. + +**IDEAS** -Case-study: support of strength-reduction. +- En pratique, il faudra peut-être affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. + +- Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). ## Support of SSA-optimizations @@ -346,6 +365,4 @@ This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. ## Alias analysis in the symbolic simulation -A REGARDER [papier pointé par Justus](https://vbpf.github.io/assets/prevail-paper.pdf) - - +A REGARDER [papier pointé par Justus](https://vbpf.github.io/assets/prevail-paper.pdf) -- cgit From 47599a2ea88799d654ec644fe5ba9892087adb39 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 26 May 2021 16:27:21 +0200 Subject: fix tr_sis definition --- scheduling/BTL.v | 66 +++++++++++++-------- scheduling/BTL_SEtheory.v | 147 ++++++++++++++++++++++++++++------------------ 2 files changed, 132 insertions(+), 81 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 954b4cb4..5f75231a 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -453,8 +453,8 @@ Definition inputs_exit (f:function) (tbl: list exit) (or: option reg): Regset.t | None => rs end. -Definition tr_inputs (f:function) (pc: list exit) (or:option reg): regset -> regset - := transfer_regs (Regset.elements (inputs_exit f pc or)). +Definition tr_inputs (f:function) (tbl: list exit) (or:option reg): regset -> regset + := transfer_regs (Regset.elements (inputs_exit f tbl or)). (* TODO: move this elsewhere *) Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, @@ -465,38 +465,39 @@ Proof. - induction l; simpl; intuition. Qed. -Lemma tr_inputs_exit f pc or rs r: - Regset.In r (inputs_exit f pc or) -> (tr_inputs f pc or rs)#r = rs#r. +Lemma tr_inputs_exit f tbl or rs r: + Regset.In r (inputs_exit f tbl or) -> (tr_inputs f tbl or rs)#r = rs#r. Proof. intros; eapply transfer_regs_inputs. rewrite <- SetoidList_InA_eq_equiv. eapply Regset.elements_1; eauto. Qed. -Lemma tr_inputs_dead f pc or rs r: - ~(Regset.In r (inputs_exit f pc or)) -> (tr_inputs f pc or rs)#r = Vundef. +Lemma tr_inputs_dead f tbl or rs r: + ~(Regset.In r (inputs_exit f tbl or)) -> (tr_inputs f tbl or rs)#r = Vundef. Proof. intros X; eapply transfer_regs_dead; intuition eauto. eapply X. eapply Regset.elements_2. rewrite -> SetoidList_InA_eq_equiv; eauto. Qed. -Definition Regset_In_dec r rs: - { Regset.In r rs } + { ~(Regset.In r rs)}. +Local Hint Resolve tr_inputs_exit Regset.mem_2 Regset.mem_1: core. + +Lemma tr_inputs_get f tbl or rs r: + (tr_inputs f tbl or rs)#r = if Regset.mem r (inputs_exit f tbl or) then rs#r else Vundef. Proof. - destruct (Regset.mem r rs) eqn: IsIN. - + left. abstract (eapply Regset.mem_2; auto). - + right. - abstract (intro H; exploit Regset.mem_1; eauto; congruence). -Defined. + autodestruct; eauto. + intros; apply tr_inputs_dead; eauto. + intro X; exploit Regset.mem_1; eauto. + congruence. +Qed. -Lemma tr_inputs_ext f pc or rs1 rs2: - (forall r, Regset.In r (inputs_exit f pc or) -> rs1#r = rs2#r) -> - (forall r, (tr_inputs f pc or rs1)#r = (tr_inputs f pc or rs2)#r). +Lemma tr_inputs_ext f tbl or rs1 rs2: + (forall r, Regset.In r (inputs_exit f tbl or) -> rs1#r = rs2#r) -> + (forall r, (tr_inputs f tbl or rs1)#r = (tr_inputs f tbl or rs2)#r). Proof. - intros EQ r. destruct (Regset_In_dec r (inputs_exit f pc or)). - + rewrite! tr_inputs_exit; eauto. - + rewrite! tr_inputs_dead; eauto. + intros EQ r. rewrite !tr_inputs_get. + autodestruct; auto. Qed. Definition fsem (p: program) := @@ -507,16 +508,33 @@ Local Open Scope list_scope. Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: final): A := match i with | Bgoto pc => tr f [pc] None - | Bcall sig ros args res pc => tr f [pc] (Some res) - | Btailcall sig ros args => tr f [] None - | Bbuiltin ef args res pc => tr f [pc] (reg_builtin_res res) - | Breturn or => tr f [] None - | Bjumptable reg tbl => tr f tbl None + | Bcall _ _ _ res pc => tr f [pc] (Some res) + | Btailcall _ _ args => tr f [] None + | Bbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res) + | Breturn _ => tr f [] None + | Bjumptable _ tbl => tr f tbl None end. Definition tr_regs: function -> final -> regset -> regset := poly_tr tr_inputs. +Definition liveout: function -> final -> Regset.t := + poly_tr inputs_exit. + +Lemma tr_regs_liveout_equiv f fi : tr_regs f fi = transfer_regs (Regset.elements (liveout f fi)). +Proof. + destruct fi; simpl; auto. +Qed. + +Local Hint Resolve tr_inputs_get: core. + +Lemma tr_regs_get f fi rs r: (tr_regs f fi rs)#r = if Regset.mem r (liveout f fi) then rs#r else Vundef. +Proof. + Local Opaque inputs_exit. + destruct fi; simpl; auto. +Qed. + + (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) (** Matching BTL and RTL code diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index d682f776..a066626e 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -563,20 +563,6 @@ Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop := /\ eval_smem ctx sis.(si_smem) = Some m /\ forall (r:reg), eval_sval ctx (sis.(si_sreg) r) = Some (rs#r). -(* Remark that we need to reason modulo "extensionality" wrt Regmap.get about regsets. - And, nothing in their representation as (val * PTree.t val) enforces that - (forall r, rs1#r = rs2#r) -> rs1 = rs2 -*) -Lemma sem_sistate_determ ctx sis rs1 m1 rs2 m2: - sem_sistate ctx sis rs1 m1 -> - sem_sistate ctx sis rs2 m2 -> - (forall r, rs1#r = rs2#r) /\ m1 = m2. -Proof. - intros (_&MEM1®1) (_&MEM2®2). - intuition try congruence. - generalize (REG1 r); rewrite REG2; congruence. -Qed. - (** * Symbolic execution of final step *) Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := match inputs with @@ -598,29 +584,28 @@ Qed. Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (inputs_exit f tbl or)). -Definition transfer_sis (f:function) (tbl: list exit) (or:option reg) (sis: sistate) := - {| si_pre := sis.(si_pre); si_sreg := str_inputs f tbl or sis.(si_sreg); si_smem := sis.(si_smem) |}. - -Definition sreg_eval ctx (sis: sistate) (r: reg): option val := - eval_sval ctx (sis.(si_sreg) r). - -Lemma transfer_sis_correct ctx sis rs tbl or r: +Lemma str_inputs_correct ctx sis rs tbl or r: (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) -> - sreg_eval ctx (transfer_sis (cf ctx) tbl or sis) r = Some (tr_inputs (cf ctx) tbl or rs) # r. + eval_sval ctx (str_inputs (cf ctx) tbl or (si_sreg sis) r) = Some (tr_inputs (cf ctx) tbl or rs) # r. Admitted. -Local Hint Resolve transfer_sis_correct: core. +Local Hint Resolve str_inputs_correct: core. -Definition tr_sis: function -> final -> sistate -> sistate := - poly_tr transfer_sis. +Definition tr_sis f (fi: final) (sis: sistate) := + {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); + si_sreg := poly_tr str_inputs f fi sis.(si_sreg); + si_smem := sis.(si_smem) |}. + +Definition sreg_eval ctx (sis: sistate) (r: reg): option val := + eval_sval ctx (sis.(si_sreg) r). Lemma tr_sis_regs_correct_aux ctx fin sis rs m: sem_sistate ctx sis rs m -> (forall r, sreg_eval ctx (tr_sis (cf ctx) fin sis) r = Some (tr_regs (cf ctx) fin rs) # r). Proof. - destruct 1 as (_ & _ & REG). + unfold sreg_eval; simpl. destruct 1 as (_ & _ & REG). destruct fin; simpl; eauto. -Qed. +Admitted. Lemma tr_sis_regs_correct ctx fin sis rs m: sem_sistate ctx sis rs m -> @@ -629,7 +614,7 @@ Proof. intros H. generalize (tr_sis_regs_correct_aux _ fin _ _ _ H). destruct H as (PRE & MEM & REG). - destruct fin; simpl; econstructor; simpl; intuition eauto. + econstructor; simpl; intuition eauto || congruence. Qed. Definition sexec_final_sfv (i: final) (sis: sistate): sfval := @@ -661,7 +646,7 @@ Lemma sexec_final_svf_correct ctx i sis t rs m s: final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> sem_sfval ctx (sreg_eval ctx (tr_sis (cf ctx) i sis)) (sexec_final_sfv i sis) rs m t s. Proof. - intros (PRE&MEM®). + unfold sreg_eval. intros (PRE&MEM®). destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. + (* Bcall *) intros; eapply exec_Scall; auto. - destruct ros; simpl in * |- *; auto. @@ -975,18 +960,35 @@ Qed. Local Hint Resolve tr_inputs_ext: core. -(* TODO -Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: - sem_sfval ctx sfv rs1 m t s -> - (forall r, rs1#r = rs2#r) -> - exists s', equiv_state s s' /\ sem_sfval ctx sfv rs2 m t s'. +Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f svf: A := + match svf with + | Sgoto pc => tr f [pc] None + | Scall _ _ _ res pc => tr f [pc] (Some res) + | Stailcall _ _ args => tr f [] None + | Sbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res) + | Sreturn _ => tr f [] None + | Sjumptable _ tbl => tr f tbl None + end. + + +Lemma sem_sfval_inv1 rsx ctx sfv rs m t s: + sem_sfval ctx rsx sfv rs m t s -> + (forall r, rsx r = Some (poly_str tr_inputs (cf ctx) sfv rs)#r). +Proof. + destruct 1; simpl; subst; auto. +Admitted. + +Lemma sem_sfval_equivx rsx1 rsx2 ctx sfv rs m t s: + sem_sfval ctx rsx1 sfv rs m t s -> + (forall r, rsx1 r = rsx2 r) -> + exists s', equiv_state s s' /\ sem_sfval ctx rsx2 sfv rs m t s'. Proof. Local Hint Resolve equiv_stack_refl equiv_state_refl regmap_setres_eq: core. Local Hint Constructors sem_sfval: core. - destruct 1; eexists; split; econstructor; eauto. - econstructor; eauto. + destruct 1; eexists; split; econstructor; eauto; try congruence. Qed. -*) + + Local Hint Resolve sexec_final_svf_exact: core. @@ -995,17 +997,6 @@ Definition abort_sistate ctx (sis: sistate): Prop := \/ eval_smem ctx sis.(si_smem) = None \/ exists (r: reg), eval_sval ctx (sis.(si_sreg) r) = None. -Lemma sem_sistate_exclude_abort ctx sis rs m: - sem_sistate ctx sis rs m -> - abort_sistate ctx sis -> - False. -Proof. - intros SIS ABORT. inv SIS. destruct H0 as (H0 & H0'). - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. -Qed. - -Local Hint Resolve sem_sistate_exclude_abort: core. - Lemma set_sreg_preserves_abort ctx sv dst sis: abort_sistate ctx sis -> abort_sistate ctx (set_sreg dst sv sis). @@ -1045,8 +1036,17 @@ Proof. intros; eapply set_smem_preserves_abort; eauto. Qed. +Lemma sem_sistate_tr_sis_exclude_abort ctx sis fi rs m: + sem_sistate ctx (tr_sis (cf ctx) fi sis) rs m -> + abort_sistate ctx sis -> + False. +Proof. + intros ((PRE1 & PRE2) & MEM & REG); simpl in *. + intros [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; try congruence. +Qed. + Local Hint Resolve sexec_op_preserves_abort sexec_load_preserves_abort - sexec_store_preserves_abort: core. + sexec_store_preserves_abort sem_sistate_tr_sis_exclude_abort: core. Lemma sexec_exclude_abort ctx ib t s1: forall sis k (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) @@ -1056,7 +1056,6 @@ Lemma sexec_exclude_abort ctx ib t s1: forall sis k Proof. induction ib; simpl; intros; eauto. - (* final *) inversion SEXEC; subst; eauto. - admit. - (* load *) destruct trap; eauto. inversion SEXEC. - (* seq *) @@ -1065,7 +1064,7 @@ Proof. - (* cond *) inversion SEXEC; subst; eauto. clear SEXEC. destruct b; eauto. -Admitted. +Qed. Lemma set_sreg_abort ctx dst sv sis rs m: sem_sistate ctx sis rs m -> @@ -1124,6 +1123,39 @@ Qed. Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort: core. +(* Remark that we need to reason modulo "extensionality" wrt Regmap.get about regsets. + And, nothing in their representation as (val * PTree.t val) enforces that + (forall r, rs1#r = rs2#r) -> rs1 = rs2 +*) +(* +Lemma sem_sistate_determ ctx sis rs1 m1 rs2 m2: + sem_sistate ctx sis rs1 m1 -> + sem_sistate ctx sis rs2 m2 -> + (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + intros (_&MEM1®1) (_&MEM2®2). + intuition try congruence. + generalize (REG1 r); rewrite REG2; congruence. +Qed. +*) + +Lemma sem_sistate_tr_sis_determ ctx sis rs1 m1 fi rs2 m2: + sem_sistate ctx sis rs1 m1 -> + sem_sistate ctx (tr_sis (cf ctx) fi sis) rs2 m2 -> + (forall r, rs2#r = (tr_regs (cf ctx) fi rs1)#r) + /\ m1 = m2. +Proof. + intros H1 H2. + lapply (tr_sis_regs_correct_aux ctx fi sis rs1 m1); eauto. + unfold sreg_eval; intros X. + destruct H1 as (_&MEM1®1). + destruct H2 as (_&MEM2®2); simpl in *. + intuition try congruence. + cut (Some rs2 # r = Some (tr_regs (cf ctx) fi rs1)#r). + { congruence. } + rewrite <- REG2, X. auto. +Qed. + Lemma sexec_rec_exact ctx ib t s1: forall sis k (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) rs m @@ -1140,13 +1172,14 @@ Proof. induction ib; simpl; intros; eauto. - (* final *) inv SEXEC. - admit. - (* TODO - exploit (sem_sistate_determ ctx sis rs m rs0 m0); eauto. + exploit (sem_sistate_tr_sis_determ ctx sis rs m fi rs' m0); eauto. intros (REG&MEM); subst. - exploit (sem_sfval_equiv rs0 rs); eauto. - intros (s2 & EQUIV & SFV'). - eexists; split; eauto. *) + exploit (sem_sfval_equivx rsx (sreg_eval ctx (tr_sis (cf ctx) fi sis)) ctx ); eauto. + + intros; rewrite EXT, REG. admit. + + intros (s2 & EQUIV & SFV'). admit. +(* eexists; split. + eapply sexec_final_svf_exact; eauto. +*) - (* Bop *) autodestruct; eauto. - destruct trap; [| inv SEXEC ]. repeat autodestruct; eauto. -- cgit From d76f6a446a8c4866d7e89b1b5b13b24a20e6d1dc Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 26 May 2021 21:42:33 +0200 Subject: end of BTL_SEtheory w.r.t fsem --- scheduling/BTL.v | 9 +- scheduling/BTL_SEtheory.v | 331 +++++++++++++++++++++++----------------------- scheduling/BTLroadmap.md | 36 ++--- 3 files changed, 194 insertions(+), 182 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 5f75231a..6f699cd0 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -391,10 +391,11 @@ Definition semantics (p: program) := (** The "functional" small-step semantics for a BTL program. at each exit, we only transfer register in "input_regs" (i.e. "alive" registers). *) - Definition transfer_regs (inputs: list reg) (rs: regset): regset := init_regs (rs##inputs) inputs. +(* TODO: LEMMA BELOW ON transfer_regs USEFUL ? *) + Lemma transfer_regs_inputs inputs rs r: List.In r inputs -> (transfer_regs inputs rs)#r = rs#r. Proof. @@ -456,6 +457,9 @@ Definition inputs_exit (f:function) (tbl: list exit) (or: option reg): Regset.t Definition tr_inputs (f:function) (tbl: list exit) (or:option reg): regset -> regset := transfer_regs (Regset.elements (inputs_exit f tbl or)). +(* TODO: BELOW, LEMMA on tr_inputs USEFUL ? *) + + (* TODO: move this elsewhere *) Lemma SetoidList_InA_eq_equiv A (l: list A): forall x, SetoidList.InA (fun x y => x = y) x l <-> List.In x l. @@ -518,6 +522,7 @@ Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: fina Definition tr_regs: function -> final -> regset -> regset := poly_tr tr_inputs. +(* TODO: NOT USEFUL ? Definition liveout: function -> final -> Regset.t := poly_tr inputs_exit. @@ -533,7 +538,7 @@ Proof. Local Opaque inputs_exit. destruct fi; simpl; auto. Qed. - +*) (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index a066626e..d6b4e741 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -401,57 +401,39 @@ Definition sfind_function ctx (svos : sval + ident): option fundef := Import ListNotations. Local Open Scope list_scope. -(** [sem_sfval] corresponds to [final_step tr_inputs] for symbolic final value. - - A main difference comes from [rsx] which observes the registers saved in the stackframe - of the returned state. - - Intuitively, only liveout registers are stored in the stack - (others will be associated to [Vundef] value). -*) -Inductive sem_sfval ctx (rsx: reg -> option val): sfval -> regset -> mem -> trace -> state -> Prop := - | exec_Sgoto pc rs rs' m: - rs' = tr_inputs ctx.(cf) [pc] None rs -> - (forall r, rsx r = Some rs'#r) -> - sem_sfval ctx rsx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc rs' m) +Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := + | exec_Sgoto pc rs m: + sem_sfval ctx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] None rs) m) | exec_Sreturn stk osv rs m m' v: (csp ctx) = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> match osv with Some sv => eval_sval ctx sv | None => Some Vundef end = Some v -> - (forall r, rsx r = Some Vundef) -> - sem_sfval ctx rsx (Sreturn osv) rs m + sem_sfval ctx (Sreturn osv) rs m E0 (Returnstate (cstk ctx) v m') - | exec_Scall rs m sig svos lsv args res pc fd rs': + | exec_Scall rs m sig svos lsv args res pc fd: sfind_function ctx svos = Some fd -> funsig fd = sig -> eval_list_sval ctx lsv = Some args -> - rs' = tr_inputs ctx.(cf) [pc] (Some res) rs -> - (forall r, rsx r = Some rs'#r) -> - sem_sfval ctx rsx (Scall sig svos lsv res pc) rs m - E0 (Callstate ((Stackframe res (cf ctx) (csp ctx) pc rs')::cstk ctx) fd args m) + sem_sfval ctx (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] (Some res) rs)::cstk ctx) fd args m) | exec_Stailcall stk rs m sig svos args fd m' lsv: sfind_function ctx svos = Some fd -> funsig fd = sig -> (csp ctx) = Vptr stk Ptrofs.zero -> Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> eval_list_sval ctx lsv = Some args -> - (forall r, rsx r = Some Vundef) -> - sem_sfval ctx rsx (Stailcall sig svos lsv) rs m + sem_sfval ctx (Stailcall sig svos lsv) rs m E0 (Callstate (cstk ctx) fd args m') - | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs rs': + | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: seval_builtin_args ctx m sargs vargs -> external_call ef (cge ctx) vargs m t vres m' -> - rs' = tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs -> - (forall r, rsx r = Some rs'#r) -> - sem_sfval ctx rsx (Sbuiltin ef sargs res pc) rs m - t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres rs') m') - | exec_Sjumptable sv tbl pc' n rs m rs': + sem_sfval ctx (Sbuiltin ef sargs res pc) rs m + t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres (tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs)) m') + | exec_Sjumptable sv tbl pc' n rs m: eval_sval ctx sv = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - rs' = tr_inputs ctx.(cf) tbl None rs -> - (forall r, rsx r = Some rs'#r) -> - sem_sfval ctx rsx (Sjumptable sv tbl) rs m - E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' rs' m) + sem_sfval ctx (Sjumptable sv tbl) rs m + E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' (tr_inputs ctx.(cf) tbl None rs) m) . (** * Preservation properties *) @@ -564,59 +546,6 @@ Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop := /\ forall (r:reg), eval_sval ctx (sis.(si_sreg) r) = Some (rs#r). (** * Symbolic execution of final step *) -Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := - match inputs with - | nil => fun r => Sundef - | r1::l => fun r => if Pos.eq_dec r r1 then sreg r1 else transfer_sreg l sreg r - end. - -Lemma transfer_sreg_inputs inputs sreg r: - List.In r inputs -> transfer_sreg inputs sreg r = sreg r. -Proof. - induction inputs; simpl; try autodestruct; intuition congruence. -Qed. - -Lemma transfer_sreg_dead inputs sreg r: - ~List.In r inputs -> transfer_sreg inputs sreg r = Sundef. -Proof. - induction inputs; simpl; try autodestruct; intuition congruence. -Qed. - -Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (inputs_exit f tbl or)). - -Lemma str_inputs_correct ctx sis rs tbl or r: - (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) -> - eval_sval ctx (str_inputs (cf ctx) tbl or (si_sreg sis) r) = Some (tr_inputs (cf ctx) tbl or rs) # r. -Admitted. - -Local Hint Resolve str_inputs_correct: core. - -Definition tr_sis f (fi: final) (sis: sistate) := - {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); - si_sreg := poly_tr str_inputs f fi sis.(si_sreg); - si_smem := sis.(si_smem) |}. - -Definition sreg_eval ctx (sis: sistate) (r: reg): option val := - eval_sval ctx (sis.(si_sreg) r). - -Lemma tr_sis_regs_correct_aux ctx fin sis rs m: - sem_sistate ctx sis rs m -> - (forall r, sreg_eval ctx (tr_sis (cf ctx) fin sis) r = Some (tr_regs (cf ctx) fin rs) # r). -Proof. - unfold sreg_eval; simpl. destruct 1 as (_ & _ & REG). - destruct fin; simpl; eauto. -Admitted. - -Lemma tr_sis_regs_correct ctx fin sis rs m: - sem_sistate ctx sis rs m -> - sem_sistate ctx (tr_sis (cf ctx) fin sis) (tr_regs (cf ctx) fin rs) m. -Proof. - intros H. - generalize (tr_sis_regs_correct_aux _ fin _ _ _ H). - destruct H as (PRE & MEM & REG). - econstructor; simpl; intuition eauto || congruence. -Qed. - Definition sexec_final_sfv (i: final) (sis: sistate): sfval := match i with | Bgoto pc => Sgoto pc @@ -644,9 +573,9 @@ Local Hint Constructors sem_sfval: core. Lemma sexec_final_svf_correct ctx i sis t rs m s: sem_sistate ctx sis rs m -> final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> - sem_sfval ctx (sreg_eval ctx (tr_sis (cf ctx) i sis)) (sexec_final_sfv i sis) rs m t s. + sem_sfval ctx (sexec_final_sfv i sis) rs m t s. Proof. - unfold sreg_eval. intros (PRE&MEM®). + intros (PRE&MEM®). destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. + (* Bcall *) intros; eapply exec_Scall; auto. - destruct ros; simpl in * |- *; auto. @@ -666,7 +595,7 @@ Local Hint Resolve seval_builtin_args_exact: core. Lemma sexec_final_svf_exact ctx i sis t rs m s: sem_sistate ctx sis rs m -> - sem_sfval ctx (sreg_eval ctx (tr_sis (cf ctx) i sis)) (sexec_final_sfv i sis) rs m t s + sem_sfval ctx (sexec_final_sfv i sis) rs m t s -> final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. Proof. intros (PRE&MEM®). @@ -797,6 +726,76 @@ Proof. erewrite MEM; auto. Qed. +(** * Compute sistate associated to final values *) +Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := + match inputs with + | nil => fun r => Sundef + | r1::l => fun r => if Pos.eq_dec r r1 then sreg r1 else transfer_sreg l sreg r + end. + +Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (inputs_exit f tbl or)). + +Lemma str_inputs_correct ctx sis rs tbl or r: + (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) -> + eval_sval ctx (str_inputs (cf ctx) tbl or (si_sreg sis) r) = Some (tr_inputs (cf ctx) tbl or rs) # r. +Proof. + intros H. + unfold str_inputs, tr_inputs, transfer_regs. + induction (Regset.elements _) as [|x l]; simpl. + + rewrite Regmap.gi; auto. + + autodestruct; intros; subst. + * rewrite Regmap.gss; auto. + * rewrite Regmap.gso; eauto. +Qed. + +Local Hint Resolve str_inputs_correct: core. + +Definition tr_sis f (fi: final) (sis: sistate) := + {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); + si_sreg := poly_tr str_inputs f fi sis.(si_sreg); + si_smem := sis.(si_smem) |}. + +Definition sreg_eval ctx (sis: sistate) (r: reg): option val := + eval_sval ctx (sis.(si_sreg) r). + +Lemma tr_sis_regs_correct_aux ctx fin sis rs m: + sem_sistate ctx sis rs m -> + (forall r, sreg_eval ctx (tr_sis (cf ctx) fin sis) r = Some (tr_regs (cf ctx) fin rs) # r). +Proof. + Local Opaque str_inputs. + unfold sreg_eval; simpl. destruct 1 as (_ & _ & REG). + destruct fin; simpl; eauto. +Qed. + +Lemma tr_sis_regs_correct ctx fin sis rs m: + sem_sistate ctx sis rs m -> + sem_sistate ctx (tr_sis (cf ctx) fin sis) (tr_regs (cf ctx) fin rs) m. +Proof. + intros H. + generalize (tr_sis_regs_correct_aux _ fin _ _ _ H). + destruct H as (PRE & MEM & REG). + econstructor; simpl; intuition eauto || congruence. +Qed. + +Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f (svf: sfval): A := + match svf with + | Sgoto pc => tr f [pc] None + | Scall _ _ _ res pc => tr f [pc] (Some res) + | Stailcall _ _ args => tr f [] None + | Sbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res) + | Sreturn _ => tr f [] None + | Sjumptable _ tbl => tr f tbl None + end. + +Definition str_regs: function -> sfval -> regset -> regset := + poly_str tr_inputs. + +Lemma str_tr_regs_equiv f fin sis: + str_regs f (sexec_final_sfv fin sis) = tr_regs f fin. +Proof. + destruct fin; simpl; auto. +Qed. + (** * symbolic execution of blocks *) (* symbolic state *) @@ -808,10 +807,9 @@ Inductive sstate := (* transition (t,s) produced by a sstate in initial context ctx *) Inductive sem_sstate ctx t s: sstate -> Prop := - | sem_Sfinal sis sfv rs rs' rsx m - (SIS: sem_sistate ctx sis rs' m) - (EXT: forall r, rsx r = Some (rs'#r)) - (SFV: sem_sfval ctx rsx sfv rs m t s) + | sem_Sfinal sis sfv rs m + (SIS: sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m) + (SFV: sem_sfval ctx sfv rs m t s) : sem_sstate ctx t s (Sfinal sis sfv) | sem_Scond b cond args sm ifso ifnot (SEVAL: seval_condition ctx cond args sm = Some b) @@ -867,6 +865,9 @@ Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin sem_sstate ctx t s (sexec_rec (cf ctx) ib sis k). Proof. induction ISTEP; simpl; try autodestruct; eauto. + (* final value *) + intros; econstructor; eauto. + rewrite str_tr_regs_equiv; eauto. (* condition *) all: intros; eapply sem_Scond; eauto; [ @@ -875,6 +876,7 @@ Proof. try autodestruct; eauto ]. Qed. + (* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) (sexec is a correct over-approximation) *) @@ -886,6 +888,27 @@ Proof. eapply sexec_rec_correct; simpl; eauto. Qed. +(* Remark that we want to reason modulo "extensionality" wrt Regmap.get about regsets. + And, nothing in their representation as (val * PTree.t val) enforces that + (forall r, rs1#r = rs2#r) -> rs1 = rs2 +*) +Lemma sem_sistate_tr_sis_determ ctx sis rs1 m1 fi rs2 m2: + sem_sistate ctx sis rs1 m1 -> + sem_sistate ctx (tr_sis (cf ctx) fi sis) rs2 m2 -> + (forall r, rs2#r = (tr_regs (cf ctx) fi rs1)#r) + /\ m1 = m2. +Proof. + intros H1 H2. + lapply (tr_sis_regs_correct_aux ctx fi sis rs1 m1); eauto. + unfold sreg_eval; intros X. + destruct H1 as (_&MEM1®1). + destruct H2 as (_&MEM2®2); simpl in *. + intuition try congruence. + cut (Some rs2 # r = Some (tr_regs (cf ctx) fi rs1)#r). + { congruence. } + rewrite <- REG2, X. auto. +Qed. + (* TODO: déplacer les trucs sur equiv_stackframe -> regmap_setres_eq dans BTL! *) Inductive equiv_stackframe: stackframe -> stackframe -> Prop := @@ -958,38 +981,19 @@ Proof. - repeat (rewrite Regmap.gso); auto. Qed. -Local Hint Resolve tr_inputs_ext: core. +(* Local Hint Resolve tr_inputs_ext: core. *) +Local Hint Resolve regmap_setres_eq: core. -Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f svf: A := - match svf with - | Sgoto pc => tr f [pc] None - | Scall _ _ _ res pc => tr f [pc] (Some res) - | Stailcall _ _ args => tr f [] None - | Sbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res) - | Sreturn _ => tr f [] None - | Sjumptable _ tbl => tr f tbl None - end. - - -Lemma sem_sfval_inv1 rsx ctx sfv rs m t s: - sem_sfval ctx rsx sfv rs m t s -> - (forall r, rsx r = Some (poly_str tr_inputs (cf ctx) sfv rs)#r). -Proof. - destruct 1; simpl; subst; auto. -Admitted. - -Lemma sem_sfval_equivx rsx1 rsx2 ctx sfv rs m t s: - sem_sfval ctx rsx1 sfv rs m t s -> - (forall r, rsx1 r = rsx2 r) -> - exists s', equiv_state s s' /\ sem_sfval ctx rsx2 sfv rs m t s'. -Proof. - Local Hint Resolve equiv_stack_refl equiv_state_refl regmap_setres_eq: core. - Local Hint Constructors sem_sfval: core. - destruct 1; eexists; split; econstructor; eauto; try congruence. +Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: + sem_sfval ctx sfv rs1 m t s -> + (forall r, (str_regs (cf ctx) sfv rs1)#r = (str_regs (cf ctx) sfv rs2)#r) -> + exists s', sem_sfval ctx sfv rs2 m t s' /\ equiv_state s s'. +Proof. + unfold str_regs. + destruct 1; simpl in *; intros; subst; eexists; split; econstructor; eauto; try congruence. + constructor; eauto. Qed. - - Local Hint Resolve sexec_final_svf_exact: core. Definition abort_sistate ctx (sis: sistate): Prop := @@ -1123,39 +1127,6 @@ Qed. Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort: core. -(* Remark that we need to reason modulo "extensionality" wrt Regmap.get about regsets. - And, nothing in their representation as (val * PTree.t val) enforces that - (forall r, rs1#r = rs2#r) -> rs1 = rs2 -*) -(* -Lemma sem_sistate_determ ctx sis rs1 m1 rs2 m2: - sem_sistate ctx sis rs1 m1 -> - sem_sistate ctx sis rs2 m2 -> - (forall r, rs1#r = rs2#r) /\ m1 = m2. -Proof. - intros (_&MEM1®1) (_&MEM2®2). - intuition try congruence. - generalize (REG1 r); rewrite REG2; congruence. -Qed. -*) - -Lemma sem_sistate_tr_sis_determ ctx sis rs1 m1 fi rs2 m2: - sem_sistate ctx sis rs1 m1 -> - sem_sistate ctx (tr_sis (cf ctx) fi sis) rs2 m2 -> - (forall r, rs2#r = (tr_regs (cf ctx) fi rs1)#r) - /\ m1 = m2. -Proof. - intros H1 H2. - lapply (tr_sis_regs_correct_aux ctx fi sis rs1 m1); eauto. - unfold sreg_eval; intros X. - destruct H1 as (_&MEM1®1). - destruct H2 as (_&MEM2®2); simpl in *. - intuition try congruence. - cut (Some rs2 # r = Some (tr_regs (cf ctx) fi rs1)#r). - { congruence. } - rewrite <- REG2, X. auto. -Qed. - Lemma sexec_rec_exact ctx ib t s1: forall sis k (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) rs m @@ -1172,14 +1143,11 @@ Proof. induction ib; simpl; intros; eauto. - (* final *) inv SEXEC. - exploit (sem_sistate_tr_sis_determ ctx sis rs m fi rs' m0); eauto. + exploit (sem_sistate_tr_sis_determ ctx sis rs m fi); eauto. intros (REG&MEM); subst. - exploit (sem_sfval_equivx rsx (sreg_eval ctx (tr_sis (cf ctx) fi sis)) ctx ); eauto. - + intros; rewrite EXT, REG. admit. - + intros (s2 & EQUIV & SFV'). admit. -(* eexists; split. - eapply sexec_final_svf_exact; eauto. -*) + exploit (sem_sfval_equiv rs0 rs); eauto. + * intros; rewrite REG, str_tr_regs_equiv; auto. + * intros (s2 & EQUIV & SFV'); eauto. - (* Bop *) autodestruct; eauto. - destruct trap; [| inv SEXEC ]. repeat autodestruct; eauto. @@ -1205,7 +1173,7 @@ Proof. destruct b. + exploit IHib1; eauto. + exploit IHib2; eauto. -Admitted. +Qed. (* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution @@ -1227,3 +1195,42 @@ Proof. inversion SEXEC. Qed. +(** * High-Level specification of the symbolic simulation test as predicate [sstate_simu] *) + +Record simu_proof_context := Sctx { + sge1: BTL.genv; + sge2: BTL.genv; + sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; + sstk1: list BTL.stackframe; + sstk2: list BTL.stackframe; + sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; + sf1: BTL.function; + sf2: BTL.function; + ssp: val; + srs0: regset; + sm0: mem +}. + +Definition bctx1 (ctx: simu_proof_context):= Bctx ctx.(sge1) ctx.(sstk1) ctx.(sf1) ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx2 (ctx: simu_proof_context):= Bctx ctx.(sge2) ctx.(sstk2) ctx.(sf2) ctx.(ssp) ctx.(srs0) ctx.(sm0). + +Definition sstate_simu (ctx: simu_proof_context) (st1 st2: sstate) := + forall t s1, sem_sstate (bctx1 ctx) t s1 st1 -> + exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2. + +Theorem sstate_simu_correct ctx ib1 ib2: + sstate_simu ctx (sexec (sf1 ctx) ib1) (sexec (sf2 ctx) ib2) -> + forall t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) (sf1 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> + exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) (sf2 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. +Proof. + unfold sstate_simu. + intros SIMU t s1 STEP1. + exploit (sexec_correct (bctx1 ctx)); simpl; eauto. + intros; exploit SIMU; eauto. + intros (s2 & SEM1 & EQ1). + exploit (sexec_exact (bctx2 ctx)); simpl; eauto. + intros (s3 & STEP2 & EQ2). + clear STEP1; eexists; split; eauto. + eapply equiv_state_trans; eauto. +Qed. + diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 6042803d..0cbcd7d6 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -117,37 +117,37 @@ partir du bloc d'entrée de la fonction. Ça semble complémentaire de l'analyse de "liveout". Mais utilisable uniquement dans le cadre d'une combinaison "exécution symbolique"+"value analysis" (donc pas tout de suite). -**IMPLEM PLAN** - -1. définir the "functional semantics" of BTL (dans un premier temps, avec uniquement Vundef en sortie de bloc). Rem: en pratique, la "functional semantics" est définie pour être une "concrétisation" de la sémantique symbolique. Les définitions des deux sémantiques (symbolique et "functional") se font donc en simultanée. -2. prouver la simulation BTL.FunctionalSemantics -> BTL.CFGSemantics (facile normalement: c'est du "less_undef lockstep"). -3. implémenter une passe BTL.CFGSemantics -> BTL.FunctionalSemantics. - en gros: on implémente un vérificateur des infos de liveness. - c'est la correction du "input_regs" qui garantit que la simulation est correct. - La preuve devrait normalement être très similaire à RTLpathLivegenproof. +Donc, dans un premier temps, la "functional semantics" encode/abstrait la notion de "liveout" pour le test d'exécution symbolique. +Les définitions des deux sémantiques (symbolique et "functional") se font donc en simultanée. **STATUS** 1. See [BTL.fsem] (Pour l'instant CFGSemantics => BTL.semantics et FunctionalSemantics => BTL.fsem) +**TODO** + +1. prouver la simulation BTL.FunctionalSemantics -> BTL.CFGSemantics (facile normalement: c'est du "less_undef lockstep"). +2. implémenter une passe BTL.CFGSemantics -> BTL.FunctionalSemantics. + en gros: on implémente un vérificateur des infos de liveness. + c'est la correction du "input_regs" qui garantit que la simulation est correct. + La preuve devrait normalement être très similaire à RTLpathLivegenproof. + ## "Basic" symbolic execution / symbolic simulation We will implement a "basic" (e.g without rewriting rules) simulation test for BTL.FunctionalSemantics. Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter de tester l'égalité des registres de chaque côté: les registres hors liveout seront égaux à Vundef de chaque côté. -**IMPLEM NOTE** - -- use Continuation Passing Style for an easy recursive generation of "all execution paths". -- dans l'etat concret, prévoir un booleen qui indique la valeur symbolique par defaut des registres: deux cas, soit "Input" (valeur initiale du registre), soit "Sundef" (valeur Vundef). +**STATUS** -**CURRENT STATUS** +- BTL_SEtheory: DONE + - model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms wrt BTL.fsem + - high-level specification [sstate_simu] of "symbolic simulation" over iblocks (wrt BTL.fsem). -- model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms wrt BTL.semantics +**TODO** -**next steps** -- preservation proof w.r.t BTL.fsem. -- high-level specification of "symbolic simulation" -- ... +1. Verif du "scheduling" (aka analogue de RTLpathScheduler & RTLSchedulerproof). +2. raffinement intermediaire avant le hash-consing ? ça permettrait de décomposer encore davantage la preuve ? +Y a-t-il un moyen simple de faire le tests d'inclusion des fails potentiels à coût linéaire plutôt que quadratique (contrairement à RTLpath) ? ## Port of RTLpath optimizations to BTL -- cgit From e30376ce891d0710757c65e85a24e7d2550a37ed Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 27 May 2021 13:49:35 +0200 Subject: cleaning BTL_SEtheory --- scheduling/BTL.v | 73 +++++++++++ scheduling/BTL_SEtheory.v | 304 +++++++++++++++++----------------------------- 2 files changed, 187 insertions(+), 190 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 6f699cd0..bf19f4ac 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -540,6 +540,79 @@ Proof. Qed. *) +(* * Reasoning modulo extensionality of [regset] into BTL semantics *) +Inductive equiv_stackframe: stackframe -> stackframe -> Prop := + | equiv_stackframe_intro res f sp pc rs1 rs2 + (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) + . + +Inductive equiv_state: state -> state -> Prop := + | State_equiv stack f sp pc rs1 m rs2 + (EQUIV: forall r, rs1#r = rs2#r): + equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) + | Call_equiv stk stk' f args m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Callstate stk f args m) (Callstate stk' f args m) + | Return_equiv stk stk' v m + (STACKS: list_forall2 equiv_stackframe stk stk'): + equiv_state (Returnstate stk v m) (Returnstate stk' v m) + . + +Local Hint Constructors equiv_stackframe equiv_state: core. + +Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf. +Proof. + destruct stf; eauto. +Qed. + +Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. +Proof. + Local Hint Resolve equiv_stackframe_refl: core. + induction stk; simpl; constructor; auto. +Qed. + +Lemma equiv_state_refl s: equiv_state s s. +Proof. + Local Hint Resolve equiv_stack_refl: core. + induction s; simpl; constructor; auto. +Qed. + +Lemma equiv_stackframe_trans stf1 stf2 stf3: + equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3. +Proof. + destruct 1; intros EQ; inv EQ; try econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +Lemma equiv_stack_trans stk1 stk2: + list_forall2 equiv_stackframe stk1 stk2 -> + forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> + list_forall2 equiv_stackframe stk1 stk3. +Proof. + Local Hint Resolve equiv_stackframe_trans: core. + induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. +Qed. + +Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3. +Proof. + Local Hint Resolve equiv_stack_trans: core. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +Lemma regmap_setres_eq (rs rs': regset) res vres: + (forall r, rs # r = rs' # r) -> + forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r. +Proof. + intros RSEQ r. destruct res; simpl; try congruence. + destruct (peq x r). + - subst. repeat (rewrite Regmap.gss). reflexivity. + - repeat (rewrite Regmap.gso); auto. +Qed. + + + (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) (** Matching BTL and RTL code diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index d6b4e741..79db99c8 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1,20 +1,19 @@ -(* A theory of symbolic execution on BTL +(** A theory of symbolic simulation (i.e. simulation of symbolic executions) on BTL blocks. NB: an efficient implementation with hash-consing will be defined in another file (some day) -*) +The main theorem of this file is [symbolic_simu_correct] stating +that the abstract definition of symbolic simulation of two BTL blocks +implies the simulation for BTL.fsem block-steps. + -(* TODO: A REVOIR - remplacer les [tid] par [tr_inputs] pour montrer la bisimulation avec [fsem] plutôt. *) +*) Require Import Coqlib Maps Floats. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL BTL OptionMonad. -(* TODO remove this, when copy-paste of RTLpathSE_theory is clearly over... *) -Ltac inversion_SOME := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) -Ltac inversion_ASSERT := fail. (* deprecated tactic of OptionMonad: use autodestruct instead *) - Record iblock_exec_context := Bctx { cge: BTL.genv; cstk: list stackframe; @@ -113,7 +112,7 @@ Definition seval_condition ctx (cond: condition) (lsv: list_sval) (sm: smem) : o (** * Auxiliary definitions on Builtins *) -(* TODO: clean this. Some (cge ctx)neric stuffs could be put in [AST.v] *) +(* TODO: clean this. Some generic stuffs could be put in [AST.v] *) Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *) @@ -170,7 +169,7 @@ Qed. End SEVAL_BUILTIN_ARG. -(* NB: (cge ctx)neric function that could be put into [AST] file *) +(* NB: generic function that could be put into [AST] file *) Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B := match arg with | BA x => BA (f x) @@ -436,105 +435,6 @@ Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' (tr_inputs ctx.(cf) tbl None rs) m) . -(** * Preservation properties *) - -Section SymbValPreserved. - -Variable ge ge': BTL.genv. - -Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. - -Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. - -Lemma senv_find_symbol_preserved id: - Senv.find_symbol ge id = Senv.find_symbol ge' id. -Proof. - destruct senv_preserved_BTL as (A & B & C). congruence. -Qed. - -Lemma senv_symbol_address_preserved id ofs: - Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. -Proof. - unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. - reflexivity. -Qed. - -Variable stk stk': list stackframe. -Variable f f': function. -Variable sp: val. -Variable rs0: regset. -Variable m0: mem. - -Lemma eval_sval_preserved sv: - eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. -Proof. - induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) - (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. - erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. - rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. - rewrite IHsv1; auto. -Qed. - -Lemma seval_builtin_arg_preserved m: forall bs varg, - seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> - seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. -Proof. - induction 1; simpl. - all: try (constructor; auto). - - rewrite <- eval_sval_preserved. assumption. - - rewrite <- senv_symbol_address_preserved. assumption. - - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. -Qed. - -Lemma seval_builtin_args_preserved m lbs vargs: - seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> - seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. -Proof. - induction 1; constructor; eauto. - eapply seval_builtin_arg_preserved; auto. -Qed. - -Lemma list_sval_eval_preserved lsv: - eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. -Proof. - induction lsv; simpl; auto. - rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. - rewrite IHlsv; auto. -Qed. - -Lemma smem_eval_preserved sm: - eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. -Proof. - induction sm; simpl; auto. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. - rewrite eval_sval_preserved; auto. -Qed. - -Lemma seval_condition_preserved cond lsv sm: - seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. -Proof. - unfold seval_condition. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - rewrite smem_eval_preserved; auto. -Qed. - -End SymbValPreserved. - - (* Syntax and Semantics of symbolic internal states *) (* [si_pre] is a precondition on initial context *) Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg: reg -> sval; si_smem: smem }. @@ -909,80 +809,8 @@ Proof. rewrite <- REG2, X. auto. Qed. - -(* TODO: déplacer les trucs sur equiv_stackframe -> regmap_setres_eq dans BTL! *) -Inductive equiv_stackframe: stackframe -> stackframe -> Prop := - | equiv_stackframe_intro res f sp pc rs1 rs2 - (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): - equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) - . - -Inductive equiv_state: state -> state -> Prop := - | State_equiv stack f sp pc rs1 m rs2 - (EQUIV: forall r, rs1#r = rs2#r): - equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) - | Call_equiv stk stk' f args m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Callstate stk f args m) (Callstate stk' f args m) - | Return_equiv stk stk' v m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Returnstate stk v m) (Returnstate stk' v m) - . - -Local Hint Constructors equiv_stackframe equiv_state: core. - -Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf. -Proof. - destruct stf; eauto. -Qed. - -Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. -Proof. - Local Hint Resolve equiv_stackframe_refl: core. - induction stk; simpl; constructor; auto. -Qed. - -Lemma equiv_state_refl s: equiv_state s s. -Proof. - Local Hint Resolve equiv_stack_refl: core. - induction s; simpl; constructor; auto. -Qed. - -Lemma equiv_stackframe_trans stf1 stf2 stf3: - equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3. -Proof. - destruct 1; intros EQ; inv EQ; try econstructor; eauto. - intros; eapply eq_trans; eauto. -Qed. - -Lemma equiv_stack_trans stk1 stk2: - list_forall2 equiv_stackframe stk1 stk2 -> - forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> - list_forall2 equiv_stackframe stk1 stk3. -Proof. - Local Hint Resolve equiv_stackframe_trans: core. - induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. -Qed. - -Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3. -Proof. - Local Hint Resolve equiv_stack_trans: core. - destruct 1; intros EQ; inv EQ; econstructor; eauto. - intros; eapply eq_trans; eauto. -Qed. - -Lemma regmap_setres_eq (rs rs': regset) res vres: - (forall r, rs # r = rs' # r) -> - forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r. -Proof. - intros RSEQ r. destruct res; simpl; try congruence. - destruct (peq x r). - - subst. repeat (rewrite Regmap.gss). reflexivity. - - repeat (rewrite Regmap.gso); auto. -Qed. - -(* Local Hint Resolve tr_inputs_ext: core. *) -Local Hint Resolve regmap_setres_eq: core. +Local Hint Constructors equiv_stackframe list_forall2: core. +Local Hint Resolve regmap_setres_eq equiv_stack_refl equiv_stack_refl: core. Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: sem_sfval ctx sfv rs1 m t s -> @@ -991,11 +819,8 @@ Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: Proof. unfold str_regs. destruct 1; simpl in *; intros; subst; eexists; split; econstructor; eauto; try congruence. - constructor; eauto. Qed. -Local Hint Resolve sexec_final_svf_exact: core. - Definition abort_sistate ctx (sis: sistate): Prop := ~(sis.(si_pre) ctx) \/ eval_smem ctx sis.(si_smem) = None @@ -1125,7 +950,7 @@ Proof. intros; rewrite REG; autodestruct; try_simplify_someHyps. Qed. -Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort: core. +Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort sexec_final_svf_exact: core. Lemma sexec_rec_exact ctx ib t s1: forall sis k (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) @@ -1195,7 +1020,7 @@ Proof. inversion SEXEC. Qed. -(** * High-Level specification of the symbolic simulation test as predicate [sstate_simu] *) +(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *) Record simu_proof_context := Sctx { sge1: BTL.genv; @@ -1218,12 +1043,14 @@ Definition sstate_simu (ctx: simu_proof_context) (st1 st2: sstate) := forall t s1, sem_sstate (bctx1 ctx) t s1 st1 -> exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2. -Theorem sstate_simu_correct ctx ib1 ib2: - sstate_simu ctx (sexec (sf1 ctx) ib1) (sexec (sf2 ctx) ib2) -> +Definition symbolic_simu ctx ib1 ib2: Prop := sstate_simu ctx (sexec (sf1 ctx) ib1) (sexec (sf2 ctx) ib2). + +Theorem symbolic_simu_correct ctx ib1 ib2: + symbolic_simu ctx ib1 ib2 -> forall t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) (sf1 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) (sf2 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. Proof. - unfold sstate_simu. + unfold symbolic_simu, sstate_simu. intros SIMU t s1 STEP1. exploit (sexec_correct (bctx1 ctx)); simpl; eauto. intros; exploit SIMU; eauto. @@ -1234,3 +1061,100 @@ Proof. eapply equiv_state_trans; eauto. Qed. +(** * Preservation properties *) + +Section SymbValPreserved. + +Variable ge ge': BTL.genv. + +Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. + +Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol ge id = Senv.find_symbol ge' id. +Proof. + destruct senv_preserved_BTL as (A & B & C). congruence. +Qed. + +Lemma senv_symbol_address_preserved id ofs: + Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. +Proof. + unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. + reflexivity. +Qed. + +Variable stk stk': list stackframe. +Variable f f': function. +Variable sp: val. +Variable rs0: regset. +Variable m0: mem. + +Lemma eval_sval_preserved sv: + eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. +Proof. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) + (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma seval_builtin_arg_preserved m: forall bs varg, + seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> + seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. +Proof. + induction 1; simpl. + all: try (constructor; auto). + - rewrite <- eval_sval_preserved. assumption. + - rewrite <- senv_symbol_address_preserved. assumption. + - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. +Qed. + +Lemma seval_builtin_args_preserved m lbs vargs: + seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> + seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. +Proof. + induction 1; constructor; eauto. + eapply seval_builtin_arg_preserved; auto. +Qed. + +Lemma list_sval_eval_preserved lsv: + eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. +Proof. + induction lsv; simpl; auto. + rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma smem_eval_preserved sm: + eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. + rewrite eval_sval_preserved; auto. +Qed. + +Lemma seval_condition_preserved cond lsv sm: + seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. +Proof. + unfold seval_condition. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + +End SymbValPreserved. -- cgit From 1a78c940f46273b7146d2111b1e2da309434f021 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 27 May 2021 16:55:18 +0200 Subject: [disabled checker] BTL Scheduling and Renumbering OK! --- scheduling/BTL.v | 4 +- scheduling/BTLRenumber.ml | 111 ++++++++++++++++++++++++++++++++++++++++++ scheduling/BTLScheduleraux.ml | 94 +++++++++++++++++++++++------------ scheduling/BTLaux.ml | 3 -- scheduling/BTLcommonaux.ml | 84 ++++++++++++++++++++++++++++++++ scheduling/BTLrenumber.ml | 50 ------------------- scheduling/BTLtoRTL.v | 2 +- scheduling/BTLtoRTLaux.ml | 51 ++++++++----------- scheduling/BTLtoRTLproof.v | 4 +- scheduling/BTLtypes.ml | 7 +++ scheduling/PrintBTL.ml | 67 +++++++++++++------------ scheduling/RTLtoBTL.v | 2 +- scheduling/RTLtoBTLaux.ml | 46 +++++++---------- scheduling/RTLtoBTLproof.v | 4 +- 14 files changed, 348 insertions(+), 181 deletions(-) create mode 100644 scheduling/BTLRenumber.ml delete mode 100644 scheduling/BTLaux.ml create mode 100644 scheduling/BTLcommonaux.ml delete mode 100644 scheduling/BTLrenumber.ml create mode 100644 scheduling/BTLtypes.ml (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 0f9ef44f..991361ca 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -25,11 +25,11 @@ Definition exit := node. (* we may generalize this with register renamings at e (* inst_info is a ghost record to provide instruction information through oracles *) Parameter inst_info: Set. -Extract Constant inst_info => "BTLaux.inst_info". +Extract Constant inst_info => "BTLtypes.inst_info". (* block_info is a ghost record to provide block information through oracles *) Parameter block_info: Set. -Extract Constant block_info => "BTLaux.block_info". +Extract Constant block_info => "BTLtypes.block_info". (** final instructions (that stops block execution) *) Inductive final: Type := diff --git a/scheduling/BTLRenumber.ml b/scheduling/BTLRenumber.ml new file mode 100644 index 00000000..36f3bcf5 --- /dev/null +++ b/scheduling/BTLRenumber.ml @@ -0,0 +1,111 @@ +open Maps +open BTL +open RTLcommonaux +open BTLcommonaux +open BTLtypes + +let recompute_inumbs btl entry = + let btl = reset_visited_ib (reset_visited_ibf btl) in + let last_used = ref 0 in + let ibf = get_some @@ PTree.get entry btl in + let ipos () = + last_used := !last_used + 1; + !last_used + in + let rec walk ib k = + (* heuristic: try to explore the lower numbered branch first *) + let walk_smallest_child s1 s2 ib1 ib2 = + if s1 < s2 then ( + walk ib1 None; + walk ib2 None) + else ( + walk ib2 None; + walk ib1 None) + in + if jump_visit ib then () + else + match ib with + | BF (Bcall (_, _, _, _, s), iinfo) | BF (Bbuiltin (_, _, _, s), iinfo) -> + let ib' = (get_some @@ PTree.get s btl).entry in + walk ib' None; + iinfo.inumb <- ipos () + | BF (Bgoto s, _) -> + let ib' = (get_some @@ PTree.get s btl).entry in + walk ib' None + | BF (Bjumptable (_, tbl), iinfo) -> + List.iter + (fun s -> + let ib' = (get_some @@ PTree.get s btl).entry in + walk ib' None) + tbl; + iinfo.inumb <- ipos () + | BF (Btailcall (_, _, _), iinfo) | BF (Breturn _, iinfo) -> + iinfo.inumb <- ipos () + | Bnop None -> + failwith + "recompute_inumbs: Bnop None can only be in the right child of \ + Bcond" + | Bnop (Some iinfo) + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) -> + let succ = get_some @@ k in + walk succ None; + iinfo.inumb <- ipos () + | Bseq (ib1, ib2) -> walk ib1 (Some ib2) + | Bcond (_, _, BF (Bgoto s1, iinfoL), BF (Bgoto s2, iinfoR), iinfo) -> + iinfoL.visited <- true; + iinfoR.visited <- true; + let ib1 = get_some @@ PTree.get s1 btl in + let ib2 = get_some @@ PTree.get s2 btl in + walk_smallest_child (p2i s1) (p2i s2) ib1.entry ib2.entry; + iinfo.inumb <- ipos () + | Bcond (_, _, BF (Bgoto s1, iinfoL), Bnop None, iinfoF) -> + iinfoL.visited <- true; + let ib1 = get_some @@ PTree.get s1 btl in + let ib2 = get_some @@ k in + walk_smallest_child (p2i s1) (get_inumb_or_next ib2) ib1.entry ib2; + iinfoF.inumb <- ipos () + | Bcond (_, _, _, _, _) -> failwith "recompute_inumbs: unsupported Bcond" + in + walk ibf.entry None; + btl + +let regenerate_btl_tree btl entry = + let new_entry = ref entry in + let rec renumber_iblock ib = + let get_new_succ s = + let sentry = get_some @@ PTree.get s btl in + i2p (get_inumb_or_next sentry.entry) + in + match ib with + | BF (Bcall (sign, fn, lr, rd, s), iinfo) -> + BF (Bcall (sign, fn, lr, rd, get_new_succ s), iinfo) + | BF (Bbuiltin (sign, fn, lr, s), iinfo) -> + BF (Bbuiltin (sign, fn, lr, get_new_succ s), iinfo) + | BF (Bgoto s, iinfo) -> BF (Bgoto (get_new_succ s), iinfo) + | BF (Bjumptable (arg, tbl), iinfo) -> + let tbl' = List.map (fun s -> get_new_succ s) tbl in + BF (Bjumptable (arg, tbl'), iinfo) + | Bcond (cond, lr, ib1, ib2, iinfo) -> + Bcond (cond, lr, renumber_iblock ib1, renumber_iblock ib2, iinfo) + | Bseq (ib1, ib2) -> Bseq (renumber_iblock ib1, renumber_iblock ib2) + | _ -> ib + in + let ord_btl = + PTree.fold + (fun ord_btl old_n old_ibf -> + let ib = renumber_iblock old_ibf.entry in + let n = get_inumb_or_next ib in + let n_pos = i2p n in + let bi = mk_binfo n in + let ibf = { entry = ib; input_regs = old_ibf.input_regs; binfo = bi } in + if old_n = entry then new_entry := n_pos; + PTree.set n_pos ibf ord_btl) + btl PTree.empty + in + (ord_btl, !new_entry) + +let renumber btl entry = + let btl' = recompute_inumbs btl entry in + regenerate_btl_tree btl' entry diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 699538ca..4b5ebb32 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -2,6 +2,7 @@ open AST open Maps open Registers open BTL +open BTLtypes open DebugPrint open RTLcommonaux open InstructionScheduler @@ -108,7 +109,7 @@ let build_constraints_and_resources (opweights : opweights) insts btl = add_output_mem i; let rs = opweights.resources_of_store chk addr (List.length lr) in resources := rs :: !resources - | Bcond (cond, lr, BF (Bgoto s, _), Bnop _, _) -> + | Bcond (cond, lr, BF (Bgoto s, _), ibnot, _) -> (* TODO gourdinl test with/out this line *) let live = (get_some @@ PTree.get s btl).input_regs in add_input_regs i (Regset.elements live); @@ -132,11 +133,7 @@ let define_problem (opweights : opweights) ibf btl = max_latency = -1; resource_bounds = opweights.pipelined_resource_bounds; instruction_usages = resources; - latency_constraints = - (* if (use_alias_analysis ()) - then (get_alias_dependencies seqa) @ simple_deps - else *) - simple_deps; + latency_constraints = simple_deps; } let zigzag_scheduler problem early_ones = @@ -161,9 +158,16 @@ let zigzag_scheduler problem early_ones = { problem with latency_constraints = !constraints' } | None -> None -let prepass_scheduler_by_name name problem early_ones = +let prepass_scheduler_by_name name problem insts = match name with - | "zigzag" -> zigzag_scheduler problem early_ones + | "zigzag" -> + let early_ones = + Array.map + (fun inst -> + match inst with Bcond (_, _, _, _, _) -> true | _ -> false) + insts + in + zigzag_scheduler problem early_ones | _ -> scheduler_by_name name problem let schedule_sequence insts btl = @@ -174,13 +178,7 @@ let schedule_sequence insts btl = let nr_instructions = Array.length insts in let problem = define_problem opweights insts btl in match - prepass_scheduler_by_name - !Clflags.option_fprepass_sched - problem - (Array.map - (fun inst -> - match inst with Bcond (_, _, _, _, _) -> true | _ -> false) - insts) + prepass_scheduler_by_name !Clflags.option_fprepass_sched problem insts with | None -> Printf.printf "no solution in prepass scheduling\n"; @@ -199,23 +197,55 @@ let schedule_sequence insts btl = let flatten_blk_basics ibf = let ib = ibf.entry in + let last = ref None in let rec traverse_blk ib = match ib with - | BF (_, _) - | Bcond (_, _, BF (Bgoto _, _), BF (Bgoto _, _), _) -> [] - | Bseq (ib1, ib2) -> - traverse_blk ib1 @ traverse_blk ib2 - | _ -> [ib] - in - Array.of_list (traverse_blk ib) + | BF (_, _) -> + last := Some ib; + [] + | Bseq (ib1, ib2) -> traverse_blk ib1 @ traverse_blk ib2 + | Bcond (_, _, _, _, iinfo) -> ( + match iinfo.pcond with + | Some _ -> [ ib ] + | None -> + last := Some ib; + []) + | _ -> [ ib ] + in + let ibl = traverse_blk ib in + (Array.of_list ibl, !last) -let btl_scheduler btl entry = - List.iter (fun (n,ibf) -> - let bseq = flatten_blk_basics ibf in - match schedule_sequence bseq btl with - | Some positions -> - Array.iter (fun p -> debug "%d " p) positions - | None -> () - ) (PTree.elements btl); - (*let seqs = get_sequences seqs in*) - () +let apply_schedule bseq olast positions = + let ibl = Array.to_list (Array.map (fun i -> bseq.(i)) positions) in + let rec build_iblock = function + | [] -> failwith "build_iblock: empty list" + | [ ib ] -> ( match olast with Some last -> Bseq (ib, last) | None -> ib) + | ib1 :: ib2 :: k -> Bseq (ib1, build_iblock (ib2 :: k)) + in + build_iblock ibl + +let schedule_blk n ibf btl = + let bseq, olast = flatten_blk_basics ibf in + match schedule_sequence bseq btl with + | Some positions -> + debug "%d," (p2i n); + Array.iter (fun p -> debug "%d " p) positions; + debug "\n"; + let new_ib = apply_schedule bseq olast positions in + let new_ibf = + { entry = new_ib; binfo = ibf.binfo; input_regs = ibf.input_regs } + in + PTree.set n new_ibf btl + | None -> btl + +let rec do_schedule btl = function + | [] -> btl + | (n, ibf) :: blks -> + let btl' = schedule_blk n ibf btl in + do_schedule btl' blks + +let btl_scheduler btl = + (*debug_flag := true;*) + let btl' = do_schedule btl (PTree.elements btl) in + (*debug_flag := false;*) + btl' diff --git a/scheduling/BTLaux.ml b/scheduling/BTLaux.ml deleted file mode 100644 index ca34c21c..00000000 --- a/scheduling/BTLaux.ml +++ /dev/null @@ -1,3 +0,0 @@ -type inst_info = { mutable inumb : int; mutable pcond : bool option } - -type block_info = { mutable bnumb : int; mutable visited : bool } diff --git a/scheduling/BTLcommonaux.ml b/scheduling/BTLcommonaux.ml new file mode 100644 index 00000000..dabf760a --- /dev/null +++ b/scheduling/BTLcommonaux.ml @@ -0,0 +1,84 @@ +open Maps +open BTL +open BTLtypes +open RTLcommonaux + +let undef_node = -1 + +let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond; visited = false } + +let def_iinfo () = { inumb = undef_node; pcond = None; visited = false } + +let mk_binfo _bnumb = { bnumb = _bnumb; visited = false } + +let reset_visited_ibf btl = + PTree.map + (fun n ibf -> + ibf.binfo.visited <- false; + ibf) + btl + +let reset_visited_ib btl = + List.iter + (fun (n, ibf) -> + let ib = ibf.entry in + let rec reset_visited_ib_rec ib = + match ib with + | Bseq (ib1, ib2) -> + reset_visited_ib_rec ib1; + reset_visited_ib_rec ib2 + | Bcond (_, _, ib1, ib2, iinfo) -> + reset_visited_ib_rec ib1; + reset_visited_ib_rec ib2; + iinfo.visited <- false + | Bnop (Some iinfo) + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) + | BF (_, iinfo) -> + iinfo.visited <- false + | _ -> () + in + reset_visited_ib_rec ib) + (PTree.elements btl); + btl + +let jump_visit = function + | Bcond (_, _, _, _, iinfo) + | Bnop (Some iinfo) + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) + | BF (_, iinfo) -> + if iinfo.visited then true + else ( + iinfo.visited <- true; + false) + | Bseq (_, _) -> false + | Bnop None -> true + +let rec get_inumb_or_next = function + | BF (Bgoto s, _) -> p2i s + | BF (_, iinfo) + | Bnop (Some iinfo) + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) + | Bcond (_, _, _, _, iinfo) -> + iinfo.inumb + | Bseq (ib1, _) -> get_inumb_or_next ib1 + | _ -> failwith "get_inumb_or_next: Bnop None" + +let rec set_next_inumb btl pos = function + | BF (Bgoto s, _) -> + let ib' = (get_some @@ PTree.get s btl).entry in + set_next_inumb btl pos ib' + | BF (_, iinfo) + | Bnop (Some iinfo) + | Bop (_, _, _, iinfo) + | Bload (_, _, _, _, _, iinfo) + | Bstore (_, _, _, _, iinfo) + | Bcond (_, _, _, _, iinfo) -> + iinfo.inumb <- pos + | Bseq (ib1, _) -> set_next_inumb btl pos ib1 + | _ -> failwith "get_inumb_or_next: Bnop None" diff --git a/scheduling/BTLrenumber.ml b/scheduling/BTLrenumber.ml deleted file mode 100644 index f9cacd6a..00000000 --- a/scheduling/BTLrenumber.ml +++ /dev/null @@ -1,50 +0,0 @@ -(* XXX uncomment -open !BTL -open DebugPrint -open RTLcommonaux -open Maps*) - -(** A quick note on the BTL renumber algorithm - This is a simple first version where we try to reuse the entry numbering from RTL. - In the futur, it would be great to implement a strongly connected components partitionning. - The numbering is done by a postorder traversal. - TODO gourdinl : note perso - - faire un comptage global du nombre d'instructions rtl à générer - - utiliser la structure d'info créée lors de la génération pour tenter de préserver au max - la relation d'ordre, avec une heuristique genre (1 + max des succs) pour l'insertion - - quand il y a un branchement conditionnel, privilégier le parcour du fils gauche en premier - (afin d'avoir de plus petits numéros à gauche) -*) -(* -let rec get_max_rtl_ninsts code entry = -let rec postorder_blk ib = -*) - -(* XXX uncomment -let postorder code entry = - let renumbered_code = ref PTree.empty in - let rec renum_ibf e = - let ibf = get_some @@ PTree.get e code in - if ibf.binfo.visited then () - else ( - ibf.binfo.visited <- true; - let next_nodes = ref [] in - let ib = ibf.entry in - let rec renum_iblock ib = - match ib with - in - let rib = renum_iblock ib in - ibf.entry <- rib; - renumbered_code := PTree.set e ibf !renumbered_code; - let succs = !next_nodes in - List.iter renum_ibf succs) - in - renum_ibf entry - -let renumber (f: BTL.coq_function) = - debug_flag := true; - let code = f.fn_code in - let entry = f.fn_entrypoint in - let renumbered_code = postorder code entry in - debug_flag := false; - renumbered_code*) diff --git a/scheduling/BTLtoRTL.v b/scheduling/BTLtoRTL.v index 82ad47b1..1333b406 100644 --- a/scheduling/BTLtoRTL.v +++ b/scheduling/BTLtoRTL.v @@ -15,7 +15,7 @@ Definition transf_function (f: function) : res RTL.function := let (tcte, dupmap) := btl2rtl f in let (tc, te) := tcte in let f' := RTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - do u <- verify_function dupmap f f'; + (*do u <- verify_function dupmap f f';*) OK f'. Definition transf_fundef (f: fundef) : res RTL.fundef := diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 42c20942..6d8c3d87 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -4,33 +4,19 @@ open RTL open Camlcoq open RTLcommonaux open DebugPrint -open BTLaux +open BTLcommonaux +open BTLtypes +open BTLRenumber let get_inumb iinfo = P.of_int iinfo.inumb -let rec get_ib_num = function - | BF (Bgoto s, _) -> s - | Bnop Some iinfo - | Bop (_, _, _, iinfo) - | Bload (_, _, _, _, _, iinfo) - | Bstore (_, _, _, _, iinfo) - | Bcond (_, _, _, _, iinfo) - | BF (_, iinfo) -> - get_inumb iinfo - | Bseq (ib1, _) -> get_ib_num ib1 - | Bnop None -> failwith "get_ib_num: Bnop None found" +let get_ib_num ib = P.of_int (get_inumb_or_next ib) -let translate_function code entry = +let translate_function btl entry = let rtl_code = ref PTree.empty in - let code = - PTree.map - (fun n ibf -> - ibf.binfo.visited <- false; - ibf) - code - in + let btl = reset_visited_ibf btl in let rec build_rtl_tree e = - let ibf = get_some @@ PTree.get e code in + let ibf = get_some @@ PTree.get e btl in if ibf.binfo.visited then () else ( ibf.binfo.visited <- true; @@ -38,6 +24,7 @@ let translate_function code entry = let next_nodes = ref [] in let rec translate_btl_block ib k = let rtli = + (* TODO gourdinl remove assertions *) match ib with | Bcond (cond, lr, BF (Bgoto s1, _), BF (Bgoto s2, _), iinfo) -> next_nodes := s1 :: s2 :: !next_nodes; @@ -48,15 +35,18 @@ let translate_function code entry = Some ( Icond (cond, lr, s1, get_ib_num (get_some k), iinfo.pcond), get_inumb iinfo ) - (* TODO gourdinl remove dynamic check *) | Bcond (_, _, _, _, _) -> - failwith "translate_function: invalid Bcond" + failwith "translate_function: unsupported Bcond" | Bseq (ib1, ib2) -> translate_btl_block ib1 (Some ib2); translate_btl_block ib2 None; None - | Bnop Some iinfo -> Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) - | Bnop None -> failwith "translate_function: invalid Bnop" + | Bnop (Some iinfo) -> + Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) + | Bnop None -> + failwith + "translate_function: Bnop None can only be in the right child \ + of Bcond" | Bop (op, lr, rd, iinfo) -> Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo) | Bload (trap, chk, addr, lr, rd, iinfo) -> @@ -76,7 +66,7 @@ let translate_function code entry = next_nodes := s :: !next_nodes; Some (Ibuiltin (ef, lr, rd, s), get_inumb iinfo) | BF (Bjumptable (arg, tbl), iinfo) -> - next_nodes := !next_nodes @ tbl; + next_nodes := tbl @ !next_nodes; Some (Ijumptable (arg, tbl), get_inumb iinfo) | BF (Breturn oreg, iinfo) -> Some (Ireturn oreg, get_inumb iinfo) | BF (Bgoto s, iinfo) -> @@ -96,13 +86,14 @@ let translate_function code entry = let btl2rtl (f : BTL.coq_function) = (*debug_flag := true;*) - let code = f.fn_code in + let btl = f.fn_code in let entry = f.fn_entrypoint in - let rtl = translate_function code entry in - let dm = PTree.map (fun n i -> n) code in + let ordered_btl, new_entry = renumber btl entry in + let rtl = translate_function ordered_btl new_entry in + let dm = PTree.map (fun n i -> n) btl in debug "Entry %d\n" (p2i entry); debug "RTL Code: "; - (*print_code rtl;*) + print_code rtl; debug "Dupmap:\n"; print_ptree print_pint dm; debug "\n"; diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 2e8d960c..7b62a844 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -58,11 +58,11 @@ Qed. Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. -Proof. +Proof. Admitted. (* unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. -Qed. + Qed.*) Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. diff --git a/scheduling/BTLtypes.ml b/scheduling/BTLtypes.ml new file mode 100644 index 00000000..3972fd6b --- /dev/null +++ b/scheduling/BTLtypes.ml @@ -0,0 +1,7 @@ +type inst_info = { + mutable inumb : int; + mutable pcond : bool option; + mutable visited : bool; +} + +type block_info = { mutable bnumb : int; mutable visited : bool } diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 502565c2..52178064 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -4,8 +4,8 @@ open Datatypes open Maps open BTL open PrintAST -open BTLaux open DebugPrint +open BTLtypes (* Printing of BTL code *) @@ -26,39 +26,44 @@ let print_pref pp pref = fprintf pp "%s" pref let rec print_iblock pp is_rec pref ib = match ib with - | Bnop _ -> + | Bnop None -> print_pref pp pref; - fprintf pp "Bnop\n" - | Bop (op, args, res, _) -> + fprintf pp "??: Bnop None\n" + | Bnop (Some iinfo) -> print_pref pp pref; - fprintf pp "Bop: %a = %a\n" reg res + fprintf pp "%d: Bnop\n" iinfo.inumb + | Bop (op, args, res, iinfo) -> + print_pref pp pref; + fprintf pp "%d: Bop: %a = %a\n" iinfo.inumb reg res (PrintOp.print_operation reg) (op, args) - | Bload (trap, chunk, addr, args, dst, _) -> + | Bload (trap, chunk, addr, args, dst, iinfo) -> print_pref pp pref; - fprintf pp "Bload: %a = %s[%a]%a\n" reg dst (name_of_chunk chunk) + fprintf pp "%d: Bload: %a = %s[%a]%a\n" iinfo.inumb reg dst + (name_of_chunk chunk) (PrintOp.print_addressing reg) (addr, args) print_trapping_mode trap - | Bstore (chunk, addr, args, src, _) -> + | Bstore (chunk, addr, args, src, iinfo) -> print_pref pp pref; - fprintf pp "Bstore: %s[%a] = %a\n" (name_of_chunk chunk) + fprintf pp "%d: Bstore: %s[%a] = %a\n" iinfo.inumb (name_of_chunk chunk) (PrintOp.print_addressing reg) (addr, args) reg src - | BF (Bcall (sg, fn, args, res, s), _) -> + | BF (Bcall (sg, fn, args, res, s), iinfo) -> print_pref pp pref; - fprintf pp "Bcall: %a = %a(%a)\n" reg res ros fn regs args; + fprintf pp "%d: Bcall: %a = %a(%a)\n" iinfo.inumb reg res ros fn regs args; print_succ pp s - | BF (Btailcall (sg, fn, args), _) -> + | BF (Btailcall (sg, fn, args), iinfo) -> print_pref pp pref; - fprintf pp "Btailcall: %a(%a)\n" ros fn regs args - | BF (Bbuiltin (ef, args, res, s), _) -> + fprintf pp "%d: Btailcall: %a(%a)\n" iinfo.inumb ros fn regs args + | BF (Bbuiltin (ef, args, res, s), iinfo) -> print_pref pp pref; - fprintf pp "Bbuiltin: %a = %s(%a)\n" (print_builtin_res reg) res - (name_of_external ef) (print_builtin_args reg) args; + fprintf pp "%d: Bbuiltin: %a = %s(%a)\n" iinfo.inumb + (print_builtin_res reg) res (name_of_external ef) + (print_builtin_args reg) args; print_succ pp s | Bcond (cond, args, ib1, ib2, iinfo) -> print_pref pp pref; - fprintf pp "Bcond: (%a) (prediction: %s)\n" + fprintf pp "%d: Bcond: (%a) (prediction: %s)\n" iinfo.inumb (PrintOp.print_condition reg) (cond, args) (match iinfo.pcond with @@ -72,39 +77,41 @@ let rec print_iblock pp is_rec pref ib = fprintf pp "%sifnot = [\n" pref; if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n"; fprintf pp "%s]\n" pref - | BF (Bjumptable (arg, tbl), _) -> + | BF (Bjumptable (arg, tbl), iinfo) -> let tbl = Array.of_list tbl in print_pref pp pref; - fprintf pp "Bjumptable: (%a)\n" reg arg; + fprintf pp "%d: Bjumptable: (%a)\n" iinfo.inumb reg arg; for i = 0 to Array.length tbl - 1 do fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i)) done - | BF (Breturn None, _) -> + | BF (Breturn None, iinfo) -> print_pref pp pref; - fprintf pp "Breturn\n" - | BF (Breturn (Some arg), _) -> + fprintf pp "%d: Breturn\n" iinfo.inumb + | BF (Breturn (Some arg), iinfo) -> print_pref pp pref; - fprintf pp "Breturn: %a\n" reg arg - | BF (Bgoto s, _) -> + fprintf pp "%d: Breturn: %a\n" iinfo.inumb reg arg + | BF (Bgoto s, iinfo) -> print_pref pp pref; - fprintf pp "Bgoto: %d\n" (P.to_int s) + fprintf pp "%d: Bgoto: %d\n" iinfo.inumb (P.to_int s) | Bseq (ib1, ib2) -> if is_rec then ( print_iblock pp is_rec pref ib1; print_iblock pp is_rec pref ib2) else fprintf pp "Bseq...\n" -let print_btl_inst pp ib = print_iblock pp false " " ib +let print_btl_inst pp ib = + if !debug_flag then print_iblock pp false " " ib else () -let print_btl_code pp btl is_rec = +let print_btl_code pp btl = if !debug_flag then ( fprintf pp "\n"; List.iter (fun (n, ibf) -> - fprintf pp "[BTL Liveness]\n"; + fprintf pp "[BTL Liveness] "; print_regset ibf.input_regs; - fprintf pp "[BTL block %d]\n" (P.to_int n); - print_iblock pp is_rec " " ibf.entry; + fprintf pp "\n"; + fprintf pp "[BTL block %d with inumb %d]\n" (P.to_int n) ibf.binfo.bnumb; + print_iblock pp true " " ibf.entry; fprintf pp "\n") (PTree.elements btl); fprintf pp "\n") diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index b3585157..2ac02597 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -15,7 +15,7 @@ Definition transf_function (f: RTL.function) : res BTL.function := let (tcte, dupmap) := rtl2btl f in let (tc, te) := tcte in let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) tc te in - do u <- verify_function dupmap f' f; + (*do u <- verify_function dupmap f' f;*) OK f'. Definition transf_fundef (f: RTL.fundef) : res fundef := diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index e932d766..056fe213 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -4,23 +4,16 @@ open BTL open PrintBTL open RTLcommonaux open DebugPrint -open BTLaux +open BTLtypes +open BTLcommonaux open BTLScheduleraux -let undef_node = -1 - -let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond } - -let def_iinfo = { inumb = undef_node; pcond = None } - -let mk_binfo _bnumb = { bnumb = _bnumb; visited = false } - let encaps_final inst osucc = match inst with | BF _ | Bcond _ -> inst - | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo)) + | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo ())) -let translate_inst (iinfo: BTL.inst_info) inst is_final = +let translate_inst (iinfo : BTL.inst_info) inst is_final = let osucc = ref None in let btli = match inst with @@ -44,8 +37,8 @@ let translate_inst (iinfo: BTL.inst_info) inst is_final = Bcond ( cond, lr, - BF (Bgoto ifso, def_iinfo), - BF (Bgoto ifnot, def_iinfo), + BF (Bgoto ifso, def_iinfo ()), + BF (Bgoto ifnot, def_iinfo ()), iinfo ) | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl), iinfo) | Ireturn oreg -> BF (Breturn oreg, iinfo) @@ -67,9 +60,7 @@ let translate_function code entry join_points liveness = let succ = match psucc with | Some ps -> - if get_some @@ PTree.get ps join_points then ( - None) - else Some ps + if get_some @@ PTree.get ps join_points then None else Some ps | None -> None in match succ with @@ -80,13 +71,10 @@ let translate_function code entry join_points liveness = assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; iinfo.pcond <- info; - Bseq ( - Bcond - ( cond, - lr, - BF (Bgoto ifso, def_iinfo), - Bnop None, - iinfo ), build_btl_block s) + Bseq + ( Bcond + (cond, lr, BF (Bgoto ifso, def_iinfo ()), Bnop None, iinfo), + build_btl_block s ) | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s)) | None -> debug "BLOCK END.\n"; @@ -113,17 +101,19 @@ let rtl2btl (f : RTL.coq_function) = let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in (* TODO gourdinl move elsewhere *) - btl_scheduler btl entry; + let btl' = btl_scheduler btl in debug "Entry %d\n" (p2i entry); + (*debug_flag := true;*) debug "RTL Code: "; print_code code; - (*debug_flag := true;*) - debug "BTL Code: "; - print_btl_code stderr btl true; + debug "BTL Code:\n"; + print_btl_code stderr btl; + debug "Scheduled BTL Code:\n"; + print_btl_code stderr btl'; (*debug_flag := false;*) debug "Dupmap:\n"; print_ptree print_pint dm; debug "Join points: "; print_true_nodes join_points; debug "\n"; - ((btl, entry), dm) + ((btl', entry), dm) diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index cd6c4dae..5a5b3a86 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -40,11 +40,11 @@ Qed. Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. -Proof. +Proof. Admitted. (* unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. -Qed. + Qed.*) Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. -- cgit From 43ab0b948ac379e55bbe334a0a541c1680437fbf Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 10:40:54 +0200 Subject: most of the proof BTL.fsem -> BTL.cfgsem. --- scheduling/BTL.v | 308 +++++++++++++++++++++++++++++++++++++++------ scheduling/BTL_SEtheory.v | 2 +- scheduling/BTLtoRTLproof.v | 12 +- scheduling/RTLtoBTLproof.v | 4 +- 4 files changed, 280 insertions(+), 46 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index bf19f4ac..389d8d93 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -17,6 +17,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad. +Require Import Lia. Import ListNotations. @@ -380,22 +381,20 @@ Inductive final_state: state -> int -> Prop := | final_state_intro: forall r m, final_state (Returnstate nil (Vint r) m) r. -(** The basic small-step semantics for a BTL program. *) +(** The "raw" CFG small-step semantics for a BTL program (e.g. close to RTL semantics) *) (* tid = transfer_identity *) Definition tid (_:function) (_:list exit) (_: option reg) (rs:regset) := rs. -Definition semantics (p: program) := +Definition cfgsem (p: program) := Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). -(** The "functional" small-step semantics for a BTL program. +(** The full "functional" small-step semantics for a BTL program. at each exit, we only transfer register in "input_regs" (i.e. "alive" registers). *) Definition transfer_regs (inputs: list reg) (rs: regset): regset := init_regs (rs##inputs) inputs. -(* TODO: LEMMA BELOW ON transfer_regs USEFUL ? *) - Lemma transfer_regs_inputs inputs rs r: List.In r inputs -> (transfer_regs inputs rs)#r = rs#r. Proof. @@ -414,18 +413,20 @@ Proof. - rewrite Regmap.gso; auto. Qed. -Fixpoint inputs_exitrec (f:function) (tbl: list exit): Regset.t := +Fixpoint union_inputs (f:function) (tbl: list exit): Regset.t := match tbl with | nil => Regset.empty - | pc::l => let rs:= inputs_exitrec f l in + | pc::l => let rs:= union_inputs f l in match f.(fn_code)!pc with | None => rs | Some ib => Regset.union rs ib.(input_regs) end end. -Lemma inputs_exitrec_In f tbl r: - Regset.In r (inputs_exitrec f tbl) -> (exists pc, List.In pc tbl /\ exists ib, f.(fn_code)!pc = Some ib /\ Regset.In r ib.(input_regs)). +(* TODO: lemma not yet used + +Lemma union_inputs_In f tbl r: + Regset.In r (union_inputs f tbl) -> (exists pc, List.In pc tbl /\ exists ib, f.(fn_code)!pc = Some ib /\ Regset.In r ib.(input_regs)). Proof. induction tbl as [|pc l]; simpl; intros. - exploit Regset.empty_1; eauto. intuition. @@ -437,27 +438,33 @@ Proof. + destruct IHl as (pc' & H1 & H2); eauto. Qed. -Lemma inputs_exitrec_notIn f tbl r: +Lemma union_inputs_notIn f tbl r: (forall pc ib, List.In pc tbl -> f.(fn_code)!pc = Some ib -> ~Regset.In r ib.(input_regs)) - -> ~Regset.In r (inputs_exitrec f tbl). + -> ~Regset.In r (union_inputs f tbl). Proof. induction tbl as [|pc l]; simpl; intuition eauto. - exploit Regset.empty_1; eauto. - destruct ((fn_code f) ! pc) eqn:ATpc; intuition eauto. exploit Regset.union_1; intuition eauto. Qed. +*) -Definition inputs_exit (f:function) (tbl: list exit) (or: option reg): Regset.t := - let rs := inputs_exitrec f tbl in +(* This function computes the union of the inputs associated to the exits + minus the optional register in [or] (typically the result of a call or a builtin). + i.e. similar to [pre_output_regs] in RTLpath. +*) +Definition pre_inputs (f:function) (tbl: list exit) (or: option reg): Regset.t := + let rs := union_inputs f tbl in match or with | Some r => Regset.remove r rs | None => rs - end. + end + . -Definition tr_inputs (f:function) (tbl: list exit) (or:option reg): regset -> regset - := transfer_regs (Regset.elements (inputs_exit f tbl or)). +(* TODO: lemma pre_inputs_In + pre_inputs_notIn ? *) -(* TODO: BELOW, LEMMA on tr_inputs USEFUL ? *) +Definition tr_inputs (f:function) (tbl: list exit) (or:option reg): regset -> regset + := transfer_regs (Regset.elements (pre_inputs f tbl or)). (* TODO: move this elsewhere *) @@ -469,8 +476,8 @@ Proof. - induction l; simpl; intuition. Qed. -Lemma tr_inputs_exit f tbl or rs r: - Regset.In r (inputs_exit f tbl or) -> (tr_inputs f tbl or rs)#r = rs#r. +Lemma tr_pre_inputs f tbl or rs r: + Regset.In r (pre_inputs f tbl or) -> (tr_inputs f tbl or rs)#r = rs#r. Proof. intros; eapply transfer_regs_inputs. rewrite <- SetoidList_InA_eq_equiv. @@ -478,17 +485,17 @@ Proof. Qed. Lemma tr_inputs_dead f tbl or rs r: - ~(Regset.In r (inputs_exit f tbl or)) -> (tr_inputs f tbl or rs)#r = Vundef. + ~(Regset.In r (pre_inputs f tbl or)) -> (tr_inputs f tbl or rs)#r = Vundef. Proof. intros X; eapply transfer_regs_dead; intuition eauto. eapply X. eapply Regset.elements_2. rewrite -> SetoidList_InA_eq_equiv; eauto. Qed. -Local Hint Resolve tr_inputs_exit Regset.mem_2 Regset.mem_1: core. +Local Hint Resolve tr_pre_inputs Regset.mem_2 Regset.mem_1: core. Lemma tr_inputs_get f tbl or rs r: - (tr_inputs f tbl or rs)#r = if Regset.mem r (inputs_exit f tbl or) then rs#r else Vundef. + (tr_inputs f tbl or rs)#r = if Regset.mem r (pre_inputs f tbl or) then rs#r else Vundef. Proof. autodestruct; eauto. intros; apply tr_inputs_dead; eauto. @@ -496,13 +503,15 @@ Proof. congruence. Qed. +(* TODO: not yet used Lemma tr_inputs_ext f tbl or rs1 rs2: - (forall r, Regset.In r (inputs_exit f tbl or) -> rs1#r = rs2#r) -> + (forall r, Regset.In r (pre_inputs f tbl or) -> rs1#r = rs2#r) -> (forall r, (tr_inputs f tbl or rs1)#r = (tr_inputs f tbl or rs2)#r). Proof. intros EQ r. rewrite !tr_inputs_get. autodestruct; auto. Qed. +*) Definition fsem (p: program) := Semantics (step tr_inputs) (initial_state p) final_state (Genv.globalenv p). @@ -524,7 +533,7 @@ Definition tr_regs: function -> final -> regset -> regset := (* TODO: NOT USEFUL ? Definition liveout: function -> final -> Regset.t := - poly_tr inputs_exit. + poly_tr pre_inputs. Lemma tr_regs_liveout_equiv f fi : tr_regs f fi = transfer_regs (Regset.elements (liveout f fi)). Proof. @@ -535,30 +544,34 @@ Local Hint Resolve tr_inputs_get: core. Lemma tr_regs_get f fi rs r: (tr_regs f fi rs)#r = if Regset.mem r (liveout f fi) then rs#r else Vundef. Proof. - Local Opaque inputs_exit. + Local Opaque pre_inputs. destruct fi; simpl; auto. Qed. *) -(* * Reasoning modulo extensionality of [regset] into BTL semantics *) +(* * Comparing BTL semantics modulo extensionality of [regset]. + +NB: This is at least used in BTL_SEtheory (and probably in the future BTL_SchedulerProof). + +*) Inductive equiv_stackframe: stackframe -> stackframe -> Prop := - | equiv_stackframe_intro res f sp pc rs1 rs2 - (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): - equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) + | equiv_stackframe_intro res f sp pc (rs1 rs2: regset) + (EQUIV: forall r, rs1#r = rs2#r) + :equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) . Inductive equiv_state: state -> state -> Prop := - | State_equiv stack f sp pc rs1 m rs2 - (EQUIV: forall r, rs1#r = rs2#r): - equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) - | Call_equiv stk stk' f args m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Callstate stk f args m) (Callstate stk' f args m) - | Return_equiv stk stk' v m - (STACKS: list_forall2 equiv_stackframe stk stk'): - equiv_state (Returnstate stk v m) (Returnstate stk' v m) + | State_equiv stk1 stk2 f sp pc rs1 m rs2 + (STACKS: list_forall2 equiv_stackframe stk1 stk2) + (EQUIV: forall r, rs1#r = rs2#r) + :equiv_state (State stk1 f sp pc rs1 m) (State stk2 f sp pc rs2 m) + | Call_equiv stk1 stk2 f args m + (STACKS: list_forall2 equiv_stackframe stk1 stk2) + :equiv_state (Callstate stk1 f args m) (Callstate stk2 f args m) + | Return_equiv stk1 stk2 v m + (STACKS: list_forall2 equiv_stackframe stk1 stk2) + :equiv_state (Returnstate stk1 v m) (Returnstate stk2 v m) . - Local Hint Constructors equiv_stackframe equiv_state: core. Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf. @@ -571,10 +584,10 @@ Proof. Local Hint Resolve equiv_stackframe_refl: core. induction stk; simpl; constructor; auto. Qed. +Local Hint Resolve equiv_stack_refl: core. Lemma equiv_state_refl s: equiv_state s s. Proof. - Local Hint Resolve equiv_stack_refl: core. induction s; simpl; constructor; auto. Qed. @@ -584,13 +597,13 @@ Proof. destruct 1; intros EQ; inv EQ; try econstructor; eauto. intros; eapply eq_trans; eauto. Qed. +Local Hint Resolve equiv_stackframe_trans: core. Lemma equiv_stack_trans stk1 stk2: list_forall2 equiv_stackframe stk1 stk2 -> forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> list_forall2 equiv_stackframe stk1 stk3. Proof. - Local Hint Resolve equiv_stackframe_trans: core. induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. Qed. @@ -611,7 +624,220 @@ Proof. - repeat (rewrite Regmap.gso); auto. Qed. +(* * Simulation of BTL fsem by BTL cfgsem: a lock-step less-def simulation. +We start by extending the previous [equiv_*] stuff for [Val.lessdef]: we need to also compare memories +for Mem.extends, because Vundef values are propagated in memory through stores, allocation, etc. + +*) + +Inductive lessdef_stackframe: stackframe -> stackframe -> Prop := + | lessdef_stackframe_intro res f sp pc rs1 rs2 + (REGS: forall r, Val.lessdef rs1#r rs2#r) + :lessdef_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2) + . + +Inductive lessdef_state: state -> state -> Prop := + | State_lessdef stk1 stk2 f sp pc rs1 m1 rs2 m2 + (STACKS: list_forall2 lessdef_stackframe stk1 stk2) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS: Mem.extends m1 m2) + :lessdef_state (State stk1 f sp pc rs1 m1) (State stk2 f sp pc rs2 m2) + | Call_lessdef stk1 stk2 f args1 args2 m1 m2 + (STACKS: list_forall2 lessdef_stackframe stk1 stk2) + (ARGS: Val.lessdef_list args1 args2) + (MEMS: Mem.extends m1 m2) + :lessdef_state (Callstate stk1 f args1 m1) (Callstate stk2 f args2 m2) + | Return_lessdef stk1 stk2 v1 v2 m1 m2 + (STACKS: list_forall2 lessdef_stackframe stk1 stk2) + (REGS: Val.lessdef v1 v2) + (MEMS: Mem.extends m1 m2) + :lessdef_state (Returnstate stk1 v1 m1) (Returnstate stk2 v2 m2) + . +Local Hint Constructors lessdef_stackframe lessdef_state: core. + +Lemma lessdef_stackframe_refl stf: lessdef_stackframe stf stf. +Proof. + destruct stf; eauto. +Qed. + +Local Hint Resolve lessdef_stackframe_refl: core. +Lemma lessdef_stack_refl stk: list_forall2 lessdef_stackframe stk stk. +Proof. + induction stk; simpl; constructor; auto. +Qed. + +Local Hint Resolve lessdef_stack_refl: core. + +Lemma lessdef_list_refl args: Val.lessdef_list args args. +Proof. + induction args; simpl; auto. +Qed. + +Local Hint Resolve lessdef_list_refl Mem.extends_refl: core. + +Lemma lessdef_state_refl s: lessdef_state s s. +Proof. + induction s; simpl; constructor; auto. +Qed. + +Local Hint Resolve Val.lessdef_trans: core. +Lemma lessdef_stackframe_trans stf1 stf2 stf3: + lessdef_stackframe stf1 stf2 -> lessdef_stackframe stf2 stf3 -> lessdef_stackframe stf1 stf3. +Proof. + destruct 1. intros LD; inv LD; try econstructor; eauto. +Qed. + +Local Hint Resolve lessdef_stackframe_trans: core. +Lemma lessdef_stack_trans stk1 stk2: + list_forall2 lessdef_stackframe stk1 stk2 -> + forall stk3, list_forall2 lessdef_stackframe stk2 stk3 -> + list_forall2 lessdef_stackframe stk1 stk3. +Proof. + induction 1; intros stk3 LD; inv LD; econstructor; eauto. +Qed. + +Local Hint Resolve lessdef_stack_trans Mem.extends_extends_compose Val.lessdef_list_trans: core. +Lemma lessdef_state_trans s1 s2 s3: lessdef_state s1 s2 -> lessdef_state s2 s3 -> lessdef_state s1 s3. +Proof. + destruct 1; intros LD; inv LD; econstructor; eauto. +Qed. + +Lemma init_regs_lessdef_preserv args1 args2 + (ARGS : Val.lessdef_list args1 args2) + : forall rl r, Val.lessdef (init_regs args1 rl)#r (init_regs args2 rl)#r. +Proof. + induction ARGS; simpl; auto. + intros rl r1; destruct rl; simpl in *; auto. + eapply set_reg_lessdef; eauto. + intros r2; eauto. +Qed. +Local Hint Resolve init_regs_lessdef_preserv: core. + +Lemma tr_inputs_lessdef f rs1 rs2 tbl or r + (REGS: forall r, Val.lessdef rs1#r rs2#r) + :Val.lessdef (tr_inputs f tbl or rs1)#r (tid f tbl or rs2)#r. +Proof. + unfold tid; rewrite tr_inputs_get. + autodestruct. +Qed. +Local Hint Resolve tr_inputs_lessdef: core. + +Lemma find_function_lessdef ge ros rs1 rs2 fd + (FIND: find_function ge ros rs1 = Some fd) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + : find_function ge ros rs2 = Some fd. +Proof. + destruct ros; simpl in *; auto. + exploit Genv.find_funct_inv; eauto. + intros (b & EQ). + destruct (REGS r); try congruence. +Qed. +Local Hint Resolve find_function_lessdef regs_lessdef_regs: core. + + +Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 ib rs1' m1' of + (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of) + rs2 m2 + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS: Mem.extends m1 m2) + : exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) + /\ (forall r, Val.lessdef rs1'#r rs2'#r) + /\ (Mem.extends m1' m2'). +Proof. + induction ISTEP; simpl; try_simplify_someHyps. +Admitted. + +Local Hint Constructors final_step list_forall2: core. +Local Hint Unfold regs_lessdef: core. + +Lemma fsem2cfgsem_finalstep_simu ge stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 + (FSTEP: final_step tr_inputs ge stk1 f sp rs1 m1 fin t s1) + (STACKS: list_forall2 lessdef_stackframe stk1 stk2) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS: Mem.extends m1 m2) + : exists s2, final_step tid ge stk2 f sp rs2 m2 fin t s2 + /\ lessdef_state s1 s2. +Proof. + destruct FSTEP; try (eexists; split; simpl; econstructor; eauto; fail). + - (* return *) + exploit Mem.free_parallel_extends; eauto. + intros (m2' & FREE & MEMS2). + eexists; split; simpl; econstructor; eauto. + destruct or; simpl; auto. + - (* tailcall *) + exploit Mem.free_parallel_extends; eauto. + intros (m2' & FREE & MEMS2). + eexists; split; simpl; econstructor; eauto. + - (* builtin *) + exploit (eval_builtin_args_lessdef (ge:=ge) (e1:=(fun r => rs1 # r)) (fun r => rs2 # r)); eauto. + intros (vl2' & BARGS & VARGS). + exploit external_call_mem_extends; eauto. + intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). + eexists; split; simpl; econstructor; eauto. + intros; apply set_res_lessdef; eauto. + - (* jumptable *) + eexists; split; simpl; econstructor; eauto. + destruct (REGS arg); try congruence. +Qed. + +Lemma fsem2cfgsem_ibstep_simu ge stk1 stk2 f sp rs1 m1 ib t s1 rs2 m2 + (STEP: iblock_step tr_inputs ge stk1 f sp rs1 m1 ib t s1) + (STACKS : list_forall2 lessdef_stackframe stk1 stk2) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS : Mem.extends m1 m2) + : exists s2, iblock_step tid ge stk2 f sp rs2 m2 ib t s2 + /\ lessdef_state s1 s2. +Proof. + destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP). + exploit fsem2cfgsem_ibistep_simu; eauto. + intros (rs2' & m2' & ISTEP2 & REGS2 & MEMS2). + rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP REGS MEMS. + exploit fsem2cfgsem_finalstep_simu; eauto. + intros (s2 & FSTEP2 & LESSDEF). clear FSTEP. + unfold iblock_step; eexists; split; eauto. +Qed. + +Local Hint Constructors step: core. + +Lemma fsem2cfgsem_step_simu ge s1 t s1' s2 + (STEP: step tr_inputs ge s1 t s1') + (REGS: lessdef_state s1 s2) + :exists s2' : state, + step tid ge s2 t s2' /\ + lessdef_state s1' s2'. +Proof. + destruct STEP; inv REGS. + - (* iblock *) + intros; exploit fsem2cfgsem_ibstep_simu; eauto. clear STEP. + intros (s2 & STEP2 & REGS2). + eexists; split; eauto. + - (* internal call *) + exploit (Mem.alloc_extends m m2 0 (fn_stacksize f) stk m' 0 (fn_stacksize f)); eauto. + 1-2: lia. + intros (m2' & ALLOC2 & MEMS2). clear ALLOC MEMS. + eexists; split; econstructor; eauto. + - (* external call *) + exploit external_call_mem_extends; eauto. + intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). clear EXTCALL MEMS. + eexists; split; econstructor; eauto. + - (* return *) + inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst. + inv STF2. + eexists; split; econstructor; eauto. + intros; eapply set_reg_lessdef; eauto. +Qed. + +Theorem fsem2cfgsem prog: + forward_simulation (fsem prog) (cfgsem prog). +Proof. + eapply forward_simulation_step with lessdef_state; simpl; auto. + - destruct 1; intros; eexists; intuition eauto. econstructor; eauto. + - intros s1 s2 r REGS FINAL; destruct FINAL. + inv REGS; inv STACKS; inv REGS0. + econstructor; eauto. + - intros; eapply fsem2cfgsem_step_simu; eauto. +Qed. (** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 79db99c8..c9f50d66 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -633,7 +633,7 @@ Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := | r1::l => fun r => if Pos.eq_dec r r1 then sreg r1 else transfer_sreg l sreg r end. -Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (inputs_exit f tbl or)). +Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (pre_inputs f tbl or)). Lemma str_inputs_correct ctx sis rs tbl or r: (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) -> diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index bbb984c4..14e986aa 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -390,8 +390,8 @@ Proof. + inv H1. econstructor; eauto. Qed. -Theorem transf_program_correct: - forward_simulation (semantics prog) (RTL.semantics tprog). +Theorem transf_program_correct_cfg: + forward_simulation (BTL.cfgsem prog) (RTL.semantics tprog). Proof. eapply forward_simulation_plus with match_states. - eapply senv_preserved. @@ -400,4 +400,12 @@ Proof. - eapply plus_simulation. Qed. +Theorem transf_program_correct: + forward_simulation (BTL.fsem prog) (RTL.semantics tprog). +Proof. + eapply compose_forward_simulations. + - eapply fsem2cfgsem. + - eapply transf_program_correct_cfg. +Qed. + End RTL_SIMULATES_BTL. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 60edea49..feaac26f 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -727,9 +727,9 @@ Qed. Local Hint Resolve plus_one star_refl: core. Theorem transf_program_correct: - forward_simulation (RTL.semantics prog) (BTL.semantics tprog). + forward_simulation (RTL.semantics prog) (BTL.cfgsem tprog). Proof. - eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=semantics tprog) (ltof _ omeasure) match_states). + eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=cfgsem tprog) (ltof _ omeasure) match_states). constructor 1; simpl. - apply well_founded_ltof. - eapply transf_initial_states. -- cgit From 05b24fdb11414100b9b04867e6e2d3a1a9054162 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 28 May 2021 11:44:11 +0200 Subject: Improvements in scheduling and renumbering BTL code --- scheduling/BTLRenumber.ml | 16 +++--- scheduling/BTLScheduleraux.ml | 7 ++- scheduling/BTLcommonaux.ml | 14 ----- scheduling/BTLtoRTLaux.ml | 117 ++++++++++++++++++------------------------ scheduling/RTLtoBTLaux.ml | 5 +- 5 files changed, 65 insertions(+), 94 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLRenumber.ml b/scheduling/BTLRenumber.ml index 36f3bcf5..58d4f7ac 100644 --- a/scheduling/BTLRenumber.ml +++ b/scheduling/BTLRenumber.ml @@ -3,6 +3,8 @@ open BTL open RTLcommonaux open BTLcommonaux open BTLtypes +open DebugPrint +open PrintBTL let recompute_inumbs btl entry = let btl = reset_visited_ib (reset_visited_ibf btl) in @@ -53,13 +55,6 @@ let recompute_inumbs btl entry = walk succ None; iinfo.inumb <- ipos () | Bseq (ib1, ib2) -> walk ib1 (Some ib2) - | Bcond (_, _, BF (Bgoto s1, iinfoL), BF (Bgoto s2, iinfoR), iinfo) -> - iinfoL.visited <- true; - iinfoR.visited <- true; - let ib1 = get_some @@ PTree.get s1 btl in - let ib2 = get_some @@ PTree.get s2 btl in - walk_smallest_child (p2i s1) (p2i s2) ib1.entry ib2.entry; - iinfo.inumb <- ipos () | Bcond (_, _, BF (Bgoto s1, iinfoL), Bnop None, iinfoF) -> iinfoL.visited <- true; let ib1 = get_some @@ PTree.get s1 btl in @@ -104,8 +99,13 @@ let regenerate_btl_tree btl entry = PTree.set n_pos ibf ord_btl) btl PTree.empty in + debug "Renumbered BTL with new_entry=%d:\n" (p2i !new_entry); + print_btl_code stderr ord_btl; (ord_btl, !new_entry) let renumber btl entry = + (*debug_flag := true;*) let btl' = recompute_inumbs btl entry in - regenerate_btl_tree btl' entry + let ord_btl, new_entry = regenerate_btl_tree btl' entry in + (*debug_flag := false;*) + (ord_btl, new_entry) diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 4b5ebb32..ad0c307d 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -80,7 +80,6 @@ let build_constraints_and_resources (opweights : opweights) insts btl = in Array.iteri (fun i inst -> - (* TODO gourdinl liveins for Bcond *) match inst with | Bnop _ -> let rs = Array.map (fun _ -> 0) opweights.pipelined_resource_bounds in @@ -203,13 +202,13 @@ let flatten_blk_basics ibf = | BF (_, _) -> last := Some ib; [] - | Bseq (ib1, ib2) -> traverse_blk ib1 @ traverse_blk ib2 - | Bcond (_, _, _, _, iinfo) -> ( + | Bseq ((Bcond (_, _, _, _, iinfo) as ib1), ib2) -> ( match iinfo.pcond with - | Some _ -> [ ib ] + | Some _ -> [ ib1 ] @ traverse_blk ib2 | None -> last := Some ib; []) + | Bseq (ib1, ib2) -> traverse_blk ib1 @ traverse_blk ib2 | _ -> [ ib ] in let ibl = traverse_blk ib in diff --git a/scheduling/BTLcommonaux.ml b/scheduling/BTLcommonaux.ml index dabf760a..4605d613 100644 --- a/scheduling/BTLcommonaux.ml +++ b/scheduling/BTLcommonaux.ml @@ -68,17 +68,3 @@ let rec get_inumb_or_next = function iinfo.inumb | Bseq (ib1, _) -> get_inumb_or_next ib1 | _ -> failwith "get_inumb_or_next: Bnop None" - -let rec set_next_inumb btl pos = function - | BF (Bgoto s, _) -> - let ib' = (get_some @@ PTree.get s btl).entry in - set_next_inumb btl pos ib' - | BF (_, iinfo) - | Bnop (Some iinfo) - | Bop (_, _, _, iinfo) - | Bload (_, _, _, _, _, iinfo) - | Bstore (_, _, _, _, iinfo) - | Bcond (_, _, _, _, iinfo) -> - iinfo.inumb <- pos - | Bseq (ib1, _) -> set_next_inumb btl pos ib1 - | _ -> failwith "get_inumb_or_next: Bnop None" diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 6d8c3d87..8e728bd2 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -4,6 +4,7 @@ open RTL open Camlcoq open RTLcommonaux open DebugPrint +open PrintBTL open BTLcommonaux open BTLtypes open BTLRenumber @@ -14,74 +15,58 @@ let get_ib_num ib = P.of_int (get_inumb_or_next ib) let translate_function btl entry = let rtl_code = ref PTree.empty in - let btl = reset_visited_ibf btl in - let rec build_rtl_tree e = - let ibf = get_some @@ PTree.get e btl in - if ibf.binfo.visited then () - else ( + let rec translate_btl_block ib k = + debug "Entering translate_btl_block...\n"; + print_btl_inst stderr ib; + let rtli = + match ib with + | Bcond (cond, lr, BF (Bgoto s1, _), Bnop None, iinfo) -> + Some + ( Icond (cond, lr, s1, get_ib_num (get_some k), iinfo.pcond), + get_inumb iinfo ) + | Bcond (_, _, _, _, _) -> + failwith "translate_function: unsupported Bcond" + | Bseq (ib1, ib2) -> + translate_btl_block ib1 (Some ib2); + translate_btl_block ib2 None; + None + | Bnop (Some iinfo) -> + Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) + | Bnop None -> + failwith + "translate_function: Bnop None can only be in the right child of \ + Bcond" + | Bop (op, lr, rd, iinfo) -> + Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo) + | Bload (trap, chk, addr, lr, rd, iinfo) -> + Some + ( Iload (trap, chk, addr, lr, rd, get_ib_num (get_some k)), + get_inumb iinfo ) + | Bstore (chk, addr, lr, src, iinfo) -> + Some + ( Istore (chk, addr, lr, src, get_ib_num (get_some k)), + get_inumb iinfo ) + | BF (Bcall (sign, fn, lr, rd, s), iinfo) -> + Some (Icall (sign, fn, lr, rd, s), get_inumb iinfo) + | BF (Btailcall (sign, fn, lr), iinfo) -> + Some (Itailcall (sign, fn, lr), get_inumb iinfo) + | BF (Bbuiltin (ef, lr, rd, s), iinfo) -> + Some (Ibuiltin (ef, lr, rd, s), get_inumb iinfo) + | BF (Bjumptable (arg, tbl), iinfo) -> + Some (Ijumptable (arg, tbl), get_inumb iinfo) + | BF (Breturn oreg, iinfo) -> Some (Ireturn oreg, get_inumb iinfo) + | BF (Bgoto s, iinfo) -> None + in + match rtli with + | Some (inst, inumb) -> rtl_code := PTree.set inumb inst !rtl_code + | None -> () + in + List.iter + (fun (n, ibf) -> ibf.binfo.visited <- true; let ib = ibf.entry in - let next_nodes = ref [] in - let rec translate_btl_block ib k = - let rtli = - (* TODO gourdinl remove assertions *) - match ib with - | Bcond (cond, lr, BF (Bgoto s1, _), BF (Bgoto s2, _), iinfo) -> - next_nodes := s1 :: s2 :: !next_nodes; - Some (Icond (cond, lr, s1, s2, iinfo.pcond), get_inumb iinfo) - | Bcond (cond, lr, BF (Bgoto s1, _), Bnop None, iinfo) -> - assert (iinfo.pcond = Some false); - next_nodes := s1 :: !next_nodes; - Some - ( Icond (cond, lr, s1, get_ib_num (get_some k), iinfo.pcond), - get_inumb iinfo ) - | Bcond (_, _, _, _, _) -> - failwith "translate_function: unsupported Bcond" - | Bseq (ib1, ib2) -> - translate_btl_block ib1 (Some ib2); - translate_btl_block ib2 None; - None - | Bnop (Some iinfo) -> - Some (Inop (get_ib_num (get_some k)), get_inumb iinfo) - | Bnop None -> - failwith - "translate_function: Bnop None can only be in the right child \ - of Bcond" - | Bop (op, lr, rd, iinfo) -> - Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo) - | Bload (trap, chk, addr, lr, rd, iinfo) -> - Some - ( Iload (trap, chk, addr, lr, rd, get_ib_num (get_some k)), - get_inumb iinfo ) - | Bstore (chk, addr, lr, src, iinfo) -> - Some - ( Istore (chk, addr, lr, src, get_ib_num (get_some k)), - get_inumb iinfo ) - | BF (Bcall (sign, fn, lr, rd, s), iinfo) -> - next_nodes := s :: !next_nodes; - Some (Icall (sign, fn, lr, rd, s), get_inumb iinfo) - | BF (Btailcall (sign, fn, lr), iinfo) -> - Some (Itailcall (sign, fn, lr), get_inumb iinfo) - | BF (Bbuiltin (ef, lr, rd, s), iinfo) -> - next_nodes := s :: !next_nodes; - Some (Ibuiltin (ef, lr, rd, s), get_inumb iinfo) - | BF (Bjumptable (arg, tbl), iinfo) -> - next_nodes := tbl @ !next_nodes; - Some (Ijumptable (arg, tbl), get_inumb iinfo) - | BF (Breturn oreg, iinfo) -> Some (Ireturn oreg, get_inumb iinfo) - | BF (Bgoto s, iinfo) -> - next_nodes := s :: !next_nodes; - None - in - match rtli with - | Some (inst, inumb) -> rtl_code := PTree.set inumb inst !rtl_code - | None -> () - in - translate_btl_block ib None; - let succs = !next_nodes in - List.iter build_rtl_tree succs) - in - build_rtl_tree entry; + translate_btl_block ib None) + (PTree.elements btl); !rtl_code let btl2rtl (f : BTL.coq_function) = diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 056fe213..d04326ea 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -10,7 +10,7 @@ open BTLScheduleraux let encaps_final inst osucc = match inst with - | BF _ | Bcond _ -> inst + | BF _ -> inst | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo ())) let translate_inst (iinfo : BTL.inst_info) inst is_final = @@ -33,12 +33,13 @@ let translate_inst (iinfo : BTL.inst_info) inst is_final = | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr), iinfo) | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s), iinfo) | Icond (cond, lr, ifso, ifnot, info) -> + osucc := Some ifnot; iinfo.pcond <- info; Bcond ( cond, lr, BF (Bgoto ifso, def_iinfo ()), - BF (Bgoto ifnot, def_iinfo ()), + Bnop None, iinfo ) | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl), iinfo) | Ireturn oreg -> BF (Breturn oreg, iinfo) -- cgit From 8752cea89f862d92d49183fe10d1918588143ab1 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 28 May 2021 12:38:05 +0200 Subject: fix some merge errors --- scheduling/BTL.v | 4 ++-- scheduling/BTL_SEtheory.v | 2 +- scheduling/RTLtoBTLproof.v | 1 + 3 files changed, 4 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 2be4fcac..2c1dda5b 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -94,7 +94,7 @@ Coercion BF: final >-> Funclass. Record iblock_info := { entry: iblock; - input_regs: Regset.t (* extra liveness information for BTL functional semantics *) + input_regs: Regset.t; (* extra liveness information for BTL functional semantics *) binfo: block_info (* Ghost field used in oracles *) }. @@ -1285,4 +1285,4 @@ Definition is_goto (ib: iblock): bool := match ib with | BF (Bgoto _) _ => true | _ => false - end. \ No newline at end of file + end. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index cea69f55..5b15fe9b 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -731,7 +731,7 @@ pour représenter l'ensemble des chemins. Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := match ib with - | BF fin => Sfinal (tr_sis f fin sis) (sexec_final_sfv fin sis) + | BF fin _ => Sfinal (tr_sis f fin sis) (sexec_final_sfv fin sis) (* basic instructions *) | Bnop _ => k sis | Bop op args res _ => k (sexec_op op args res sis) diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 7d83931b..f2f99ef5 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -554,6 +554,7 @@ Proof. eexists; left; eexists; split; eauto. repeat econstructor; eauto. apply iblock_istep_run_equiv in BTL_RUN; eauto. + econstructor. - (* Others *) exists (Some ib2); right; split. simpl; auto. -- cgit From 25595a7b34b70011dcb77aae277ee1cdb8920c60 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 14:28:56 +0200 Subject: splitting BTL by introducing BTLmatchRTL reduce also copy-paste between BTLtoRTLproof and RTLtoBTLproof sharing is done in BTLmatchRTL --- scheduling/BTL.v | 565 +------------------------------------------ scheduling/BTLmatchRTL.v | 587 +++++++++++++++++++++++++++++++++++++++++++++ scheduling/BTLtoRTL.v | 7 +- scheduling/BTLtoRTLproof.v | 42 +--- scheduling/RTLtoBTL.v | 3 +- scheduling/RTLtoBTLproof.v | 44 +--- 6 files changed, 614 insertions(+), 634 deletions(-) create mode 100644 scheduling/BTLmatchRTL.v (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 2c1dda5b..6536addb 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -90,8 +90,6 @@ Coercion BF: final >-> Funclass. - the same trick appears for all "basic" instructions and [Icond]. *) - - Record iblock_info := { entry: iblock; input_regs: Regset.t; (* extra liveness information for BTL functional semantics *) @@ -187,6 +185,8 @@ Section RELSEM. Here, [or] is an optional register that will be assigned after exiting the iblock, but before entering in [pc]: e.g. the register set by a function call before entering in return address. + See [tr_inputs] implementation below. + *) Variable tr_exit: function -> list exit -> option reg -> regset -> regset. @@ -390,14 +390,6 @@ Inductive final_state: state -> int -> Prop := | final_state_intro: forall r m, final_state (Returnstate nil (Vint r) m) r. -(** The "raw" CFG small-step semantics for a BTL program (e.g. close to RTL semantics) *) - -(* tid = transfer_identity *) -Definition tid (_:function) (_:list exit) (_: option reg) (rs:regset) := rs. - -Definition cfgsem (p: program) := - Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). - (** The full "functional" small-step semantics for a BTL program. at each exit, we only transfer register in "input_regs" (i.e. "alive" registers). *) @@ -633,9 +625,9 @@ Proof. - repeat (rewrite Regmap.gso); auto. Qed. -(* * Simulation of BTL fsem by BTL cfgsem: a lock-step less-def simulation. +(* * Comparing BTL semantics modulo [regs_lessdef]. -We start by extending the previous [equiv_*] stuff for [Val.lessdef]: we need to also compare memories +This extends the previous [equiv_*] stuff for [Val.lessdef]. Here, we need to also compare memories for Mem.extends, because Vundef values are propagated in memory through stores, allocation, etc. *) @@ -721,16 +713,6 @@ Proof. eapply set_reg_lessdef; eauto. intros r2; eauto. Qed. -Local Hint Resolve init_regs_lessdef_preserv: core. - -Lemma tr_inputs_lessdef f rs1 rs2 tbl or r - (REGS: forall r, Val.lessdef rs1#r rs2#r) - :Val.lessdef (tr_inputs f tbl or rs1)#r (tid f tbl or rs2)#r. -Proof. - unfold tid; rewrite tr_inputs_get. - autodestruct. -Qed. -Local Hint Resolve tr_inputs_lessdef: core. Lemma find_function_lessdef ge ros rs1 rs2 fd (FIND: find_function ge ros rs1 = Some fd) @@ -742,547 +724,12 @@ Proof. intros (b & EQ). destruct (REGS r); try congruence. Qed. -Local Hint Resolve find_function_lessdef regs_lessdef_regs: core. - - -Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 ib rs1' m1' of - (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of) - rs2 m2 - (REGS: forall r, Val.lessdef rs1#r rs2#r) - (MEMS: Mem.extends m1 m2) - : exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) - /\ (forall r, Val.lessdef rs1'#r rs2'#r) - /\ (Mem.extends m1' m2'). -Proof. - induction ISTEP; simpl; try_simplify_someHyps. -Admitted. - -Local Hint Constructors final_step list_forall2: core. -Local Hint Unfold regs_lessdef: core. - -Lemma fsem2cfgsem_finalstep_simu ge stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 - (FSTEP: final_step tr_inputs ge stk1 f sp rs1 m1 fin t s1) - (STACKS: list_forall2 lessdef_stackframe stk1 stk2) - (REGS: forall r, Val.lessdef rs1#r rs2#r) - (MEMS: Mem.extends m1 m2) - : exists s2, final_step tid ge stk2 f sp rs2 m2 fin t s2 - /\ lessdef_state s1 s2. -Proof. - destruct FSTEP; try (eexists; split; simpl; econstructor; eauto; fail). - - (* return *) - exploit Mem.free_parallel_extends; eauto. - intros (m2' & FREE & MEMS2). - eexists; split; simpl; econstructor; eauto. - destruct or; simpl; auto. - - (* tailcall *) - exploit Mem.free_parallel_extends; eauto. - intros (m2' & FREE & MEMS2). - eexists; split; simpl; econstructor; eauto. - - (* builtin *) - exploit (eval_builtin_args_lessdef (ge:=ge) (e1:=(fun r => rs1 # r)) (fun r => rs2 # r)); eauto. - intros (vl2' & BARGS & VARGS). - exploit external_call_mem_extends; eauto. - intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). - eexists; split; simpl; econstructor; eauto. - intros; apply set_res_lessdef; eauto. - - (* jumptable *) - eexists; split; simpl; econstructor; eauto. - destruct (REGS arg); try congruence. -Qed. - -Lemma fsem2cfgsem_ibstep_simu ge stk1 stk2 f sp rs1 m1 ib t s1 rs2 m2 - (STEP: iblock_step tr_inputs ge stk1 f sp rs1 m1 ib t s1) - (STACKS : list_forall2 lessdef_stackframe stk1 stk2) - (REGS: forall r, Val.lessdef rs1#r rs2#r) - (MEMS : Mem.extends m1 m2) - : exists s2, iblock_step tid ge stk2 f sp rs2 m2 ib t s2 - /\ lessdef_state s1 s2. -Proof. - destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP). - exploit fsem2cfgsem_ibistep_simu; eauto. - intros (rs2' & m2' & ISTEP2 & REGS2 & MEMS2). - rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP REGS MEMS. - exploit fsem2cfgsem_finalstep_simu; eauto. - intros (s2 & FSTEP2 & LESSDEF). clear FSTEP. - unfold iblock_step; eexists; split; eauto. -Qed. - -Local Hint Constructors step: core. - -Lemma fsem2cfgsem_step_simu ge s1 t s1' s2 - (STEP: step tr_inputs ge s1 t s1') - (REGS: lessdef_state s1 s2) - :exists s2' : state, - step tid ge s2 t s2' /\ - lessdef_state s1' s2'. -Proof. - destruct STEP; inv REGS. - - (* iblock *) - intros; exploit fsem2cfgsem_ibstep_simu; eauto. clear STEP. - intros (s2 & STEP2 & REGS2). - eexists; split; eauto. - - (* internal call *) - exploit (Mem.alloc_extends m m2 0 (fn_stacksize f) stk m' 0 (fn_stacksize f)); eauto. - 1-2: lia. - intros (m2' & ALLOC2 & MEMS2). clear ALLOC MEMS. - eexists; split; econstructor; eauto. - - (* external call *) - exploit external_call_mem_extends; eauto. - intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). clear EXTCALL MEMS. - eexists; split; econstructor; eauto. - - (* return *) - inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst. - inv STF2. - eexists; split; econstructor; eauto. - intros; eapply set_reg_lessdef; eauto. -Qed. - -Theorem fsem2cfgsem prog: - forward_simulation (fsem prog) (cfgsem prog). -Proof. - eapply forward_simulation_step with lessdef_state; simpl; auto. - - destruct 1; intros; eexists; intuition eauto. econstructor; eauto. - - intros s1 s2 r REGS FINAL; destruct FINAL. - inv REGS; inv STACKS; inv REGS0. - econstructor; eauto. - - intros; eapply fsem2cfgsem_step_simu; eauto. -Qed. - -(** TODO: est-ce qu'il faudrait déplacer ceci dans une librairie du genre "BTLlib" ? *) - -(** Matching BTL and RTL code - -We should be able to define a single verifier able to prove a bisimulation between BTL and RTL code. - -NB 1: the proof of BTL -> RTL (plus simulation) should be much easier than RTL -> BTL (star simulation). - -NB 2: our scheme allows the BTL to duplicate some RTL code. -- in other words: RTL -> BTL allows tail duplication, loop unrolling, etc. Exactly like "Duplicate" on RTL. -- BTL -> RTL allows to "undo" some duplications (e.g. remove duplications that have not enabled interesting optimizations) ! - -Hence, in a sense, our verifier imitates the approach of Duplicate, where [dupmap] maps the BTL nodes to the RTL nodes. - -The [match_*] definitions gives a "relational" specification of the verifier... -*) - -Require Import Errors. - -Inductive match_final_inst (dupmap: PTree.t node): final -> instruction -> Prop := -(* TODO: it may be simplify the oracle for BTL -> RTL, but may makes the verifier more complex ? - | 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: - list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' -> - match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln') -. - -Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := - | ijo_None_left o: is_join_opt None o o - | ijo_None_right o: is_join_opt o None o - | ijo_Some x: is_join_opt (Some x) (Some x) (Some x) - . - -(* [match_iblock dupmap cfg isfst pc ib opc] means that [ib] match a block in a RTL code starting at [pc], with: - - [isfst] (in "input") indicates that no step in the surrounding block has been executed before entering [pc] - - if [opc] (in "ouput") is [None], this means that all branches of the block ends on a final instruction - - if [opc] is [Some pc'], this means that all branches of the block that do not exit, join on [pc']. -*) -Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> iblock -> (option node) -> Prop := - | mib_BF isfst fi pc i iinfo: - cfg!pc = Some i -> - match_final_inst dupmap fi i -> - match_iblock dupmap cfg isfst pc (BF fi iinfo) None - | mib_nop_on_rtl isfst pc pc' iinfo: - cfg!pc = Some (Inop pc') -> - match_iblock dupmap cfg isfst pc (Bnop (Some iinfo)) (Some pc') - | mib_nop_skip pc: - match_iblock dupmap cfg false pc (Bnop None) (Some pc) - | mib_op isfst pc pc' op lr r iinfo: - cfg!pc = Some (Iop op lr r pc') -> - match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc') - | mib_load isfst pc pc' m a lr r iinfo: - cfg!pc = Some (Iload TRAP m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r iinfo) (Some pc') - | mib_store isfst pc pc' m a lr r iinfo: - cfg!pc = Some (Istore m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bstore m a lr r iinfo) (Some pc') - | mib_exit pc pc' iinfo: - dupmap!pc = (Some pc') -> - match_iblock dupmap cfg false pc' (BF (Bgoto pc) iinfo) None - (* NB: on RTL side, we exit the block by a "basic" instruction (or Icond). - Thus some step should have been executed before [pc'] in the RTL code *) - | mib_seq_Some isfst b1 b2 pc1 pc2 opc: - match_iblock dupmap cfg isfst pc1 b1 (Some pc2) -> - match_iblock dupmap cfg false pc2 b2 opc -> - match_iblock dupmap cfg isfst pc1 (Bseq b1 b2) opc -(* | mib_seq_None isfst b1 b2 pc: - match_iblock dupmap cfg isfst pc b1 None -> - match_iblock dupmap cfg isfst (Bseq b1 b2) pc None - (* TODO: here [b2] is dead code ! Is this case really useful ? - The oracle could remove this dead code. - 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 iinfo: - 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 iinfo) opc - . - -Definition match_cfg dupmap (cfg: code) (cfg':RTL.code): Prop := - forall pc pc', dupmap!pc = Some pc' -> - exists ib, cfg!pc = Some ib /\ match_iblock dupmap cfg' true pc' ib.(entry) None. - -(** Shared verifier between RTL -> BTL and BTL -> RTL *) - -Local Open Scope error_monad_scope. - -Definition verify_is_copy dupmap n n' := - match dupmap!n with - | None => Error(msg "BTL.verify_is_copy None") - | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "BTL.verify_is_copy invalid map") end - end. - -Fixpoint verify_is_copy_list dupmap ln ln' := - match ln with - | n::ln => match ln' with - | n'::ln' => do _ <- verify_is_copy dupmap n n'; - verify_is_copy_list dupmap ln ln' - | nil => Error (msg "BTL.verify_is_copy_list: ln' bigger than ln") end - | nil => match ln' with - | n :: ln' => Error (msg "BTL.verify_is_copy_list: ln bigger than ln'") - | nil => OK tt end - end. - -Lemma verify_is_copy_correct dupmap n n' tt: - verify_is_copy dupmap n n' = OK tt -> - dupmap ! n = Some n'. -Proof. - unfold verify_is_copy; repeat autodestruct. - intros NP H; destruct (_ ! n) eqn:REVM; [|inversion H]. - eapply Pos.compare_eq in NP. congruence. -Qed. -Local Hint Resolve verify_is_copy_correct: core. - -Lemma verify_is_copy_list_correct dupmap ln: forall ln' tt, - verify_is_copy_list dupmap ln ln' = OK tt -> - list_forall2 (fun n n' => dupmap ! n = Some n') ln ln'. -Proof. - induction ln. - - intros. destruct ln'; monadInv H. constructor. - - intros. destruct ln'; monadInv H. constructor; eauto. -Qed. - -(* TODO Copied from duplicate, should we import ? *) -Lemma product_eq {A B: Type} : - (forall (a b: A), {a=b} + {a<>b}) -> - (forall (c d: B), {c=d} + {c<>d}) -> - forall (x y: A+B), {x=y} + {x<>y}. -Proof. - intros H H'. intros. decide equality. -Qed. - -(* TODO Copied from duplicate, should we import ? *) -(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application" - * error when doing so *) -Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}. -Proof. - intros. - apply (builtin_arg_eq Pos.eq_dec). -Defined. -Global Opaque builtin_arg_eq_pos. - -(* TODO Copied from duplicate, should we import ? *) -Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}. -Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed. -Global Opaque builtin_res_eq_pos. - -Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := - match ib with - | BF fi _ => - match fi with - | Bgoto pc1 => - do u <- verify_is_copy dupmap pc1 pc; - if negb isfst then - OK None - else Error (msg "BTL.verify_block: isfst is true Bgoto") - | 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 => - match cfg!pc with - | Some (Icall s' ri' lr' r' pc2) => - do u <- verify_is_copy dupmap pc1 pc2; - if (signature_eq s s') then - if (product_eq Pos.eq_dec ident_eq ri ri') then - if (list_eq_dec Pos.eq_dec lr lr') then - if (Pos.eq_dec r r') then OK None - else Error (msg "BTL.verify_block: different r r' in Bcall") - else Error (msg "BTL.verify_block: different lr in Bcall") - else Error (msg "BTL.verify_block: different ri in Bcall") - else Error (msg "BTL.verify_block: different signatures in Bcall") - | _ => Error (msg "BTL.verify_block: incorrect cfg Bcall") - end - | Btailcall s ri lr => - match cfg!pc with - | Some (Itailcall s' ri' lr') => - if (signature_eq s s') then - if (product_eq Pos.eq_dec ident_eq ri ri') then - if (list_eq_dec Pos.eq_dec lr lr') then OK None - else Error (msg "BTL.verify_block: different lr in Btailcall") - else Error (msg "BTL.verify_block: different ri in Btailcall") - else Error (msg "BTL.verify_block: different signatures in Btailcall") - | _ => Error (msg "BTL.verify_block: incorrect cfg Btailcall") - end - | Bbuiltin ef la br pc1 => - match cfg!pc with - | Some (Ibuiltin ef' la' br' pc2) => - do u <- verify_is_copy dupmap pc1 pc2; - if (external_function_eq ef ef') then - if (list_eq_dec builtin_arg_eq_pos la la') then - if (builtin_res_eq_pos br br') then OK None - else Error (msg "BTL.verify_block: different brr in Bbuiltin") - else Error (msg "BTL.verify_block: different lbar in Bbuiltin") - else Error (msg "BTL.verify_block: different ef in Bbuiltin") - | _ => Error (msg "BTL.verify_block: incorrect cfg Bbuiltin") - end - | Bjumptable r ln => - match cfg!pc with - | Some (Ijumptable r' ln') => - do u <- verify_is_copy_list dupmap ln ln'; - if (Pos.eq_dec r r') then OK None - else Error (msg "BTL.verify_block: different r in Bjumptable") - | _ => Error (msg "BTL.verify_block: incorrect cfg Bjumptable") - end - end - | Bnop oiinfo => - match oiinfo with - | Some _ => - match cfg!pc with - | Some (Inop pc') => OK (Some pc') - | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") - end - | None => - if negb isfst then OK (Some pc) - else Error (msg "BTL.verify_block: isfst is true Bnop (on_rtl is false)") - end - | Bop op lr r _ => - match cfg!pc with - | Some (Iop op' lr' r' pc') => - if (eq_operation op op') then - if (list_eq_dec Pos.eq_dec lr lr') then - if (Pos.eq_dec r r') then - OK (Some pc') - else Error (msg "BTL.verify_block: different r in Bop") - else Error (msg "BTL.verify_block: different lr in Bop") - 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 _ => - match cfg!pc with - | Some (Iload tm' m' a' lr' r' pc') => - if (trapping_mode_eq tm TRAP && trapping_mode_eq tm' TRAP) then - if (chunk_eq m m') then - if (eq_addressing a a') then - if (list_eq_dec Pos.eq_dec lr lr') then - if (Pos.eq_dec r r') then - OK (Some pc') - else Error (msg "BTL.verify_block: different r in Bload") - else Error (msg "BTL.verify_block: different lr in Bload") - else Error (msg "BTL.verify_block: different addressing in Bload") - else Error (msg "BTL.verify_block: different mchunk in Bload") - 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 _ => - match cfg!pc with - | Some (Istore m' a' lr' r' pc') => - if (chunk_eq m m') then - if (eq_addressing a a') then - if (list_eq_dec Pos.eq_dec lr lr') then - if (Pos.eq_dec r r') then OK (Some pc') - else Error (msg "BTL.verify_block: different r in Bstore") - else Error (msg "BTL.verify_block: different lr in Bstore") - else Error (msg "BTL.verify_block: different addressing in Bstore") - else Error (msg "BTL.verify_block: different mchunk in Bstore") - | _ => Error (msg "BTL.verify_block: incorrect cfg Bstore") - end - | Bseq b1 b2 => - do opc <- verify_block dupmap cfg isfst pc b1; - match opc with - | Some pc' => - 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 _ => - match cfg!pc with - | Some (Icond c' lr' pcso pcnot i') => - if (list_eq_dec Pos.eq_dec lr lr') then - if (eq_condition c c') then - do opc1 <- verify_block dupmap cfg false pcso bso; - do opc2 <- verify_block dupmap cfg false pcnot bnot; - match opc1, opc2 with - | None, o => OK o - | o, None => OK o - | Some x, Some x' => - if Pos.eq_dec x x' then OK (Some x) - else Error (msg "BTL.verify_block: is_join_opt incorrect for Bcond") - end - else Error (msg "BTL.verify_block: incompatible conditions in Bcond") - else Error (msg "BTL.verify_block: different lr in Bcond") - | _ => Error (msg "BTL.verify_block: incorrect cfg Bcond") - end - end. - -(* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *) -Lemma verify_block_correct dupmap cfg ib: forall pc isfst fin, - verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin. -Proof. - induction ib; intros; - try (unfold verify_block in H; destruct (cfg!pc) eqn:EQCFG; [ idtac | discriminate; fail ]). - - (* BF *) - destruct fi; unfold verify_block in H. - + (* Bgoto *) - monadInv H. - destruct (isfst); simpl in EQ0; inv EQ0. - eapply verify_is_copy_correct in EQ. - constructor; assumption. - + (* Breturn *) - destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. - destruct (option_eq _ _ _); try discriminate. inv H. - eapply mib_BF; eauto. constructor. - + (* Bcall *) - destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. - monadInv H. - eapply verify_is_copy_correct in EQ. - destruct (signature_eq _ _); try discriminate. - destruct (product_eq _ _ _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (Pos.eq_dec _ _); try discriminate. subst. - inv EQ0. eapply mib_BF; eauto. constructor; assumption. - + (* Btailcall *) - destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. - destruct (signature_eq _ _); try discriminate. - destruct (product_eq _ _ _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. subst. - inv H. eapply mib_BF; eauto. constructor. - + (* Bbuiltin *) - destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. - monadInv H. - eapply verify_is_copy_correct in EQ. - destruct (external_function_eq _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (builtin_res_eq_pos _ _); try discriminate. subst. - inv EQ0. eapply mib_BF; eauto. constructor; assumption. - + (* Bjumptable *) - destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. - monadInv H. - eapply verify_is_copy_list_correct in EQ. - destruct (Pos.eq_dec _ _); try discriminate. subst. - inv EQ0. eapply mib_BF; eauto. constructor; assumption. - - (* Bnop *) - destruct oiinfo; simpl in *. - + destruct (cfg!pc) eqn:EQCFG; try discriminate. - destruct i0; inv H. constructor; assumption. - + destruct isfst; try discriminate. inv H. - constructor; assumption. - - (* Bop *) - destruct i; try discriminate. - destruct (eq_operation _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (Pos.eq_dec _ _); try discriminate. inv H. - constructor; assumption. - - (* Bload *) - destruct i; try discriminate. - do 2 (destruct (trapping_mode_eq _ _); try discriminate). - simpl in H. - destruct (chunk_eq _ _); try discriminate. - destruct (eq_addressing _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (Pos.eq_dec _ _); try discriminate. inv H. - constructor; assumption. - - (* Bstore *) - destruct i; try discriminate. - destruct (chunk_eq _ _); try discriminate. - destruct (eq_addressing _ _); try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (Pos.eq_dec _ _); try discriminate. inv H. - constructor; assumption. - - (* Bseq *) - monadInv H. - destruct x; try discriminate. - eapply mib_seq_Some. - eapply IHib1; eauto. - eapply IHib2; eauto. - - (* Bcond *) - destruct i; try discriminate. - destruct (list_eq_dec _ _ _); try discriminate. - destruct (eq_condition _ _); try discriminate. - fold (verify_block dupmap cfg false n ib1) in H. - fold (verify_block dupmap cfg false n0 ib2) in H. - monadInv H. - destruct x, x0; try destruct (Pos.eq_dec _ _); try discriminate. - all: inv EQ2; eapply mib_cond; eauto; econstructor. -Qed. -Local Hint Resolve verify_block_correct: core. - -Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit := - match l with - | nil => OK tt - | (pc, pc')::l => - match cfg!pc with - | Some ib => do o <- verify_block dupmap cfg' true pc' ib.(entry); - match o with - | None => verify_blocks dupmap cfg cfg' l - | _ => Error(msg "BTL.verify_blocks.end") - end - | _ => Error(msg "BTL.verify_blocks.entry") - end - end. -Definition verify_cfg dupmap (cfg: code) (cfg':RTL.code): res unit := - verify_blocks dupmap cfg cfg' (PTree.elements dupmap). - -Lemma verify_cfg_correct dupmap cfg cfg' tt: - verify_cfg dupmap cfg cfg' = OK tt -> - match_cfg dupmap cfg cfg'. -Proof. - unfold verify_cfg. - intros X pc pc' H; generalize X; clear X. - exploit PTree.elements_correct; eauto. - generalize tt pc pc' H; clear tt pc pc' H. - generalize (PTree.elements dupmap). - induction l as [|[pc1 pc1']l]; simpl. - - tauto. - - intros pc pc' DUP u H. - unfold bind. - repeat autodestruct. - intros; subst. - destruct H as [H|H]; eauto. - inversion H; subst. - eexists; split; eauto. -Qed. - -Definition verify_function dupmap f f' : res unit := - do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); - verify_cfg dupmap (fn_code f) (RTL.fn_code f'). +(** * Auxiliary general purpose functions on BTL *) Definition is_goto (ib: iblock): bool := match ib with | BF (Bgoto _) _ => true | _ => false end. + diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v new file mode 100644 index 00000000..23ff6681 --- /dev/null +++ b/scheduling/BTLmatchRTL.v @@ -0,0 +1,587 @@ +(** General notions about the bisimulation BTL <-> RTL. + +This bisimulation is proved on the "CFG semantics" of BTL called [cfgsem] defined below, +such that the full state of registers is preserved when exiting blocks. + +*) + +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad Lia. +Require Export BTL. + +(* tid = transfer_identity *) +Definition tid (_:function) (_:list exit) (_: option reg) (rs:regset) := rs. + +Definition cfgsem (p: program) := + Semantics (step tid) (initial_state p) final_state (Genv.globalenv p). + +(* * Simulation of BTL fsem by BTL cfgsem: a lock-step less-def simulation. +*) + + +Lemma tr_inputs_lessdef f rs1 rs2 tbl or r + (REGS: forall r, Val.lessdef rs1#r rs2#r) + :Val.lessdef (tr_inputs f tbl or rs1)#r (tid f tbl or rs2)#r. +Proof. + unfold tid; rewrite tr_inputs_get. + autodestruct. +Qed. + +Local Hint Resolve tr_inputs_lessdef init_regs_lessdef_preserv find_function_lessdef regs_lessdef_regs + lessdef_state_refl: core. + +Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 ib rs1' m1' of + (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of) + rs2 m2 + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS: Mem.extends m1 m2) + : exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) + /\ (forall r, Val.lessdef rs1'#r rs2'#r) + /\ (Mem.extends m1' m2'). +Proof. + induction ISTEP; simpl; try_simplify_someHyps. +Admitted. + +Local Hint Constructors lessdef_stackframe lessdef_state final_step list_forall2 step: core. +Local Hint Unfold regs_lessdef: core. + +Lemma fsem2cfgsem_finalstep_simu ge stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 + (FSTEP: final_step tr_inputs ge stk1 f sp rs1 m1 fin t s1) + (STACKS: list_forall2 lessdef_stackframe stk1 stk2) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS: Mem.extends m1 m2) + : exists s2, final_step tid ge stk2 f sp rs2 m2 fin t s2 + /\ lessdef_state s1 s2. +Proof. + destruct FSTEP; try (eexists; split; simpl; econstructor; eauto; fail). + - (* return *) + exploit Mem.free_parallel_extends; eauto. + intros (m2' & FREE & MEMS2). + eexists; split; simpl; econstructor; eauto. + destruct or; simpl; auto. + - (* tailcall *) + exploit Mem.free_parallel_extends; eauto. + intros (m2' & FREE & MEMS2). + eexists; split; simpl; econstructor; eauto. + - (* builtin *) + exploit (eval_builtin_args_lessdef (ge:=ge) (e1:=(fun r => rs1 # r)) (fun r => rs2 # r)); eauto. + intros (vl2' & BARGS & VARGS). + exploit external_call_mem_extends; eauto. + intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). + eexists; split; simpl; econstructor; eauto. + intros; apply set_res_lessdef; eauto. + - (* jumptable *) + eexists; split; simpl; econstructor; eauto. + destruct (REGS arg); try congruence. +Qed. + +Lemma fsem2cfgsem_ibstep_simu ge stk1 stk2 f sp rs1 m1 ib t s1 rs2 m2 + (STEP: iblock_step tr_inputs ge stk1 f sp rs1 m1 ib t s1) + (STACKS : list_forall2 lessdef_stackframe stk1 stk2) + (REGS: forall r, Val.lessdef rs1#r rs2#r) + (MEMS : Mem.extends m1 m2) + : exists s2, iblock_step tid ge stk2 f sp rs2 m2 ib t s2 + /\ lessdef_state s1 s2. +Proof. + destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP). + exploit fsem2cfgsem_ibistep_simu; eauto. + intros (rs2' & m2' & ISTEP2 & REGS2 & MEMS2). + rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP REGS MEMS. + exploit fsem2cfgsem_finalstep_simu; eauto. + intros (s2 & FSTEP2 & LESSDEF). clear FSTEP. + unfold iblock_step; eexists; split; eauto. +Qed. + +Local Hint Constructors step: core. + +Lemma fsem2cfgsem_step_simu ge s1 t s1' s2 + (STEP: step tr_inputs ge s1 t s1') + (REGS: lessdef_state s1 s2) + :exists s2' : state, + step tid ge s2 t s2' /\ + lessdef_state s1' s2'. +Proof. + destruct STEP; inv REGS. + - (* iblock *) + intros; exploit fsem2cfgsem_ibstep_simu; eauto. clear STEP. + intros (s2 & STEP2 & REGS2). + eexists; split; eauto. + - (* internal call *) + exploit (Mem.alloc_extends m m2 0 (fn_stacksize f) stk m' 0 (fn_stacksize f)); eauto. + 1-2: lia. + intros (m2' & ALLOC2 & MEMS2). clear ALLOC MEMS. + eexists; split; econstructor; eauto. + - (* external call *) + exploit external_call_mem_extends; eauto. + intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). clear EXTCALL MEMS. + eexists; split; econstructor; eauto. + - (* return *) + inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst. + inv STF2. + eexists; split; econstructor; eauto. + intros; eapply set_reg_lessdef; eauto. +Qed. + +Theorem fsem2cfgsem prog: + forward_simulation (fsem prog) (cfgsem prog). +Proof. + eapply forward_simulation_step with lessdef_state; simpl; auto. + - destruct 1; intros; eexists; intuition eauto. econstructor; eauto. + - intros s1 s2 r REGS FINAL; destruct FINAL. + inv REGS; inv STACKS; inv REGS0. + econstructor; eauto. + - intros; eapply fsem2cfgsem_step_simu; eauto. +Qed. + +(** * Matching BTL (for cfgsem) and RTL code + +We define a single verifier able to prove a bisimulation between BTL and RTL code. + +Hence, in a sense, our verifier imitates the approach of Duplicate, where [dupmap] maps the BTL nodes to the RTL nodes. + +The [match_function] definition gives a "relational" specification of the verifier... +*) + +Require Import Errors. + +Inductive match_final_inst (dupmap: PTree.t node): final -> instruction -> Prop := + | 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: + list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' -> + match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln') +. + +Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop := + | ijo_None_left o: is_join_opt None o o + | ijo_None_right o: is_join_opt o None o + | ijo_Some x: is_join_opt (Some x) (Some x) (Some x) + . + +(* [match_iblock dupmap cfg isfst pc ib opc] means that [ib] match a block in a RTL code starting at [pc], with: + - [isfst] (in "input") indicates that no step in the surrounding block has been executed before entering [pc] + - if [opc] (in "ouput") is [None], this means that all branches of the block ends on a final instruction + - if [opc] is [Some pc'], this means that all branches of the block that do not exit, join on [pc']. +*) +Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> iblock -> (option node) -> Prop := + | mib_BF isfst fi pc i iinfo: + cfg!pc = Some i -> + match_final_inst dupmap fi i -> + match_iblock dupmap cfg isfst pc (BF fi iinfo) None + | mib_nop_on_rtl isfst pc pc' iinfo: + cfg!pc = Some (Inop pc') -> + match_iblock dupmap cfg isfst pc (Bnop (Some iinfo)) (Some pc') + | mib_nop_skip pc: + match_iblock dupmap cfg false pc (Bnop None) (Some pc) + | mib_op isfst pc pc' op lr r iinfo: + cfg!pc = Some (Iop op lr r pc') -> + match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc') + | mib_load isfst pc pc' m a lr r iinfo: + cfg!pc = Some (Iload TRAP m a lr r pc') -> + match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r iinfo) (Some pc') + | mib_store isfst pc pc' m a lr r iinfo: + cfg!pc = Some (Istore m a lr r pc') -> + match_iblock dupmap cfg isfst pc (Bstore m a lr r iinfo) (Some pc') + | mib_exit pc pc' iinfo: + dupmap!pc = (Some pc') -> + match_iblock dupmap cfg false pc' (BF (Bgoto pc) iinfo) None + (* NB: on RTL side, we exit the block by a "basic" instruction (or Icond). + Thus some step should have been executed before [pc'] in the RTL code *) + | mib_seq_Some isfst b1 b2 pc1 pc2 opc: + match_iblock dupmap cfg isfst pc1 b1 (Some pc2) -> + match_iblock dupmap cfg false pc2 b2 opc -> + match_iblock dupmap cfg isfst pc1 (Bseq b1 b2) opc +(* | mib_seq_None isfst b1 b2 pc: + match_iblock dupmap cfg isfst pc b1 None -> + match_iblock dupmap cfg isfst (Bseq b1 b2) pc None + (* here [b2] is dead code ! Our verifier rejects such dead code! + *) +*) + | mib_cond isfst c lr bso bnot pcso pcnot pc opc1 opc2 opc i iinfo: + 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 iinfo) opc + . + +Definition match_cfg dupmap (cfg: BTL.code) (cfg':RTL.code): Prop := + forall pc pc', dupmap!pc = Some pc' -> + exists ib, cfg!pc = Some ib /\ match_iblock dupmap cfg' true pc' ib.(entry) None. + +Record match_function dupmap f f': Prop := { + dupmap_correct: match_cfg dupmap (BTL.fn_code f) (RTL.fn_code f'); + dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f'); + preserv_fnsig: fn_sig f = RTL.fn_sig f'; + preserv_fnparams: fn_params f = RTL.fn_params f'; + preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' +}. + +Inductive match_fundef: BTL.fundef -> RTL.fundef -> Prop := + | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc = Some pc') + : match_stackframes (BTL.Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). + + +(** * Shared verifier between RTL -> BTL and BTL -> RTL *) + +Local Open Scope error_monad_scope. + +Definition verify_is_copy dupmap n n' := + match dupmap!n with + | None => Error(msg "BTL.verify_is_copy None") + | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "BTL.verify_is_copy invalid map") end + end. + +Fixpoint verify_is_copy_list dupmap ln ln' := + match ln with + | n::ln => match ln' with + | n'::ln' => do _ <- verify_is_copy dupmap n n'; + verify_is_copy_list dupmap ln ln' + | nil => Error (msg "BTL.verify_is_copy_list: ln' bigger than ln") end + | nil => match ln' with + | n :: ln' => Error (msg "BTL.verify_is_copy_list: ln bigger than ln'") + | nil => OK tt end + end. + +Lemma verify_is_copy_correct dupmap n n' tt: + verify_is_copy dupmap n n' = OK tt -> + dupmap ! n = Some n'. +Proof. + unfold verify_is_copy; repeat autodestruct. + intros NP H; destruct (_ ! n) eqn:REVM; [|inversion H]. + eapply Pos.compare_eq in NP. congruence. +Qed. +Local Hint Resolve verify_is_copy_correct: core. + +Lemma verify_is_copy_list_correct dupmap ln: forall ln' tt, + verify_is_copy_list dupmap ln ln' = OK tt -> + list_forall2 (fun n n' => dupmap ! n = Some n') ln ln'. +Proof. + induction ln. + - intros. destruct ln'; monadInv H. constructor. + - intros. destruct ln'; monadInv H. constructor; eauto. +Qed. + +(* TODO Copied from duplicate, should we import ? *) +Lemma product_eq {A B: Type} : + (forall (a b: A), {a=b} + {a<>b}) -> + (forall (c d: B), {c=d} + {c<>d}) -> + forall (x y: A+B), {x=y} + {x<>y}. +Proof. + intros H H'. intros. decide equality. +Qed. + +(* TODO Copied from duplicate, should we import ? *) +(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application" + * error when doing so *) +Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}. +Proof. + intros. + apply (builtin_arg_eq Pos.eq_dec). +Defined. +Global Opaque builtin_arg_eq_pos. + +(* TODO Copied from duplicate, should we import ? *) +Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}. +Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed. +Global Opaque builtin_res_eq_pos. + +Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) := + match ib with + | BF fi _ => + match fi with + | Bgoto pc1 => + do u <- verify_is_copy dupmap pc1 pc; + if negb isfst then + OK None + else Error (msg "BTL.verify_block: isfst is true Bgoto") + | 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 => + match cfg!pc with + | Some (Icall s' ri' lr' r' pc2) => + do u <- verify_is_copy dupmap pc1 pc2; + if (signature_eq s s') then + if (product_eq Pos.eq_dec ident_eq ri ri') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then OK None + else Error (msg "BTL.verify_block: different r r' in Bcall") + else Error (msg "BTL.verify_block: different lr in Bcall") + else Error (msg "BTL.verify_block: different ri in Bcall") + else Error (msg "BTL.verify_block: different signatures in Bcall") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bcall") + end + | Btailcall s ri lr => + match cfg!pc with + | Some (Itailcall s' ri' lr') => + if (signature_eq s s') then + if (product_eq Pos.eq_dec ident_eq ri ri') then + if (list_eq_dec Pos.eq_dec lr lr') then OK None + else Error (msg "BTL.verify_block: different lr in Btailcall") + else Error (msg "BTL.verify_block: different ri in Btailcall") + else Error (msg "BTL.verify_block: different signatures in Btailcall") + | _ => Error (msg "BTL.verify_block: incorrect cfg Btailcall") + end + | Bbuiltin ef la br pc1 => + match cfg!pc with + | Some (Ibuiltin ef' la' br' pc2) => + do u <- verify_is_copy dupmap pc1 pc2; + if (external_function_eq ef ef') then + if (list_eq_dec builtin_arg_eq_pos la la') then + if (builtin_res_eq_pos br br') then OK None + else Error (msg "BTL.verify_block: different brr in Bbuiltin") + else Error (msg "BTL.verify_block: different lbar in Bbuiltin") + else Error (msg "BTL.verify_block: different ef in Bbuiltin") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bbuiltin") + end + | Bjumptable r ln => + match cfg!pc with + | Some (Ijumptable r' ln') => + do u <- verify_is_copy_list dupmap ln ln'; + if (Pos.eq_dec r r') then OK None + else Error (msg "BTL.verify_block: different r in Bjumptable") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bjumptable") + end + end + | Bnop oiinfo => + match oiinfo with + | Some _ => + match cfg!pc with + | Some (Inop pc') => OK (Some pc') + | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop") + end + | None => + if negb isfst then OK (Some pc) + else Error (msg "BTL.verify_block: isfst is true Bnop (on_rtl is false)") + end + | Bop op lr r _ => + match cfg!pc with + | Some (Iop op' lr' r' pc') => + if (eq_operation op op') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then + OK (Some pc') + else Error (msg "BTL.verify_block: different r in Bop") + else Error (msg "BTL.verify_block: different lr in Bop") + 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 _ => + match cfg!pc with + | Some (Iload tm' m' a' lr' r' pc') => + if (trapping_mode_eq tm TRAP && trapping_mode_eq tm' TRAP) then + if (chunk_eq m m') then + if (eq_addressing a a') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then + OK (Some pc') + else Error (msg "BTL.verify_block: different r in Bload") + else Error (msg "BTL.verify_block: different lr in Bload") + else Error (msg "BTL.verify_block: different addressing in Bload") + else Error (msg "BTL.verify_block: different mchunk in Bload") + 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 _ => + match cfg!pc with + | Some (Istore m' a' lr' r' pc') => + if (chunk_eq m m') then + if (eq_addressing a a') then + if (list_eq_dec Pos.eq_dec lr lr') then + if (Pos.eq_dec r r') then OK (Some pc') + else Error (msg "BTL.verify_block: different r in Bstore") + else Error (msg "BTL.verify_block: different lr in Bstore") + else Error (msg "BTL.verify_block: different addressing in Bstore") + else Error (msg "BTL.verify_block: different mchunk in Bstore") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bstore") + end + | Bseq b1 b2 => + do opc <- verify_block dupmap cfg isfst pc b1; + match opc with + | Some pc' => + 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 _ => + match cfg!pc with + | Some (Icond c' lr' pcso pcnot i') => + if (list_eq_dec Pos.eq_dec lr lr') then + if (eq_condition c c') then + do opc1 <- verify_block dupmap cfg false pcso bso; + do opc2 <- verify_block dupmap cfg false pcnot bnot; + match opc1, opc2 with + | None, o => OK o + | o, None => OK o + | Some x, Some x' => + if Pos.eq_dec x x' then OK (Some x) + else Error (msg "BTL.verify_block: is_join_opt incorrect for Bcond") + end + else Error (msg "BTL.verify_block: incompatible conditions in Bcond") + else Error (msg "BTL.verify_block: different lr in Bcond") + | _ => Error (msg "BTL.verify_block: incorrect cfg Bcond") + end + end. + +(* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *) +Lemma verify_block_correct dupmap cfg ib: forall pc isfst fin, + verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin. +Proof. + induction ib; intros; + try (unfold verify_block in H; destruct (cfg!pc) eqn:EQCFG; [ idtac | discriminate; fail ]). + - (* BF *) + destruct fi; unfold verify_block in H. + + (* Bgoto *) + monadInv H. + destruct (isfst); simpl in EQ0; inv EQ0. + eapply verify_is_copy_correct in EQ. + constructor; assumption. + + (* Breturn *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + destruct (option_eq _ _ _); try discriminate. inv H. + eapply mib_BF; eauto. constructor. + + (* Bcall *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_correct in EQ. + destruct (signature_eq _ _); try discriminate. + destruct (product_eq _ _ _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. + + (* Btailcall *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + destruct (signature_eq _ _); try discriminate. + destruct (product_eq _ _ _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. subst. + inv H. eapply mib_BF; eauto. constructor. + + (* Bbuiltin *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_correct in EQ. + destruct (external_function_eq _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (builtin_res_eq_pos _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. + + (* Bjumptable *) + destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate. + monadInv H. + eapply verify_is_copy_list_correct in EQ. + destruct (Pos.eq_dec _ _); try discriminate. subst. + inv EQ0. eapply mib_BF; eauto. constructor; assumption. + - (* Bnop *) + destruct oiinfo; simpl in *. + + destruct (cfg!pc) eqn:EQCFG; try discriminate. + destruct i0; inv H. constructor; assumption. + + destruct isfst; try discriminate. inv H. + constructor; assumption. + - (* Bop *) + destruct i; try discriminate. + destruct (eq_operation _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bload *) + destruct i; try discriminate. + do 2 (destruct (trapping_mode_eq _ _); try discriminate). + simpl in H. + destruct (chunk_eq _ _); try discriminate. + destruct (eq_addressing _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bstore *) + destruct i; try discriminate. + destruct (chunk_eq _ _); try discriminate. + destruct (eq_addressing _ _); try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (Pos.eq_dec _ _); try discriminate. inv H. + constructor; assumption. + - (* Bseq *) + monadInv H. + destruct x; try discriminate. + eapply mib_seq_Some. + eapply IHib1; eauto. + eapply IHib2; eauto. + - (* Bcond *) + destruct i; try discriminate. + destruct (list_eq_dec _ _ _); try discriminate. + destruct (eq_condition _ _); try discriminate. + fold (verify_block dupmap cfg false n ib1) in H. + fold (verify_block dupmap cfg false n0 ib2) in H. + monadInv H. + destruct x, x0; try destruct (Pos.eq_dec _ _); try discriminate. + all: inv EQ2; eapply mib_cond; eauto; econstructor. +Qed. +Local Hint Resolve verify_block_correct: core. + +Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit := + match l with + | nil => OK tt + | (pc, pc')::l => + match cfg!pc with + | Some ib => do o <- verify_block dupmap cfg' true pc' ib.(entry); + match o with + | None => verify_blocks dupmap cfg cfg' l + | _ => Error(msg "BTL.verify_blocks.end") + end + | _ => Error(msg "BTL.verify_blocks.entry") + end + end. + +Definition verify_cfg dupmap (cfg: code) (cfg':RTL.code): res unit := + verify_blocks dupmap cfg cfg' (PTree.elements dupmap). + +Lemma verify_cfg_correct dupmap cfg cfg' tt: + verify_cfg dupmap cfg cfg' = OK tt -> + match_cfg dupmap cfg cfg'. +Proof. + unfold verify_cfg. + intros X pc pc' H; generalize X; clear X. + exploit PTree.elements_correct; eauto. + generalize tt pc pc' H; clear tt pc pc' H. + generalize (PTree.elements dupmap). + induction l as [|[pc1 pc1']l]; simpl. + - tauto. + - intros pc pc' DUP u H. + unfold bind. + repeat autodestruct. + intros; subst. + destruct H as [H|H]; eauto. + inversion H; subst. + eexists; split; eauto. +Qed. + +Definition verify_function dupmap f f' : res unit := + do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f'); + verify_cfg dupmap (fn_code f) (RTL.fn_code f'). + +Lemma verify_function_correct dupmap f f' tt: + verify_function dupmap f f' = OK tt -> + fn_sig f = RTL.fn_sig f' -> + fn_params f = RTL.fn_params f' -> + fn_stacksize f = RTL.fn_stacksize f' -> + match_function dupmap f f'. +Proof. + unfold verify_function; intro VERIF. monadInv VERIF. + constructor; eauto. + eapply verify_cfg_correct; eauto. +Qed. + diff --git a/scheduling/BTLtoRTL.v b/scheduling/BTLtoRTL.v index 1333b406..fc58533d 100644 --- a/scheduling/BTLtoRTL.v +++ b/scheduling/BTLtoRTL.v @@ -1,21 +1,22 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. +Require Export BTLmatchRTL. Require Import Errors Linking. (** External oracle *) -Axiom btl2rtl: function -> RTL.code * node * (PTree.t node). +Axiom btl2rtl: BTL.function -> RTL.code * node * (PTree.t node). Extract Constant btl2rtl => "BTLtoRTLaux.btl2rtl". Local Open Scope error_monad_scope. -Definition transf_function (f: function) : res RTL.function := +Definition transf_function (f: BTL.function) : res RTL.function := let (tcte, dupmap) := btl2rtl f in let (tc, te) := tcte in let f' := RTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - (*do u <- verify_function dupmap f f';*) + do u <- verify_function dupmap f f'; OK f'. Definition transf_fundef (f: fundef) : res RTL.fundef := diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 46f360c5..765f9cad 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -1,31 +1,12 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. -Require Import RTL Op Registers OptionMonad BTL. +Require Import RTL Op Registers OptionMonad. Require Import Errors Linking BTLtoRTL. Require Import Linking. -Record match_function dupmap f f': Prop := { - dupmap_correct: match_cfg dupmap (fn_code f) (RTL.fn_code f'); - dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f'); - preserv_fnsig: fn_sig f = RTL.fn_sig f'; - preserv_fnparams: fn_params f = RTL.fn_params f'; - preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' -}. - -Inductive match_fundef: fundef -> RTL.fundef -> Prop := - | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') - | match_External ef: match_fundef (External ef) (External ef). - -Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := - | match_stackframe_intro - dupmap res f sp pc rs f' pc' - (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc = Some pc') - : match_stackframes (Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). - -Inductive match_states: state -> RTL.state -> Prop := +Inductive match_states: BTL.state -> RTL.state -> Prop := | match_states_intro dupmap st f sp pc rs m st' f' pc' (STACKS: list_forall2 match_stackframes st st') @@ -43,26 +24,13 @@ Inductive match_states: state -> RTL.state -> Prop := : match_states (Returnstate st v m) (RTL.Returnstate st' v m) . -Lemma verify_function_correct dupmap f f' tt: - verify_function dupmap f f' = OK tt -> - fn_sig f = RTL.fn_sig f' -> - fn_params f = RTL.fn_params f' -> - fn_stacksize f = RTL.fn_stacksize f' -> - match_function dupmap f f'. -Proof. - unfold verify_function; intro VERIF. monadInv VERIF. - constructor; eauto. - - eapply verify_cfg_correct; eauto. - - eapply verify_is_copy_correct; eauto. -Qed. - Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. -Proof. Admitted. (* +Proof. unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. - Qed.*) +Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. @@ -402,7 +370,7 @@ Proof. Qed. Theorem transf_program_correct_cfg: - forward_simulation (BTL.cfgsem prog) (RTL.semantics tprog). + forward_simulation (BTLmatchRTL.cfgsem prog) (RTL.semantics tprog). Proof. eapply forward_simulation_plus with match_states. - eapply senv_preserved. diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index 2ac02597..507fc9d9 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -1,6 +1,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. +Require Export BTLmatchRTL. Require Import Errors Linking. @@ -15,7 +16,7 @@ Definition transf_function (f: RTL.function) : res BTL.function := let (tcte, dupmap) := rtl2btl f in let (tc, te) := tcte in let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) tc te in - (*do u <- verify_function dupmap f' f;*) + do u <- verify_function dupmap f' f; OK f'. Definition transf_fundef (f: RTL.fundef) : res fundef := diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index f2f99ef5..ef336de5 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -229,45 +229,21 @@ Qed. (** * Matching relation on functions *) -Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { - dupmap_correct: match_cfg dupmap (fn_code tf) (RTL.fn_code f); - dupmap_entrypoint: dupmap!(fn_entrypoint tf) = Some (RTL.fn_entrypoint f); - preserv_fnsig: fn_sig tf = RTL.fn_sig f; - preserv_fnparams: fn_params tf = RTL.fn_params f; - preserv_fnstacksize: fn_stacksize tf = RTL.fn_stacksize f -}. - -Inductive match_fundef: RTL.fundef -> fundef -> Prop := - | match_Internal dupmap f tf: match_function dupmap f tf -> match_fundef (Internal f) (Internal tf) - | match_External ef: match_fundef (External ef) (External ef). - -Inductive match_stackframes: RTL.stackframe -> stackframe -> Prop := - | match_stackframe_intro - dupmap res f sp pc rs f' pc' - (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc' = Some pc) - : match_stackframes (RTL.Stackframe res f sp pc rs) (Stackframe res f' sp pc' rs). +(* we simply switch [f] and [tf] in the usual way *) +Definition match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := + BTLmatchRTL.match_function dupmap tf f. -Lemma verify_function_correct dupmap f f' tt: - verify_function dupmap f' f = OK tt -> - fn_sig f' = RTL.fn_sig f -> - fn_params f' = RTL.fn_params f -> - fn_stacksize f' = RTL.fn_stacksize f -> - match_function dupmap f f'. -Proof. - unfold verify_function; intro VERIF. monadInv VERIF. - constructor; eauto. - - eapply verify_cfg_correct; eauto. - - eapply verify_is_copy_correct; eauto. -Qed. +Definition match_fundef f tf := BTLmatchRTL.match_fundef tf f. + +Definition match_stackframes stk stk' := BTLmatchRTL.match_stackframes stk' stk. Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. -Proof. Admitted. (* +Proof. unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. eexists; eapply verify_function_correct; simpl; eauto. - Qed.*) +Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. @@ -711,7 +687,7 @@ Proof. inv H. + (* Internal function *) inv TRANSF. - rename H0 into TRANSF. + rename H1 into TRANSF. eapply dupmap_entrypoint in TRANSF as ENTRY. eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. apply DMC in ENTRY as DMC'. @@ -739,7 +715,7 @@ Qed. Local Hint Resolve plus_one star_refl: core. Theorem transf_program_correct: - forward_simulation (RTL.semantics prog) (BTL.cfgsem tprog). + forward_simulation (RTL.semantics prog) (BTLmatchRTL.cfgsem tprog). Proof. eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=cfgsem tprog) (ltof _ omeasure) match_states). constructor 1; simpl. -- cgit From b79d0a04787d9234cf580841bf58e592fe4ab3ee Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 15:24:16 +0200 Subject: starting to extend RTLtoBTL with Liveness checking (on BTL side) --- scheduling/BTL_Livecheck.v | 44 ++++++++++++++++++++++++++++ scheduling/BTL_Scheduler.v | 0 scheduling/BTL_Schedulerproof.v | 0 scheduling/BTLmatchRTL.v | 12 -------- scheduling/BTLtoRTLproof.v | 12 ++++++++ scheduling/RTLtoBTL.v | 3 +- scheduling/RTLtoBTLproof.v | 63 ++++++++++++++++++++++++++--------------- 7 files changed, 98 insertions(+), 36 deletions(-) create mode 100644 scheduling/BTL_Livecheck.v create mode 100644 scheduling/BTL_Scheduler.v create mode 100644 scheduling/BTL_Schedulerproof.v (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v new file mode 100644 index 00000000..2baba025 --- /dev/null +++ b/scheduling/BTL_Livecheck.v @@ -0,0 +1,44 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep Op Registers OptionMonad. +Require Import Errors RTL BTL BTLmatchRTL. + +(** TODO: adapt stuff RTLpathLivegen *) + +Definition liveness_checker (f: BTL.function): res unit := OK tt. (* TODO: implement this ! *) + +Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. + +(** TODO: adapt stuff RTLpathLivegenproof *) + +Local Notation ext alive := (fun r => Regset.In r alive). + +Lemma regset_add_spec live r1 r2: Regset.In r1 (Regset.add r2 live) <-> (r1 = r2 \/ Regset.In r1 live). +Proof. + destruct (Pos.eq_dec r1 r2). + - subst. intuition; eapply Regset.add_1; auto. + - intuition. + * right. eapply Regset.add_3; eauto. + * eapply Regset.add_2; auto. +Qed. + +Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := + forall r, (alive r) -> rs1#r = rs2#r. + +(* continue to adapt stuff RTLpathLivegenproof *) + +Section FSEM_SIMULATES_CFGSEM. + +Variable prog: BTL.program. + +Let ge := Genv.globalenv prog. + +Hypothesis liveness_checker_correct: forall b f, Genv.find_funct_ptr ge b = Some (Internal f) -> liveness_ok_function f. + +Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). +Proof. + exploit liveness_checker_correct. +Admitted. + +End FSEM_SIMULATES_CFGSEM. + + diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v new file mode 100644 index 00000000..e69de29b diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v new file mode 100644 index 00000000..e69de29b diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index 23ff6681..8994579d 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -223,18 +223,6 @@ Record match_function dupmap f f': Prop := { preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f' }. -Inductive match_fundef: BTL.fundef -> RTL.fundef -> Prop := - | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') - | match_External ef: match_fundef (External ef) (External ef). - -Inductive match_stackframes: stackframe -> RTL.stackframe -> Prop := - | match_stackframe_intro - dupmap res f sp pc rs f' pc' - (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc = Some pc') - : match_stackframes (BTL.Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). - - (** * Shared verifier between RTL -> BTL and BTL -> RTL *) Local Open Scope error_monad_scope. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 765f9cad..75f67d51 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -6,6 +6,18 @@ Require Import Errors Linking BTLtoRTL. Require Import Linking. + +Inductive match_fundef: BTL.fundef -> RTL.fundef -> Prop := + | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: BTL.stackframe -> RTL.stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc = Some pc') + : match_stackframes (BTL.Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs). + Inductive match_states: BTL.state -> RTL.state -> Prop := | match_states_intro dupmap st f sp pc rs m st' f' pc' diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v index 507fc9d9..309c616e 100644 --- a/scheduling/RTLtoBTL.v +++ b/scheduling/RTLtoBTL.v @@ -1,7 +1,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. -Require Export BTLmatchRTL. +Require Export BTLmatchRTL BTL_Livecheck. Require Import Errors Linking. @@ -17,6 +17,7 @@ Definition transf_function (f: RTL.function) : res BTL.function := let (tc, te) := tcte in let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) tc te in do u <- verify_function dupmap f' f; + do u <- liveness_checker f'; OK f'. Definition transf_fundef (f: RTL.fundef) : res fundef := diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index ef336de5..18ff8d5f 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -230,20 +230,33 @@ Qed. (** * Matching relation on functions *) (* we simply switch [f] and [tf] in the usual way *) -Definition match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := - BTLmatchRTL.match_function dupmap tf f. +Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := { + matchRTL:> BTLmatchRTL.match_function dupmap tf f; + liveness_ok: liveness_ok_function tf; +}. -Definition match_fundef f tf := BTLmatchRTL.match_fundef tf f. +Local Hint Resolve matchRTL: core. -Definition match_stackframes stk stk' := BTLmatchRTL.match_stackframes stk' stk. +Inductive match_fundef: RTL.fundef -> BTL.fundef -> Prop := + | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: RTL.stackframe -> BTL.stackframe -> Prop := + | match_stackframe_intro + dupmap res f sp pc rs f' pc' + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc' = Some pc) + : match_stackframes (RTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc' rs). Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. Proof. unfold transf_function; unfold bind. repeat autodestruct. intros H _ _ X. inversion X; subst; clear X. +(* eexists; eapply verify_function_correct; simpl; eauto. -Qed. +*) +Admitted. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. @@ -362,7 +375,7 @@ Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = RTL.f Proof. intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto. erewrite preserv_fnsig; eauto. -Qed. +Admitted. Lemma transf_initial_states s1: RTL.initial_state prog s1 -> @@ -525,8 +538,8 @@ Proof. - (* Bgoto *) inv MSS2. inversion MIB; subst; try inv H4. remember H2 as ODUPLIC; clear HeqODUPLIC. - eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - apply DMC in H2 as [ib [FNC MI]]; clear DMC. + exploit dupmap_correct; eauto. + intros [ib [FNC MI]]. eexists; left; eexists; split; eauto. repeat econstructor; eauto. apply iblock_istep_run_equiv in BTL_RUN; eauto. @@ -584,9 +597,8 @@ Proof. * repeat (econstructor; eauto). eapply transf_fundef_correct; eauto. + (* Bbuiltin *) - eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - remember H1 as ODUPLIC; clear HeqODUPLIC. - apply DMC in H1 as [ib [FNC MI]]; clear DMC. + exploit dupmap_correct; eauto. + intros [ib [FNC MI]]. exists (Some (normRTL (entry ib))); left; eexists; split; eauto. econstructor; eauto. econstructor. eexists; eexists; split. @@ -598,9 +610,8 @@ Proof. + (* Bjumptable *) exploit list_nth_z_rev_dupmap; eauto. intros (pc'0 & LNZ & DM). - eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - remember DM as ODUPLIC; clear HeqODUPLIC. - apply DMC in DM as [ib [FNC MI]]; clear DMC. + exploit dupmap_correct; eauto. + intros [ib [FNC MI]]. exists (Some (normRTL (entry ib))); left; eexists; split; eauto. econstructor; eauto. econstructor. eexists; eexists; split. @@ -687,11 +698,10 @@ Proof. inv H. + (* Internal function *) inv TRANSF. - rename H1 into TRANSF. - eapply dupmap_entrypoint in TRANSF as ENTRY. - eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - apply DMC in ENTRY as DMC'. - destruct DMC' as [ib [CENTRY MI]]; clear DMC. + rename H0 into TRANSF. + exploit dupmap_entrypoint; eauto. intros ENTRY. + exploit dupmap_correct; eauto. + intros [ib [CENTRY MI]]. exists (Some (normRTL (entry ib))); left; eexists; split. * eapply exec_function_internal. erewrite preserv_fnstacksize; eauto. @@ -705,16 +715,15 @@ Proof. * econstructor; eauto. - (* Returnstate *) inv H. inv STACKS. inv H1. - eapply dupmap_correct in TRANSF as DMC. unfold match_cfg in DMC. - remember DUPLIC as ODUPLIC; clear HeqODUPLIC. - apply DMC in DUPLIC as [ib [FNC MI]]; clear DMC. + exploit dupmap_correct; eauto. + intros [ib [FNC MI]]. eexists; left; eexists; split; eauto. eapply exec_return. Qed. Local Hint Resolve plus_one star_refl: core. -Theorem transf_program_correct: +Theorem transf_program_correct_cfg: forward_simulation (RTL.semantics prog) (BTLmatchRTL.cfgsem tprog). Proof. eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=cfgsem tprog) (ltof _ omeasure) match_states). @@ -729,4 +738,12 @@ Proof. - eapply senv_preserved. Qed. +Theorem transf_program_correct: + forward_simulation (RTL.semantics prog) (BTL.fsem tprog). +Proof. + eapply compose_forward_simulations. + - eapply transf_program_correct_cfg. + - eapply cfgsem2fsem. +Admitted. + End BTL_SIMULATES_RTL. -- cgit From 5b67f8284c3a98581f4da9b065a738fc534480c4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 15:56:44 +0200 Subject: archi pour la verif du scheduler --- scheduling/BTL_SEtheory.v | 24 +++++++++++++----------- scheduling/BTL_Scheduler.v | 40 ++++++++++++++++++++++++++++++++++++++++ scheduling/BTL_Schedulerproof.v | 29 +++++++++++++++++++++++++++++ 3 files changed, 82 insertions(+), 11 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 5b15fe9b..ea7082a9 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1020,7 +1020,10 @@ Proof. inversion SEXEC. Qed. -(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *) +(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] + +TODO: à revoir complètement. Il faut passer le dupmap en paramètre et match les états symboliques modulo le dupmap. +*) Record simu_proof_context := Sctx { sge1: BTL.genv; @@ -1028,24 +1031,23 @@ Record simu_proof_context := Sctx { sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; sstk1: list BTL.stackframe; sstk2: list BTL.stackframe; - sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; - sf1: BTL.function; - sf2: BTL.function; + sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; (* REM: equiv_stackframe n'est pas suffisant, il faut le dupmap dedans ! *) ssp: val; srs0: regset; sm0: mem }. -Definition bctx1 (ctx: simu_proof_context):= Bctx ctx.(sge1) ctx.(sstk1) ctx.(sf1) ctx.(ssp) ctx.(srs0) ctx.(sm0). -Definition bctx2 (ctx: simu_proof_context):= Bctx ctx.(sge2) ctx.(sstk2) ctx.(sf2) ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx1 f1 (ctx: simu_proof_context):= Bctx ctx.(sge1) ctx.(sstk1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx2 f2 (ctx: simu_proof_context):= Bctx ctx.(sge2) ctx.(sstk2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). -Definition sstate_simu (ctx: simu_proof_context) (st1 st2: sstate) := - forall t s1, sem_sstate (bctx1 ctx) t s1 st1 -> - exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2. +(* TODO: A REVOIR ! *) +Definition sstate_simu f1 f2 (ctx: simu_proof_context) (st1 st2: sstate) := + forall t s1, sem_sstate (bctx1 f1 ctx) t s1 st1 -> + exists s2, sem_sstate (bctx2 f2 ctx) t s2 st2 /\ equiv_state s1 s2. -Definition symbolic_simu ctx ib1 ib2: Prop := sstate_simu ctx (sexec (sf1 ctx) ib1) (sexec (sf2 ctx) ib2). +Definition symbolic_simu f1 f2 ctx ib1 ib2: sstate_simu f1 f2 ctx (sexec (f1 ctx) ib1) (sexec (f2 ctx) ib2). -Theorem symbolic_simu_correct ctx ib1 ib2: +Theorem symbolic_simu_correct f1 f2 ctx ib1 ib2: symbolic_simu ctx ib1 ib2 -> forall t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) (sf1 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) (sf2 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index e69de29b..8cf9635c 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -0,0 +1,40 @@ +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import RTL Op Registers OptionMonad BTL. + +Require Import Errors Linking BTL_SEtheory. + +(** External oracle *) +Axiom scheduler: BTL.function -> BTL.code * node * (PTree.t node). + +(* FAKE: A REVOIR avec BTL_SEtheory... *) +Definition symbolic_simu (dupmap: PTree.t node) (f1 f2: BTL.function) (pc1 pc2: node): Prop := True. + +Record match_function (dupmap: PTree.t node) (f1 f2: BTL.function): Prop := { + preserv_fnsig: fn_sig f1 = fn_sig f2; + preserv_fnparams: fn_params f1 = fn_params f2; + preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2; + preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> symbolic_simu dupmap f1 f2 pc1 pc2; +}. + +Definition verified_scheduler (f: BTL.function) := + let (tcte, dupmap) := scheduler f in + let (tc, te) := tcte in + let f' := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in + OK (f', dupmap). (* TODO: fixme *) + +Program Definition transf_function (f: BTL.function): + { r : res BTL.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := + match (verified_scheduler f) with + | Error e => Error e + | OK (tf, dm) => OK tf + end. +Next Obligation. +Admitted. + +Definition transf_fundef (f: fundef) : res fundef := + transf_partial_fundef (fun f => proj1_sig (transf_function f)) f. + +Definition transf_program (p: program) : res program := + transform_partial_program transf_fundef p. diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index e69de29b..049dc6b1 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -0,0 +1,29 @@ +Require Import AST Linking Values Maps Globalenvs Smallstep Registers. +Require Import Coqlib Maps Events Errors Op. +Require Import RTL BTL BTL_SEtheory. +Require Import BTL_Scheduler. + +Definition match_prog (p tp: program) := + match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp. + +Lemma transf_program_match: + forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog. +Proof. + intros. eapply match_transform_partial_program_contextual; eauto. +Qed. + +Section PRESERVATION. + +Variable prog: program. +Variable tprog: program. + +Hypothesis TRANSL: match_prog prog tprog. + +Let pge := Genv.globalenv prog. +Let tpge := Genv.globalenv tprog. + +Theorem transf_program_correct: + forward_simulation (fsem prog) (fsem tprog). +Admitted. + +End PRESERVATION. -- cgit From 2311cb788e6dcf0103266fd6c6aa76096c283e34 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 16:33:23 +0200 Subject: remove dupmap from BTL_Scheduler ! --- scheduling/BTL_SEtheory.v | 24 ++++++++++-------------- scheduling/BTL_Scheduler.v | 32 ++++++++++---------------------- scheduling/BTL_Schedulerproof.v | 2 ++ 3 files changed, 22 insertions(+), 36 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index ea7082a9..3083ca71 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1020,10 +1020,7 @@ Proof. inversion SEXEC. Qed. -(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] - -TODO: à revoir complètement. Il faut passer le dupmap en paramètre et match les états symboliques modulo le dupmap. -*) +(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *) Record simu_proof_context := Sctx { sge1: BTL.genv; @@ -1031,7 +1028,7 @@ Record simu_proof_context := Sctx { sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; sstk1: list BTL.stackframe; sstk2: list BTL.stackframe; - sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; (* REM: equiv_stackframe n'est pas suffisant, il faut le dupmap dedans ! *) + sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; ssp: val; srs0: regset; sm0: mem @@ -1040,24 +1037,23 @@ Record simu_proof_context := Sctx { Definition bctx1 f1 (ctx: simu_proof_context):= Bctx ctx.(sge1) ctx.(sstk1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). Definition bctx2 f2 (ctx: simu_proof_context):= Bctx ctx.(sge2) ctx.(sstk2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). -(* TODO: A REVOIR ! *) Definition sstate_simu f1 f2 (ctx: simu_proof_context) (st1 st2: sstate) := forall t s1, sem_sstate (bctx1 f1 ctx) t s1 st1 -> exists s2, sem_sstate (bctx2 f2 ctx) t s2 st2 /\ equiv_state s1 s2. -Definition symbolic_simu f1 f2 ctx ib1 ib2: sstate_simu f1 f2 ctx (sexec (f1 ctx) ib1) (sexec (f2 ctx) ib2). +Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall ctx, sstate_simu f1 f2 ctx (sexec f1 ib1) (sexec f2 ib2). -Theorem symbolic_simu_correct f1 f2 ctx ib1 ib2: - symbolic_simu ctx ib1 ib2 -> - forall t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) (sf1 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> - exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) (sf2 ctx) (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. +Theorem symbolic_simu_correct f1 f2 ib1 ib2: + symbolic_simu f1 f2 ib1 ib2 -> + forall ctx t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> + exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. Proof. unfold symbolic_simu, sstate_simu. - intros SIMU t s1 STEP1. - exploit (sexec_correct (bctx1 ctx)); simpl; eauto. + intros SIMU ctx t s1 STEP1. + exploit (sexec_correct (bctx1 f1 ctx)); simpl; eauto. intros; exploit SIMU; eauto. intros (s2 & SEM1 & EQ1). - exploit (sexec_exact (bctx2 ctx)); simpl; eauto. + exploit (sexec_exact (bctx2 f2 ctx)); simpl; eauto. intros (s3 & STEP2 & EQ2). clear STEP1; eexists; split; eauto. eapply equiv_state_trans; eauto. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 8cf9635c..faeb58cd 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -5,36 +5,24 @@ Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking BTL_SEtheory. (** External oracle *) -Axiom scheduler: BTL.function -> BTL.code * node * (PTree.t node). +Axiom scheduler: BTL.function -> BTL.code. -(* FAKE: A REVOIR avec BTL_SEtheory... *) -Definition symbolic_simu (dupmap: PTree.t node) (f1 f2: BTL.function) (pc1 pc2: node): Prop := True. - -Record match_function (dupmap: PTree.t node) (f1 f2: BTL.function): Prop := { +(* a specification of the verification to do on each function *) +Record match_function (f1 f2: BTL.function): Prop := { preserv_fnsig: fn_sig f1 = fn_sig f2; preserv_fnparams: fn_params f1 = fn_params f2; preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2; - preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> symbolic_simu dupmap f1 f2 pc1 pc2; + preserv_entrypoint: fn_entrypoint f1 = fn_entrypoint f2; + 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); }. -Definition verified_scheduler (f: BTL.function) := - let (tcte, dupmap) := scheduler f in - let (tc, te) := tcte in - let f' := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - OK (f', dupmap). (* TODO: fixme *) - -Program Definition transf_function (f: BTL.function): - { r : res BTL.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := - match (verified_scheduler f) with - | Error e => Error e - | OK (tf, dm) => OK tf - end. -Next Obligation. -Admitted. +Definition transf_function (f: BTL.function) := + (* TODO: fixme *) + OK (BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f)). Definition transf_fundef (f: fundef) : res fundef := - transf_partial_fundef (fun f => proj1_sig (transf_function f)) f. + transf_partial_fundef (fun f => transf_function f) f. Definition transf_program (p: program) : res program := transform_partial_program transf_fundef p. diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 049dc6b1..c8c2a76f 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -24,6 +24,8 @@ Let tpge := Genv.globalenv tprog. Theorem transf_program_correct: forward_simulation (fsem prog) (fsem tprog). +Proof. + eapply forward_simulation_step with equiv_state. (* lock-step with respect to equiv_states *) Admitted. End PRESERVATION. -- cgit From a97cea76d531b3ea985c284d76baa6370f7dc489 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 16:49:31 +0200 Subject: declare a checker for the symbolic simulation --- scheduling/BTL_Scheduler.v | 15 +++++++++++++-- 1 file changed, 13 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index faeb58cd..43d6dd5e 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -17,9 +17,20 @@ Record match_function (f1 f2: BTL.function): Prop := { exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2); }. +Local Open Scope error_monad_scope. + +Definition check_symbolic_simu (f tf: BTL.function): res unit := OK tt. (* TODO: fixme *) + +Lemma check_symbolic_simu_correct x f1 f2: + check_symbolic_simu f1 f2 = OK x -> + 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). +Admitted. + Definition transf_function (f: BTL.function) := - (* TODO: fixme *) - OK (BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f)). + let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in + do _ <- check_symbolic_simu f tf; + OK tf. Definition transf_fundef (f: fundef) : res fundef := transf_partial_fundef (fun f => transf_function f) f. -- cgit From 9208b0c21262839184281d9cc3bdf1e6dca7a416 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 28 May 2021 18:03:45 +0200 Subject: maj roadmap --- scheduling/BTLroadmap.md | 68 +++++++++++++++++++++++------------------------- 1 file changed, 32 insertions(+), 36 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 0cbcd7d6..815aad96 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -51,7 +51,7 @@ Two possibility of branch duplications (e.g tail-duplication, loop unrolling, et Branch factorization should also be possible in BTL -> RTL pass. Example: revert "loop-unrolling". -**IMPLEM NOTE:** in a first step a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. +**IMPLEM NOTE:** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes. **CURRENT STATUS** @@ -64,12 +64,6 @@ Branch factorization should also be possible in BTL -> RTL pass. Example: revert - BTL -> RTL: TODO. - RTL -> BTL: started. -**TODO** - -- lien BTL/RTL: autoriser le BTL à avoir des Bnop en plus du RTL, e.g. pour autoriser des "if-then" sans else. -Pour faciliter le vérificateur, faire comme le Bgoto: les Bnop en "trop" ne sont autorisés que s'il y a eu une instruction RTL avant. -Ajouter aussi un booléen (positionné par l'oracle) sur le Bnop qui indique si le "nop" existe ou pas au niveau RTL. - ## Simulation modulo liveness and "Functional" semantics of BTL blocks L'approche suivie pour réaliser la simulation modulo liveness dans @@ -83,20 +77,20 @@ On n'a pas envie de généraliser cette usine à gaz. On cherche donc une façon plus abstraite de faire... On a l'idée de coder la "simulation modulo liveness" dans une "simulation -less_undef". Ça peut rendre le cadre du test de simulation plus +less_def". Ça peut rendre le cadre du test de simulation plus simple et plus général. -**Idée_Informelle** à côté de la sémantique "à la RTL" pour BTL, on pourrait -définir une nouvelle sémantique (basée sur la même sémantique à grand -pas dans les blocs), où c'est juste "l'entrée dans un bloc" qui change -de sémantique. Intuitivement, cette nouvelle sémantique considèrerait -chaque bloc comme une sorte de fonction paramétrée par les `input_regs` -et appelée uniquement en "tailcall" (via les "goto"). C'est ce qu'on -va appeler la "functional semantics" de BTL (l'autre étant appelée -qqchose comme "CFG semantics" ?). +**Idée_Informelle** à côté de la sémantique "à la RTL" pour BTL +[BTLmatch.cfgsem], on définit une sémantique [BTL.fsem], où c'est +juste "l'entrée dans un bloc" qui change de sémantique. Intuitivement, +cette sémantique considère chaque bloc comme une sorte de +fonction paramétrée par les `input_regs` et appelée uniquement en +"tailcall" (via les "goto"). C'est ce qu'on va appeler la "functional +semantics" de BTL (l'autre étant appelée qqchose comme "CFG semantics" +?). -Autrement dit, sur l'entrée dans un bloc pour un état (rs,m), on -commence par mettre à Vundef tous les registres qui ne sont pas dans l'`input_regs`. +Autrement dit, sur l'entrée dans un bloc pour un état (rs,m), on pourrait moralement +commencer par mettre à Vundef tous les registres qui ne sont pas dans l'`input_regs`. **NOTE** cette idée de voir les blocs comme des "fonctions" correpond bien à la représentation "SSA" à la Appel/MLIR. Donc cette sémantique @@ -107,9 +101,14 @@ paramètres. **REM** pour le test d'exécution symbolique, il semble plus adapté de mettre les Vundef juste à la fin de la transition (précédente) plutôt -qu'au début (de la suivante). On pourrait aussi faire les deux (dans le détail, ça ne ferait pas exactement deux fois la même chose: -par exemple, quand on sort sur un call ou un builtin, le résultat du call/builtin n'est jamais dans le liveout puisqu'il va être écrasé de toute façon pendant la transition, -mais il peut être -- ou pas -- dans le livein de la transition suivante dans ce même bloc). +qu'au début (de la suivante): c'est d'ailleurs plus proche de la vision SSA. + +En fait, on pourrait faire les deux (dans le détail, ça ne ferait pas +exactement deux fois la même chose: par exemple, quand on sort sur un +call ou un builtin, le résultat du call/builtin n'est jamais dans le +liveout puisqu'il va être écrasé de toute façon pendant la transition, +mais il peut être -- ou pas -- dans le livein de la transition +suivante dans ce même bloc). Avoir une initialisation à Vundef en début de bloc permettrait de faire une analyse des expressions mal initialisées - par exemple à @@ -117,24 +116,20 @@ partir du bloc d'entrée de la fonction. Ça semble complémentaire de l'analyse de "liveout". Mais utilisable uniquement dans le cadre d'une combinaison "exécution symbolique"+"value analysis" (donc pas tout de suite). -Donc, dans un premier temps, la "functional semantics" encode/abstrait la notion de "liveout" pour le test d'exécution symbolique. -Les définitions des deux sémantiques (symbolique et "functional") se font donc en simultanée. +Donc, dans un premier temps, la BTL.fsem encode/abstrait la notion de +"liveout" pour le test d'exécution symbolique. Les définitions des +deux sémantiques (symbolique et "functional") se font donc en +simultanée. **STATUS** -1. See [BTL.fsem] (Pour l'instant CFGSemantics => BTL.semantics et FunctionalSemantics => BTL.fsem) - -**TODO** - -1. prouver la simulation BTL.FunctionalSemantics -> BTL.CFGSemantics (facile normalement: c'est du "less_undef lockstep"). -2. implémenter une passe BTL.CFGSemantics -> BTL.FunctionalSemantics. - en gros: on implémente un vérificateur des infos de liveness. - c'est la correction du "input_regs" qui garantit que la simulation est correct. - La preuve devrait normalement être très similaire à RTLpathLivegenproof. +1. See [BTL.fsem] +2. fsem -> cfgsem: fait (ou presque). +3. cfgsem -> fsem: a faire (via verif de liveness). ## "Basic" symbolic execution / symbolic simulation -We will implement a "basic" (e.g without rewriting rules) simulation test for BTL.FunctionalSemantics. +We will implement a "basic" (e.g without rewriting rules) simulation test for BTL.fsem. Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter de tester l'égalité des registres de chaque côté: les registres hors liveout seront égaux à Vundef de chaque côté. **STATUS** @@ -353,15 +348,16 @@ Extends the symbolic simulation test to support invariants at block entry, inclu **IDEAS** -- En pratique, il faudra peut-être affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. +- En pratique, il faudrait affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. +Une solution possible: calquer Bjumptable sur Bcond (c-a-d autoriser les Bjumptable en milieu de blocs). On pourrait étendre la prédiction de branchement par profiling aux jumptables (par exemple, avoir des superblocks avec des jumptables au milieu). Un autre intérêt: simplifier le support pour des formes SSA partielles (cf. ci-dessous). - Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). ## Support of SSA-optimizations -Extends BTL with "register renamings" at exits. +Minimum feature: extends BTL with "register renamings" at exits. This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. -This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. +Maximum feature: add a basic instruction that performs parallel renaming of registers. If we also support Bjumptable in the middle of a block (see above), this simple feature would suffice to represent a very general notion of "partial SSA forms": since they could appear in the middle of a block, or just before an exit (such a parallel assignment would be forbidden in BTL<->RTL matching). ## Alias analysis in the symbolic simulation -- cgit From e6a1df51a2a3d29c58d72453355e50a979e86297 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 29 May 2021 08:10:07 +0200 Subject: BTLroadmap: jumptable --- scheduling/BTLroadmap.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 815aad96..954143c1 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -349,7 +349,7 @@ Extends the symbolic simulation test to support invariants at block entry, inclu **IDEAS** - En pratique, il faudrait affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. -Une solution possible: calquer Bjumptable sur Bcond (c-a-d autoriser les Bjumptable en milieu de blocs). On pourrait étendre la prédiction de branchement par profiling aux jumptables (par exemple, avoir des superblocks avec des jumptables au milieu). Un autre intérêt: simplifier le support pour des formes SSA partielles (cf. ci-dessous). +Une solution possible: calquer Bjumptable sur Bcond (c-a-d autoriser les Bjumptable en milieu de blocs). On pourrait étendre la prédiction de branchement par profiling aux jumptables (par exemple, avoir des superblocks avec des jumptables au milieu). Un autre intérêt: simplifier (un peu) le support pour des formes SSA partielles (cf. ci-dessous). Ceci dit, ça compliquerait les choses à plein d'endroits (type Coq [iblock] mutuellement inductif avec [list_iblock] pour avoir les bons principes d'inductions, etc) pour des gains minimes en pratiques ? - Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). -- cgit From 271f87ba08f42340900378c0797511d4071fc1b8 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 31 May 2021 16:55:18 +0200 Subject: BTL Scheduler oracle and some drafts --- scheduling/BTLScheduleraux.ml | 6 +++++- scheduling/BTL_Livecheck.v | 47 +++++++++++++++++++++++++++++++++++++++++-- scheduling/BTL_Scheduler.v | 2 ++ scheduling/BTLmatchRTL.v | 42 ++++++++++++++++++++++++++++++++++---- scheduling/RTLtoBTLaux.ml | 6 +----- 5 files changed, 91 insertions(+), 12 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index ad0c307d..b87636e1 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -4,6 +4,7 @@ open Registers open BTL open BTLtypes open DebugPrint +open PrintBTL open RTLcommonaux open InstructionScheduler open PrepassSchedulingOracleDeps @@ -243,8 +244,11 @@ let rec do_schedule btl = function let btl' = schedule_blk n ibf btl in do_schedule btl' blks -let btl_scheduler btl = +let btl_scheduler f = + let btl = f.fn_code in (*debug_flag := true;*) let btl' = do_schedule btl (PTree.elements btl) in + debug "Scheduled BTL Code:\n"; + print_btl_code stderr btl'; (*debug_flag := false;*) btl' diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 2baba025..6064e4cd 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -4,8 +4,51 @@ Require Import Errors RTL BTL BTLmatchRTL. (** TODO: adapt stuff RTLpathLivegen *) -Definition liveness_checker (f: BTL.function): res unit := OK tt. (* TODO: implement this ! *) - +Local Open Scope lazy_bool_scope. + +Local Open Scope option_monad_scope. + +Fixpoint list_mem (rl: list reg) (alive: Regset.t) {struct rl}: bool := + match rl with + | nil => true + | r1 :: rs => Regset.mem r1 alive &&& list_mem rs alive + end. + +(* +Fixpoint iblock_checker (btl: code) (ib: iblock) (alive: Regset.t) : option Regset.t := + match ib with + | Bseq ib1 ib2 => + SOME alive' <- iblock_checker btl ib1 alive IN + iblock_checker btl ib1 alive' (Some ib2) + | Bnop _ => + iblock_checker btl alive + | Bop _ args dest _ => + ASSERT list_mem args alive IN + iblock_checker btl (Regset.add dest alive) + | Bload _ _ _ args dest _ => + ASSERT list_mem args alive IN + iblock_checker btl (Regset.add dest alive) + | Bstore _ _ args src _ => + ASSERT Regset.mem src alive IN + ASSERT list_mem args alive IN + iblock_checker btl alive + | Bcond _ args (BF (Bgoto s) _) ib2 _ => + ASSERT list_mem args alive IN + SOME _ <- iblock_checker btl ib2 alive None IN + SOME next <- btl!s IN + ASSERT Regset.subset next.(input_regs) alive IN + iblock_checker btl next.(entry) next.(input_regs) None + | Bcond _ _ _ _ _ => None + | BF _ _ => Some tt + end. + +Definition liveness_checker (f: BTL.function): res unit := + match f.(fn_code)!(f.(fn_entrypoint)) with + | Some ibf => iblock_checker f.(fn_code) ibf.(entry) ibf.(input_regs) + | None => Error (msg "liveness_checker: empty function") + end. + *) +Definition liveness_checker (f: BTL.function): res unit := OK tt. Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. (** TODO: adapt stuff RTLpathLivegenproof *) diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 43d6dd5e..b479204c 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -7,6 +7,8 @@ Require Import Errors Linking BTL_SEtheory. (** External oracle *) Axiom scheduler: BTL.function -> BTL.code. +Extract Constant scheduler => "BTLScheduleraux.btl_scheduler". + (* a specification of the verification to do on each function *) Record match_function (f1 f2: BTL.function): Prop := { preserv_fnsig: fn_sig f1 = fn_sig f2; diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index 8994579d..8081b3a6 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -30,21 +30,55 @@ Qed. Local Hint Resolve tr_inputs_lessdef init_regs_lessdef_preserv find_function_lessdef regs_lessdef_regs lessdef_state_refl: core. +Local Hint Unfold regs_lessdef: core. -Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 ib rs1' m1' of +Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 rs1' m1' of + rs2 m2 ib (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of) - rs2 m2 (REGS: forall r, Val.lessdef rs1#r rs2#r) (MEMS: Mem.extends m1 m2) : exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) /\ (forall r, Val.lessdef rs1'#r rs2'#r) /\ (Mem.extends m1' m2'). Proof. - induction ISTEP; simpl; try_simplify_someHyps. + induction ISTEP; simpl; try_simplify_someHyps; intros. + - (* Bop *) + exploit (@eval_operation_lessdef _ _ ge sp op (rs ## args)); eauto. + intros (v1 & EVAL' & LESSDEF). + do 2 eexists; rewrite EVAL'. repeat (split; eauto). + eapply set_reg_lessdef; eauto. + - (* Bload *) + exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto. + intros (v2 & EVAL' & LESSDEF). exploit Mem.loadv_extends; eauto. + intros (v3 & LOAD' & LESSDEF'). + do 2 eexists; rewrite EVAL', LOAD'. repeat (split; eauto). + eapply set_reg_lessdef; eauto. + - (* Bstore *) + exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto. + intros (v2 & EVAL' & LESSDEF). exploit Mem.storev_extends; eauto. + intros (v3 & STORE' & LESSDEF'). + do 2 eexists; rewrite EVAL', STORE'. repeat (split; eauto). + - (* Bseq stop *) + exploit IHISTEP; eauto. intros (rs2' & m2' & IBIS & REGS' & MEMS'). + destruct (iblock_istep_run _ _ _ _ _); try congruence. + destruct o, _fin; simpl in *; try congruence. eauto. + - (* Bseq continue *) + + (* TODO gourdinl induction pb? + exploit IHISTEP1; eauto. intros (rs3 & m3 & IBIS1 & REGS1 & MEMS1). + rewrite IBIS1; simpl. rewrite iblock_istep_run_equiv in ISTEP2. + exploit IHISTEP2; eauto. intros (rs4 & m4 & IBIS2 & REGS2 & MEMS2). + destruct o, _fin; simpl in *; try congruence; try_simplify_someHyps. + intros; apply IHISTEP1 in MEMS; auto. + exploit IHISTEP2; eauto. intros (rs3 & m3 & IBIS1 & REGS1 & MEMS1). + + + destruct (iblock_istep_run _ _ _ _ _) eqn:EQb1. + 2: { rewrite iblock_istep_run_equiv in ISTEP1. destruct b1; simpl in *; try congruence. + destruct o, _fin; simpl in *; try congruence; try_simplify_someHyps.*) Admitted. Local Hint Constructors lessdef_stackframe lessdef_state final_step list_forall2 step: core. -Local Hint Unfold regs_lessdef: core. Lemma fsem2cfgsem_finalstep_simu ge stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 (FSTEP: final_step tr_inputs ge stk1 f sp rs1 m1 fin t s1) diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index d04326ea..e709d846 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -6,7 +6,6 @@ open RTLcommonaux open DebugPrint open BTLtypes open BTLcommonaux -open BTLScheduleraux let encaps_final inst osucc = match inst with @@ -102,19 +101,16 @@ let rtl2btl (f : RTL.coq_function) = let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in (* TODO gourdinl move elsewhere *) - let btl' = btl_scheduler btl in debug "Entry %d\n" (p2i entry); (*debug_flag := true;*) debug "RTL Code: "; print_code code; debug "BTL Code:\n"; print_btl_code stderr btl; - debug "Scheduled BTL Code:\n"; - print_btl_code stderr btl'; (*debug_flag := false;*) debug "Dupmap:\n"; print_ptree print_pint dm; debug "Join points: "; print_true_nodes join_points; debug "\n"; - ((btl', entry), dm) + ((btl, entry), dm) -- cgit From 16859dc21531591f575bcfe747f520c334f418f9 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 31 May 2021 17:57:01 +0200 Subject: prove fsem2cfgsem_ibistep_simu --- scheduling/BTLmatchRTL.v | 32 +++++++++++++------------------- 1 file changed, 13 insertions(+), 19 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index 8081b3a6..439ba9cc 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -32,12 +32,12 @@ Local Hint Resolve tr_inputs_lessdef init_regs_lessdef_preserv find_function_les lessdef_state_refl: core. Local Hint Unfold regs_lessdef: core. -Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 rs1' m1' of - rs2 m2 ib - (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of) +Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 rs1' m1' of ib + (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of): forall + rs2 m2 (REGS: forall r, Val.lessdef rs1#r rs2#r) - (MEMS: Mem.extends m1 m2) - : exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) + (MEMS: Mem.extends m1 m2), + exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of) /\ (forall r, Val.lessdef rs1'#r rs2'#r) /\ (Mem.extends m1' m2'). Proof. @@ -63,20 +63,14 @@ Proof. destruct (iblock_istep_run _ _ _ _ _); try congruence. destruct o, _fin; simpl in *; try congruence. eauto. - (* Bseq continue *) - - (* TODO gourdinl induction pb? - exploit IHISTEP1; eauto. intros (rs3 & m3 & IBIS1 & REGS1 & MEMS1). - rewrite IBIS1; simpl. rewrite iblock_istep_run_equiv in ISTEP2. - exploit IHISTEP2; eauto. intros (rs4 & m4 & IBIS2 & REGS2 & MEMS2). - destruct o, _fin; simpl in *; try congruence; try_simplify_someHyps. - intros; apply IHISTEP1 in MEMS; auto. - exploit IHISTEP2; eauto. intros (rs3 & m3 & IBIS1 & REGS1 & MEMS1). - - - destruct (iblock_istep_run _ _ _ _ _) eqn:EQb1. - 2: { rewrite iblock_istep_run_equiv in ISTEP1. destruct b1; simpl in *; try congruence. - destruct o, _fin; simpl in *; try congruence; try_simplify_someHyps.*) -Admitted. + exploit IHISTEP1; eauto. + clear ISTEP1 REGS MEMS. + intros (rs3 & m3 & ISTEP3 & REGS3 & MEMS3). + rewrite ISTEP3; simpl. rewrite iblock_istep_run_equiv in ISTEP2. + exploit IHISTEP2; eauto. + - (* Bcond *) + erewrite (@eval_condition_lessdef cond (rs ## args)); eauto. +Qed. Local Hint Constructors lessdef_stackframe lessdef_state final_step list_forall2 step: core. -- cgit From 78fbb0e6f4c2065460a0ddb7e2e3ec94da21169f Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 31 May 2021 23:51:31 +0200 Subject: Dupmap bugfix and some advance in Livegen --- scheduling/BTLRenumber.ml | 7 ++-- scheduling/BTL_Livecheck.v | 99 ++++++++++++++++++++++++++++++++++++---------- scheduling/BTLtoRTLaux.ml | 10 +++-- scheduling/RTLtoBTLaux.ml | 2 +- 4 files changed, 89 insertions(+), 29 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLRenumber.ml b/scheduling/BTLRenumber.ml index 58d4f7ac..6ff42a27 100644 --- a/scheduling/BTLRenumber.ml +++ b/scheduling/BTLRenumber.ml @@ -87,6 +87,7 @@ let regenerate_btl_tree btl entry = | Bseq (ib1, ib2) -> Bseq (renumber_iblock ib1, renumber_iblock ib2) | _ -> ib in + let dm = ref PTree.empty in let ord_btl = PTree.fold (fun ord_btl old_n old_ibf -> @@ -96,16 +97,16 @@ let regenerate_btl_tree btl entry = let bi = mk_binfo n in let ibf = { entry = ib; input_regs = old_ibf.input_regs; binfo = bi } in if old_n = entry then new_entry := n_pos; + dm := PTree.set old_n n_pos !dm; PTree.set n_pos ibf ord_btl) btl PTree.empty in debug "Renumbered BTL with new_entry=%d:\n" (p2i !new_entry); print_btl_code stderr ord_btl; - (ord_btl, !new_entry) + ((ord_btl, !new_entry), !dm) let renumber btl entry = (*debug_flag := true;*) let btl' = recompute_inumbs btl entry in - let ord_btl, new_entry = regenerate_btl_tree btl' entry in (*debug_flag := false;*) - (ord_btl, new_entry) + regenerate_btl_tree btl' entry diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 6064e4cd..4d556711 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -14,41 +14,98 @@ Fixpoint list_mem (rl: list reg) (alive: Regset.t) {struct rl}: bool := | r1 :: rs => Regset.mem r1 alive &&& list_mem rs alive end. -(* -Fixpoint iblock_checker (btl: code) (ib: iblock) (alive: Regset.t) : option Regset.t := +Definition reg_option_mem (or: option reg) (alive: Regset.t) := + match or with None => true | Some r => Regset.mem r alive end. + +Definition reg_sum_mem (ros: reg + ident) (alive: Regset.t) := + match ros with inl r => Regset.mem r alive | inr s => true end. + +(* NB: definition following [regmap_setres] in [RTL.step] semantics *) +Definition reg_builtin_res (res: builtin_res reg) (alive: Regset.t): Regset.t := + match res with + | BR r => Regset.add r alive + | _ => alive + end. + +Definition exit_checker {A} (btl: code) (alive: Regset.t) (s: node) (v: A): option A := + SOME next <- btl!s IN + ASSERT Regset.subset next.(input_regs) alive IN + Some v. + +Fixpoint exit_list_checker (btl: code) (alive: Regset.t) (l: list node): bool := + match l with + | nil => true + | s :: l' => exit_checker btl alive s true &&& exit_list_checker btl alive l' + end. + +Definition final_inst_checker (btl: code) (alive: Regset.t) (fin: final): option unit := + match fin with + | Bgoto s => + exit_checker btl alive s tt + | Breturn oreg => + ASSERT reg_option_mem oreg alive IN Some tt + | Bcall _ ros args res s => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN + exit_checker btl (Regset.add res alive) s tt + | Btailcall _ ros args => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN Some tt + | Bbuiltin _ args res s => + ASSERT list_mem (params_of_builtin_args args) alive IN + exit_checker btl (reg_builtin_res res alive) s tt + | Bjumptable arg tbl => + ASSERT Regset.mem arg alive IN + ASSERT exit_list_checker btl alive tbl IN Some tt + end. + +Definition is_none {A} (oa: option A): bool := + match oa with + | Some _ => false + | None => true + end. + +Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option (Regset.t*option final) := match ib with | Bseq ib1 ib2 => - SOME alive' <- iblock_checker btl ib1 alive IN - iblock_checker btl ib1 alive' (Some ib2) - | Bnop _ => - iblock_checker btl alive + SOME res <- body_checker btl ib1 alive IN + ASSERT is_none (snd res) IN + body_checker btl ib2 (fst res) + | Bnop _ => Some (alive, None) | Bop _ args dest _ => ASSERT list_mem args alive IN - iblock_checker btl (Regset.add dest alive) + Some (Regset.add dest alive, None) | Bload _ _ _ args dest _ => ASSERT list_mem args alive IN - iblock_checker btl (Regset.add dest alive) + Some (Regset.add dest alive, None) | Bstore _ _ args src _ => ASSERT Regset.mem src alive IN ASSERT list_mem args alive IN - iblock_checker btl alive - | Bcond _ args (BF (Bgoto s) _) ib2 _ => + Some (alive, None) + | Bcond _ args (BF (Bgoto s) _) (Bnop None) _ => ASSERT list_mem args alive IN - SOME _ <- iblock_checker btl ib2 alive None IN - SOME next <- btl!s IN - ASSERT Regset.subset next.(input_regs) alive IN - iblock_checker btl next.(entry) next.(input_regs) None - | Bcond _ _ _ _ _ => None - | BF _ _ => Some tt + exit_checker btl alive s (alive, None) + | BF fin _ => Some (alive, Some fin) + | _ => None + end. + +Definition iblock_checker (btl: code) (ib: iblock) (alive: Regset.t) : option unit := + SOME res <- body_checker btl ib alive IN + SOME fin <- snd res IN + final_inst_checker btl (fst res) fin. + +Fixpoint list_iblock_checker (btl: code) (l: list (node*iblock_info)): bool := + match l with + | nil => true + | (_, ibf) :: l' => iblock_checker btl ibf.(entry) ibf.(input_regs) &&& list_iblock_checker btl l' end. Definition liveness_checker (f: BTL.function): res unit := - match f.(fn_code)!(f.(fn_entrypoint)) with - | Some ibf => iblock_checker f.(fn_code) ibf.(entry) ibf.(input_regs) - | None => Error (msg "liveness_checker: empty function") + match list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)) with + | true => OK tt + | false => Error (msg "BTL_Livecheck: liveness_checker failed") end. - *) -Definition liveness_checker (f: BTL.function): res unit := OK tt. + Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. (** TODO: adapt stuff RTLpathLivegenproof *) diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml index 8e728bd2..ddec991d 100644 --- a/scheduling/BTLtoRTLaux.ml +++ b/scheduling/BTLtoRTLaux.ml @@ -73,14 +73,16 @@ let btl2rtl (f : BTL.coq_function) = (*debug_flag := true;*) let btl = f.fn_code in let entry = f.fn_entrypoint in - let ordered_btl, new_entry = renumber btl entry in + let obne, dm = renumber btl entry in + let ordered_btl, new_entry = obne in let rtl = translate_function ordered_btl new_entry in - let dm = PTree.map (fun n i -> n) btl in - debug "Entry %d\n" (p2i entry); + debug "Entry %d\n" (p2i new_entry); + debug "BTL Code:\n"; + print_btl_code stderr ordered_btl; debug "RTL Code: "; print_code rtl; debug "Dupmap:\n"; print_ptree print_pint dm; debug "\n"; (*debug_flag := false;*) - ((rtl, entry), dm) + ((rtl, new_entry), dm) diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index e709d846..3de82d96 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -101,8 +101,8 @@ let rtl2btl (f : RTL.coq_function) = let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in (* TODO gourdinl move elsewhere *) - debug "Entry %d\n" (p2i entry); (*debug_flag := true;*) + debug "Entry %d\n" (p2i entry); debug "RTL Code: "; print_code code; debug "BTL Code:\n"; -- cgit From 84ddcd68504461bb2dee30ac38f0bb495b8aec38 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 1 Jun 2021 12:11:50 +0200 Subject: Lemma list_iblock_checker_correct --- scheduling/BTL_Livecheck.v | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 4d556711..62833dc3 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -100,6 +100,26 @@ Fixpoint list_iblock_checker (btl: code) (l: list (node*iblock_info)): bool := | (_, ibf) :: l' => iblock_checker btl ibf.(entry) ibf.(input_regs) &&& list_iblock_checker btl l' end. +Lemma lazy_and_Some_true A (o: option A) (b: bool): o &&& b = true <-> (exists v, o = Some v) /\ b = true. +Proof. + destruct o; simpl; intuition. + - eauto. + - firstorder. try_simplify_someHyps. +Qed. + +Lemma lazy_and_Some_tt_true (o: option unit) (b: bool): o &&& b = true <-> o = Some tt /\ b = true. +Proof. + intros; rewrite lazy_and_Some_true; firstorder. + destruct x; auto. +Qed. + +Lemma list_iblock_checker_correct btl l: + list_iblock_checker btl l = true -> forall e, List.In e l -> iblock_checker btl (snd e).(entry) (snd e).(input_regs) = Some tt. +Proof. + intros CHECKER e H; induction l as [|(n & ibf) l]; intuition. + simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto). +Qed. + Definition liveness_checker (f: BTL.function): res unit := match list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)) with | true => OK tt -- cgit From 9b74954477b9b1e5b75fbdd15215fc09941f72e4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 1 Jun 2021 12:18:51 +0200 Subject: Lemma liveness_checker_correct --- scheduling/BTL_Livecheck.v | 12 ++++++++++++ 1 file changed, 12 insertions(+) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 62833dc3..6015d91f 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -126,6 +126,18 @@ Definition liveness_checker (f: BTL.function): res unit := | false => Error (msg "BTL_Livecheck: liveness_checker failed") end. +Lemma liveness_checker_correct f n ibf: + liveness_checker f = OK tt -> + f.(fn_code)!n = Some ibf -> + iblock_checker f.(fn_code) ibf.(entry) ibf.(input_regs) = Some tt. +Proof. + unfold liveness_checker. + destruct list_iblock_checker eqn:EQL; try congruence. + intros _ P. exploit list_iblock_checker_correct; eauto. + - eapply PTree.elements_correct; eauto. + - simpl; auto. +Qed. + Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. (** TODO: adapt stuff RTLpathLivegenproof *) -- cgit From 06e5c1dfdb60e45de6b9efa9cdb82031acf5aed2 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 1 Jun 2021 14:37:59 +0200 Subject: preparing liveness proof main theorem --- scheduling/BTL_Livecheck.v | 45 +++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 43 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 6015d91f..ca7db1da 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -140,6 +140,10 @@ Qed. Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. +Inductive liveness_ok_fundef: fundef -> Prop := + | liveness_ok_Internal f: liveness_ok_function f -> liveness_ok_fundef (Internal f) + | liveness_ok_External ef: liveness_ok_fundef (External ef). + (** TODO: adapt stuff RTLpathLivegenproof *) Local Notation ext alive := (fun r => Regset.In r alive). @@ -156,6 +160,29 @@ Qed. Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := forall r, (alive r) -> rs1#r = rs2#r. +Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := + | eqlive_stackframes_intro ibf res f sp pc rs1 rs2 + (LIVE: liveness_ok_function f) + (ENTRY: f.(fn_code)!pc = Some ibf) + (EQUIV: forall v, eqlive_reg (ext ibf.(input_regs)) (rs1 # res <- v) (rs2 # res <- v)): + eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + +Inductive eqlive_states: state -> state -> Prop := + | eqlive_states_intro + ibf st1 st2 f sp pc rs1 rs2 m + (STACKS: list_forall2 eqlive_stackframes st1 st2) + (LIVE: liveness_ok_function f) + (PATH: f.(fn_code)!pc = Some ibf) + (EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2): + eqlive_states (State st1 f sp pc rs1 m) (State st2 f sp pc rs2 m) + | eqlive_states_call st1 st2 f args m + (LIVE: liveness_ok_fundef f) + (STACKS: list_forall2 eqlive_stackframes st1 st2): + eqlive_states (Callstate st1 f args m) (Callstate st2 f args m) + | eqlive_states_return st1 st2 v m + (STACKS: list_forall2 eqlive_stackframes st1 st2): + eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). + (* continue to adapt stuff RTLpathLivegenproof *) Section FSEM_SIMULATES_CFGSEM. @@ -166,11 +193,25 @@ Let ge := Genv.globalenv prog. Hypothesis liveness_checker_correct: forall b f, Genv.find_funct_ptr ge b = Some (Internal f) -> liveness_ok_function f. -Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). +Lemma cfgsem2fsem_step_simu s1 s1' t s2 + (STEP : step tid (Genv.globalenv prog) s1 t s1') + (STATES : eqlive_states s1 s2) + :exists s2' : state, + step tr_inputs (Genv.globalenv prog) s2 t s2' /\ + eqlive_states s1' s2'. Proof. - exploit liveness_checker_correct. Admitted. +Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). +Proof. + eapply forward_simulation_step with eqlive_states; simpl; eauto. + - destruct 1, f; intros; eexists; intuition eauto; + repeat econstructor; eauto. + - intros s1 s2 r STATES FINAL; destruct FINAL. + inv STATES; inv STACKS; constructor. + - intros. eapply cfgsem2fsem_step_simu; eauto. +Qed. + End FSEM_SIMULATES_CFGSEM. -- cgit From cf9824cd02c9dd5a8053c1853f26b98ad807766e Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 1 Jun 2021 15:33:12 +0200 Subject: Some others main lemmas --- scheduling/BTL_Livecheck.v | 40 +++++++++++++++++++++++++++++++++------- 1 file changed, 33 insertions(+), 7 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index ca7db1da..49660222 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -169,11 +169,11 @@ Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := Inductive eqlive_states: state -> state -> Prop := | eqlive_states_intro - ibf st1 st2 f sp pc rs1 rs2 m + st1 st2 f sp pc rs1 rs2 m (STACKS: list_forall2 eqlive_stackframes st1 st2) - (LIVE: liveness_ok_function f) - (PATH: f.(fn_code)!pc = Some ibf) - (EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2): + (LIVE: liveness_ok_function f): + (*(PATH: f.(fn_code)!pc = Some ibf)*) + (*(EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2):*) eqlive_states (State st1 f sp pc rs1 m) (State st2 f sp pc rs2 m) | eqlive_states_call st1 st2 f args m (LIVE: liveness_ok_fundef f) @@ -191,7 +191,19 @@ Variable prog: BTL.program. Let ge := Genv.globalenv prog. -Hypothesis liveness_checker_correct: forall b f, Genv.find_funct_ptr ge b = Some (Internal f) -> liveness_ok_function f. +Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. + +Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m1 rs2 m2 ib s1 t + (STEP : iblock_step tid (Genv.globalenv prog) stk1 f sp rs1 m1 ib t s1) + (STACKS : list_forall2 eqlive_stackframes stk1 stk2) + (LIVE : liveness_ok_function f) + :exists s2 : state, + iblock_step tr_inputs (Genv.globalenv prog) stk2 f sp rs2 m2 ib t s2 /\ + eqlive_states s1 s2. +Proof. +Admitted. + +Local Hint Constructors step: core. Lemma cfgsem2fsem_step_simu s1 s1' t s2 (STEP : step tid (Genv.globalenv prog) s1 t s1') @@ -200,13 +212,27 @@ Lemma cfgsem2fsem_step_simu s1 s1' t s2 step tr_inputs (Genv.globalenv prog) s2 t s2' /\ eqlive_states s1' s2'. Proof. -Admitted. + destruct STEP; inv STATES. + - (* iblock *) + exploit cfgsem2fsem_ibstep_simu; eauto. + intros (s2 & STEP2 & STATES2). + eexists; split; eauto. + - (* internal call *) + inv LIVE. + eexists; split; repeat econstructor; eauto. + - (* external_call *) + eexists; split; repeat econstructor; eauto. + - (* return *) + inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst. + inv STF2. + eexists; split; econstructor; eauto. +Qed. Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). Proof. eapply forward_simulation_step with eqlive_states; simpl; eauto. - destruct 1, f; intros; eexists; intuition eauto; - repeat econstructor; eauto. + repeat (econstructor; eauto). - intros s1 s2 r STATES FINAL; destruct FINAL. inv STATES; inv STACKS; constructor. - intros. eapply cfgsem2fsem_step_simu; eauto. -- cgit From f6a3483c7dcac5ed56a349361a19dc1f4c985de3 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 1 Jun 2021 17:16:06 +0200 Subject: starting BTL_SEsimuref --- scheduling/BTL_SEsimuref.v | 317 +++++++++++++++++++++++++++++++++++++++++++++ scheduling/BTL_SEtheory.v | 100 +------------- 2 files changed, 318 insertions(+), 99 deletions(-) create mode 100644 scheduling/BTL_SEsimuref.v (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v new file mode 100644 index 00000000..af2159cb --- /dev/null +++ b/scheduling/BTL_SEsimuref.v @@ -0,0 +1,317 @@ +(** Refinement of BTL_SEtheory data-structures + in order to introduce (and prove correct) a lower-level specification of the simulation test. + + TODO: not yet with hash-consing ? Or "fake" hash-consing only ? + + Ceci est un "bac à sable". + + - On introduit une représentation plus concrète pour les types d'état symbolique [sistate] et [sstate]. + - Etant donné une spécification intuitive "*_simu" pour tester la simulation sur cette représentation des états symboliques, + on essaye déjà de trouver les bonnes notions de raffinement "*_refines" qui permette de prouver les lemmes "*_simu_correct". + - Il faudra ensuite vérifier que l'exécution symbolique préserve ces relations de raffinement ! + +*) + +Require Import Coqlib Maps Floats. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL BTL OptionMonad BTL_SEtheory. + +(** * Preservation properties *) + +(* TODO: inherited from RTLpath. Use simu_proof_context + bctx2 + bctx1 instead of all the hypotheses/variables below ? *) + +Section SymbValPreserved. + +Variable ge ge': BTL.genv. + +Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. + +Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol ge id = Senv.find_symbol ge' id. +Proof. + destruct senv_preserved_BTL as (A & B & C). congruence. +Qed. + +Lemma senv_symbol_address_preserved id ofs: + Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. +Proof. + unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. + reflexivity. +Qed. + +Variable stk stk': list stackframe. +Variable f f': function. +Variable sp: val. +Variable rs0: regset. +Variable m0: mem. + +Lemma eval_sval_preserved sv: + eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. +Proof. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) + (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma seval_builtin_arg_preserved m: forall bs varg, + seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> + seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. +Proof. + induction 1; simpl. + all: try (constructor; auto). + - rewrite <- eval_sval_preserved. assumption. + - rewrite <- senv_symbol_address_preserved. assumption. + - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. +Qed. + +Lemma seval_builtin_args_preserved m lbs vargs: + seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> + seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. +Proof. + induction 1; constructor; eauto. + eapply seval_builtin_arg_preserved; auto. +Qed. + +Lemma list_sval_eval_preserved lsv: + eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. +Proof. + induction lsv; simpl; auto. + rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma smem_eval_preserved sm: + eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. + rewrite eval_sval_preserved; auto. +Qed. + +Lemma seval_condition_preserved cond lsv sm: + seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. +Proof. + unfold seval_condition. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + +End SymbValPreserved. + +(** * ... *) + +Record sis_ok ctx (sis: sistate): Prop := { + OK_PRE: (sis.(si_pre) ctx); + OK_SMEM: eval_smem ctx sis.(si_smem) <> None; + OK_SREG: forall (r: reg), eval_sval ctx (si_sreg sis r) <> None +}. +Local Hint Resolve OK_PRE OK_SMEM OK_SREG: core. +Local Hint Constructors sis_ok: core. + +Lemma sem_sok ctx sis rs m: + sem_sistate ctx sis rs m -> sis_ok ctx sis. +Proof. + unfold sem_sistate; + econstructor; + intuition congruence. +Qed. + +(* NB: refinement of (symbolic) internal state *) +Record ristate := + { + (** [ris_smem] represents the current smem symbolic evaluations. + (we also recover the history of smem in ris_smem) *) + ris_smem: smem; + (** For the values in registers: + 1) we store a list of sval evaluations + 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval *) + ris_input_init: bool; + ris_ok_sval: list sval; + ris_sreg:> PTree.t sval + }. + +Definition ris_sreg_get (ris: ristate) r: sval := + match PTree.get r ris with + | None => if ris_input_init ris then Sinput r else Sundef + | Some sv => sv + end. +Coercion ris_sreg_get: ristate >-> Funclass. + +Record ris_ok ctx (ris: ristate) : Prop := { + OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; + OK_RREG: forall sv, List.In sv (ris_ok_sval ris) -> eval_sval ctx sv <> None +}. +Local Hint Resolve OK_RMEM OK_RREG: core. +Local Hint Constructors ris_ok: core. + +Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := { + OK_EQUIV: sis_ok ctx sis <-> ris_ok ctx ris; + MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem); + REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r); + (* the below invariant allows to evaluate operations in the initial memory of the path instead of the current memory *) + VALID_PRESERV: forall m b ofs, eval_smem ctx sis.(si_smem) = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs +}. +Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ VALID_PRESERV: core. +Local Hint Constructors ris_refines: core. + +Record ris_simu ris1 ris2: Prop := { + SIMU_FAILS: forall sv, List.In sv ris2.(ris_ok_sval) -> List.In sv ris1.(ris_ok_sval); + SIMU_MEM: ris1.(ris_smem) = ris2.(ris_smem); + SIMU_REG: forall r, ris_sreg_get ris1 r = ris_sreg_get ris2 r +}. +Local Hint Resolve SIMU_FAILS SIMU_MEM SIMU_REG: core. +Local Hint Constructors ris_simu: core. +Local Hint Resolve sge_match: core. + +Lemma ris_simu_ok_preserv f1 f2 ris1 ris2 (ctx:simu_proof_context): + ris_simu ris1 ris2 -> ris_ok (bctx1 f1 ctx) ris1 -> ris_ok (bctx2 f2 ctx) ris2. +Proof. + intros SIMU OK; econstructor; eauto. + - erewrite <- SIMU_MEM; eauto. + unfold bctx2; erewrite smem_eval_preserved; eauto. + - unfold bctx2; intros; erewrite eval_sval_preserved; eauto. +Qed. + +Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context) sis1 sis2 rs m: + ris_simu ris1 ris2 -> + ris_refines (bctx1 f1 ctx) ris1 sis1 -> + ris_refines (bctx2 f2 ctx) ris2 sis2 -> + sem_sistate (bctx1 f1 ctx) sis1 rs m -> + sem_sistate (bctx2 f2 ctx) sis2 rs m. +Proof. + intros RIS REF1 REF2 SEM. + (* destruct REF1. destruct REF2. *) + exploit sem_sok; eauto. + erewrite OK_EQUIV; eauto. + intros ROK1. + exploit ris_simu_ok_preserv; eauto. + intros ROK2. generalize ROK2; erewrite <- OK_EQUIV; eauto. + intros SOK2. + destruct SEM as (PRE & SMEM & SREG). + unfold sem_sistate; intuition eauto. + + erewrite <- MEM_EQ, <- SIMU_MEM; eauto. + unfold bctx2; erewrite smem_eval_preserved; eauto. + erewrite MEM_EQ; eauto. + + erewrite <- REG_EQ, <- SIMU_REG; eauto. + unfold bctx2; erewrite eval_sval_preserved; eauto. + erewrite REG_EQ; eauto. +Qed. + +Definition rfv_refines ctx (rfv sfv: sfval): Prop := + forall rs m t s, sem_sfval ctx rfv rs m t s <-> sem_sfval ctx sfv rs m t s. + +Definition rfv_simu (rfv1 rfv2: sfval): Prop := rfv1 = rfv2. + +Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context) sfv1 sfv2 rs m t s: + rfv_simu rfv1 rfv2 -> + rfv_refines (bctx1 f1 ctx) rfv1 sfv1 -> + rfv_refines (bctx2 f2 ctx) rfv2 sfv2 -> + sem_sfval (bctx1 f1 ctx) sfv1 rs m t s -> + sem_sfval (bctx2 f2 ctx) sfv2 rs m t s. +Proof. + unfold rfv_simu, rfv_refines; intros X REF1 REF2 SEM. subst. + rewrite <- REF2. + assert (sem_sfval (bctx1 f1 ctx) rfv2 rs m t s). { rewrite REF1; auto. } +Admitted. + +(* refinement of (symbolic) state *) +Inductive rstate := + | Rfinal (ris: ristate) (rfv: sfval) + | Rcond (cond: condition) (rargs: list_sval) (rifso rifnot: rstate) + | Rabort + . + +Inductive rst_simu: rstate -> rstate -> Prop := + | Rfinal_simu ris1 ris2 rfv1 rfv2: + ris_simu ris1 ris2 -> + rfv_simu rfv1 rfv2 -> + rst_simu (Rfinal ris1 rfv1) (Rfinal ris2 rfv2) + | Rcond_simu cond rargs rifso1 rifnot1 rifso2 rifnot2: + rst_simu rifso1 rifso2 -> + rst_simu rifnot1 rifnot2 -> + rst_simu (Rcond cond rargs rifso1 rifnot1) (Rcond cond rargs rifso2 rifnot2) + | Rabort_simu: rst_simu Rabort Rabort +(* TODO: extension à voir dans un second temps ! + | Rcond_skip cond rargs rifso1 rifnot1 rst: + rst_simu rifso1 rst -> + rst_simu rifnot1 rst -> + rst_simu (Rcond cond rargs rifso1 rifnot1) rst +*) + . + +(* Comment prend on en compte le ris en cours d'execution ??? *) +Inductive rst_refines ctx: (* Prop -> *) rstate -> sstate -> Prop := + | refines_Sfinal ris sis rfv sfv (*ok: Prop*) + (*OK: ris_ok ctx ris -> ok*) + (RIS: ris_refines ctx ris sis) + (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) + : rst_refines ctx (*ok*) (Rfinal ris rfv) (Sfinal sis sfv) + | refines_Scond cond rargs args sm rifso rifnot ifso ifnot (*ok1 ok2: Prop*) + (*OK1: ok2 -> ok1*) + (RCOND: (* ok2 -> *) seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm) + (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx (*ok2*) rifso ifso) + (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx (*ok2*) rifnot ifnot) + : rst_refines ctx (*ok1*) (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot) + | refines_Sabort + : rst_refines ctx Rabort Sabort + . + +Lemma rst_simu_correct rst1 rst2: + rst_simu rst1 rst2 -> + forall f1 f2 ctx st1 st2 (*ok1 ok2*), + rst_refines (*ok1*) (bctx1 f1 ctx) rst1 st1 -> + rst_refines (*ok2*) (bctx2 f2 ctx) rst2 st2 -> + (* ok1 -> ok2 -> *) + sstate_simu f1 f2 ctx st1 st2. +Proof. + induction 1; simpl; auto. + - (* final *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. + inv REF1. inv REF2. inv SEM1. + exploit sem_sok; eauto. + rewrite OK_EQUIV; eauto. + intros RIS_OK1. + exploit (ris_simu_ok_preserv f1 f2); eauto. + intros RIS_OK2. intuition. + exploit ris_simu_correct; eauto. + exploit rvf_simu_correct; eauto. + simpl. + eexists; split. + + econstructor; eauto. + simpl. + admit. (* TODO: condition sur les tr_inputs: à ajouter plutôt dans le simu_proof_context ! *) + + admit. + - (* cond *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. + inv REF1. inv REF2. inv SEM1. + destruct b; simpl. + + assert (TODO1: rst_refines (bctx1 f1 ctx) rifso1 ifso). admit. + assert (TODO2: rst_refines (bctx2 f2 ctx) rifso2 ifso0). admit. + exploit IHrst_simu1; eauto. + intros (s2 & X1 & X2). + exists s2; split; eauto. + econstructor; eauto. + * assert (TODO3: seval_condition (bctx2 f2 ctx) cond args0 sm0 = Some true). admit. + eauto. + * simpl; eauto. + + admit. + - (* abort *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. + inv REF1. inv SEM1. +Admitted. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 3083ca71..c7a44153 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1057,102 +1057,4 @@ Proof. intros (s3 & STEP2 & EQ2). clear STEP1; eexists; split; eauto. eapply equiv_state_trans; eauto. -Qed. - -(** * Preservation properties *) - -Section SymbValPreserved. - -Variable ge ge': BTL.genv. - -Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. - -Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. - -Lemma senv_find_symbol_preserved id: - Senv.find_symbol ge id = Senv.find_symbol ge' id. -Proof. - destruct senv_preserved_BTL as (A & B & C). congruence. -Qed. - -Lemma senv_symbol_address_preserved id ofs: - Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. -Proof. - unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. - reflexivity. -Qed. - -Variable stk stk': list stackframe. -Variable f f': function. -Variable sp: val. -Variable rs0: regset. -Variable m0: mem. - -Lemma eval_sval_preserved sv: - eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. -Proof. - induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) - (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. - erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. - rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. - rewrite IHsv1; auto. -Qed. - -Lemma seval_builtin_arg_preserved m: forall bs varg, - seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> - seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. -Proof. - induction 1; simpl. - all: try (constructor; auto). - - rewrite <- eval_sval_preserved. assumption. - - rewrite <- senv_symbol_address_preserved. assumption. - - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. -Qed. - -Lemma seval_builtin_args_preserved m lbs vargs: - seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> - seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. -Proof. - induction 1; constructor; eauto. - eapply seval_builtin_arg_preserved; auto. -Qed. - -Lemma list_sval_eval_preserved lsv: - eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. -Proof. - induction lsv; simpl; auto. - rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. - rewrite IHlsv; auto. -Qed. - -Lemma smem_eval_preserved sm: - eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. -Proof. - induction sm; simpl; auto. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. - rewrite eval_sval_preserved; auto. -Qed. - -Lemma seval_condition_preserved cond lsv sm: - seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. -Proof. - unfold seval_condition. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - rewrite smem_eval_preserved; auto. -Qed. - -End SymbValPreserved. +Qed. \ No newline at end of file -- cgit From abfc8509ec4686924095c1328179fa06766dc496 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 2 Jun 2021 08:20:58 +0200 Subject: the current "high-level" specification of the simulation test will not work ! --- scheduling/BTL_SEsimuref.v | 94 +++++++++++++++++++++++++++++++++------------- scheduling/BTL_SEtheory.v | 58 +++++++++++++++++++++++----- scheduling/BTL_Scheduler.v | 9 ++++- 3 files changed, 124 insertions(+), 37 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index af2159cb..65052310 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -182,8 +182,8 @@ Local Hint Resolve SIMU_FAILS SIMU_MEM SIMU_REG: core. Local Hint Constructors ris_simu: core. Local Hint Resolve sge_match: core. -Lemma ris_simu_ok_preserv f1 f2 ris1 ris2 (ctx:simu_proof_context): - ris_simu ris1 ris2 -> ris_ok (bctx1 f1 ctx) ris1 -> ris_ok (bctx2 f2 ctx) ris2. +Lemma ris_simu_ok_preserv f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2): + ris_simu ris1 ris2 -> ris_ok (bctx1 ctx) ris1 -> ris_ok (bctx2 ctx) ris2. Proof. intros SIMU OK; econstructor; eauto. - erewrite <- SIMU_MEM; eauto. @@ -191,12 +191,12 @@ Proof. - unfold bctx2; intros; erewrite eval_sval_preserved; eauto. Qed. -Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context) sis1 sis2 rs m: +Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2 rs m: ris_simu ris1 ris2 -> - ris_refines (bctx1 f1 ctx) ris1 sis1 -> - ris_refines (bctx2 f2 ctx) ris2 sis2 -> - sem_sistate (bctx1 f1 ctx) sis1 rs m -> - sem_sistate (bctx2 f2 ctx) sis2 rs m. + ris_refines (bctx1 ctx) ris1 sis1 -> + ris_refines (bctx2 ctx) ris2 sis2 -> + sem_sistate (bctx1 ctx) sis1 rs m -> + sem_sistate (bctx2 ctx) sis2 rs m. Proof. intros RIS REF1 REF2 SEM. (* destruct REF1. destruct REF2. *) @@ -216,21 +216,59 @@ Proof. erewrite REG_EQ; eauto. Qed. -Definition rfv_refines ctx (rfv sfv: sfval): Prop := - forall rs m t s, sem_sfval ctx rfv rs m t s <-> sem_sfval ctx sfv rs m t s. +Definition option_refines ctx (orsv: option sval) (osv: option sval) := + match orsv, osv with + | Some rsv, Some sv => eval_sval ctx rsv = eval_sval ctx sv + | None, None => True + | _, _ => False + end. + +Definition sum_refines ctx (rsi: sval + ident) (si: sval + ident) := + match rsi, si with + | inl rsv, inl sv => eval_sval ctx rsv = eval_sval ctx sv + | inr id, inr id' => id = id' + | _, _ => False + end. + +Definition bargs_refines ctx (rargs: list (builtin_arg sval)) (args: list (builtin_arg sval)): Prop := + eval_list_builtin_sval ctx rargs = eval_list_builtin_sval ctx args. + +Inductive rfv_refines ctx: sfval -> sfval -> Prop := + | refines_Sgoto pc: rfv_refines ctx (Sgoto pc) (Sgoto pc) + | refines_Scall: forall sig rvos ros rargs args r pc + (SUM:sum_refines ctx rvos ros) + (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args) + ,rfv_refines ctx (Scall sig rvos rargs r pc) (Scall sig ros args r pc) + | refines_Stailcall: forall sig rvos ros rargs args + (SUM:sum_refines ctx rvos ros) + (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args) + ,rfv_refines ctx (Stailcall sig rvos rargs) (Stailcall sig ros args) + | refines_Sbuiltin: forall ef lbra lba br pc + (BARGS: bargs_refines ctx lbra lba) + ,rfv_refines ctx (Sbuiltin ef lbra br pc) (Sbuiltin ef lba br pc) + | refines_Sjumptable: forall rsv sv lpc + (VAL: eval_sval ctx rsv = eval_sval ctx sv) + ,rfv_refines ctx (Sjumptable rsv lpc) (Sjumptable sv lpc) + | refines_Sreturn: forall orsv osv + (OPT:option_refines ctx orsv osv) + ,rfv_refines ctx (Sreturn orsv) (Sreturn osv) +. Definition rfv_simu (rfv1 rfv2: sfval): Prop := rfv1 = rfv2. -Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context) sfv1 sfv2 rs m t s: +Local Hint Constructors sem_sfval equiv_state: core. + +Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context f1 f2) sfv1 sfv2 rs m t s: rfv_simu rfv1 rfv2 -> - rfv_refines (bctx1 f1 ctx) rfv1 sfv1 -> - rfv_refines (bctx2 f2 ctx) rfv2 sfv2 -> - sem_sfval (bctx1 f1 ctx) sfv1 rs m t s -> - sem_sfval (bctx2 f2 ctx) sfv2 rs m t s. + rfv_refines (bctx1 ctx) rfv1 sfv1 -> + rfv_refines (bctx2 ctx) rfv2 sfv2 -> + sem_sfval (bctx1 ctx) sfv1 rs m t s -> + exists s', sem_sfval (bctx2 ctx) sfv2 rs m t s' /\ equiv_state s s'. Proof. - unfold rfv_simu, rfv_refines; intros X REF1 REF2 SEM. subst. - rewrite <- REF2. - assert (sem_sfval (bctx1 f1 ctx) rfv2 rs m t s). { rewrite REF1; auto. } + unfold rfv_simu; intros X REF1 REF2 SEM. subst. + unfold bctx2; destruct REF1; inv REF2; inv SEM; simpl. + - eexists; split; eauto; simpl. + (* eapply State_equiv; eauto. NE MARCHE PAS ! *) Admitted. (* refinement of (symbolic) state *) @@ -277,11 +315,11 @@ Inductive rst_refines ctx: (* Prop -> *) rstate -> sstate -> Prop := Lemma rst_simu_correct rst1 rst2: rst_simu rst1 rst2 -> - forall f1 f2 ctx st1 st2 (*ok1 ok2*), - rst_refines (*ok1*) (bctx1 f1 ctx) rst1 st1 -> - rst_refines (*ok2*) (bctx2 f2 ctx) rst2 st2 -> + forall f1 f2 (ctx: simu_proof_context f1 f2) st1 st2 (*ok1 ok2*), + rst_refines (*ok1*) (bctx1 ctx) rst1 st1 -> + rst_refines (*ok2*) (bctx2 ctx) rst2 st2 -> (* ok1 -> ok2 -> *) - sstate_simu f1 f2 ctx st1 st2. + sstate_simu ctx st1 st2. Proof. induction 1; simpl; auto. - (* final *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. @@ -293,25 +331,29 @@ Proof. intros RIS_OK2. intuition. exploit ris_simu_correct; eauto. exploit rvf_simu_correct; eauto. - simpl. + simpl. (* eexists; split. + econstructor; eauto. simpl. - admit. (* TODO: condition sur les tr_inputs: à ajouter plutôt dans le simu_proof_context ! *) + clear SIS. + admit. (* TODO: condition sur les tr_inputs du simu_proof_context ! *) + (* TODO: le rfv_refines est trop sémantique ! *) + admit. - (* cond *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. inv REF1. inv REF2. inv SEM1. destruct b; simpl. - + assert (TODO1: rst_refines (bctx1 f1 ctx) rifso1 ifso). admit. - assert (TODO2: rst_refines (bctx2 f2 ctx) rifso2 ifso0). admit. + + assert (TODO1: rst_refines (bctx1 ctx) rifso1 ifso). admit. + assert (TODO2: rst_refines (bctx2 ctx) rifso2 ifso0). admit. exploit IHrst_simu1; eauto. intros (s2 & X1 & X2). exists s2; split; eauto. econstructor; eauto. - * assert (TODO3: seval_condition (bctx2 f2 ctx) cond args0 sm0 = Some true). admit. + * assert (TODO3: seval_condition (bctx2 ctx) cond args0 sm0 = Some true). admit. eauto. * simpl; eauto. + admit. - (* abort *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. inv REF1. inv SEM1. +*) + Admitted. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index c7a44153..72789121 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1022,38 +1022,76 @@ Qed. (** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *) -Record simu_proof_context := Sctx { +(* +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)). + +Lemma equiv_input_regs_union f1 f2: + equiv_input_regs f1 f2 -> forall tbl, union_inputs f1 tbl = union_inputs f2 tbl. +Proof. + intros (EQNone & EQSome). induction tbl as [|pc l']; simpl; auto. + generalize (EQNone pc) (EQSome pc); clear EQNone EQSome; intros EQN EQS. + do 2 autodestruct; intuition; try_simplify_someHyps. + intros; exploit EQS; eauto; clear EQS. congruence. +Qed. + +Lemma equiv_input_regs_pre f1 f2 tbl or: + equiv_input_regs f1 f2 -> pre_inputs f1 tbl or = pre_inputs f2 tbl or. +Proof. + intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto. +Qed. +*) + +Record simu_proof_context {f1 f2: BTL.function} := Sctx { sge1: BTL.genv; sge2: BTL.genv; sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; sstk1: list BTL.stackframe; sstk2: list BTL.stackframe; sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; + (* inputs_equiv: equiv_input_regs f1 f2;*) (* TODO a-t-on besoin de ça ? *) ssp: val; srs0: regset; sm0: mem }. +Arguments simu_proof_context: clear implicits. -Definition bctx1 f1 (ctx: simu_proof_context):= Bctx ctx.(sge1) ctx.(sstk1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). -Definition bctx2 f2 (ctx: simu_proof_context):= Bctx ctx.(sge2) ctx.(sstk2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx1 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge1) ctx.(sstk1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx2 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge2) ctx.(sstk2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). -Definition sstate_simu f1 f2 (ctx: simu_proof_context) (st1 st2: sstate) := - forall t s1, sem_sstate (bctx1 f1 ctx) t s1 st1 -> - exists s2, sem_sstate (bctx2 f2 ctx) t s2 st2 /\ equiv_state s1 s2. -Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall ctx, sstate_simu f1 f2 ctx (sexec f1 ib1) (sexec f2 ib2). +(* TODO: A REVOIR. L'approche suivie ne marche pas !!! + +le [equiv_state] ci-dessous n'est pas assez général + => il faut un [match_state] qui va dépendre de [BTL_Scheduler.match_function] (dans le [match_stackframe]). + +Or, le [sstate_simu] qu'on cherche à définir ici, c'était pour définir ce [match_function]: circularité ! + +Une solution: définir le match_function à l'aide d'un [iblock_step_simu] ? + +Pas sûr que ça marche, on aura aussi la dépendance circulaire entre [simu_proof_context] et [match_function] ! + +Autre solution: définir [sfval_simu] syntaxiquement comme dans [RTLpath_SEtheory] pour éviter le problème de circularité ! + +*) +Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2: sstate) := + forall t s1, sem_sstate (bctx1 ctx) t s1 st1 -> + exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2. + +Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2). Theorem symbolic_simu_correct f1 f2 ib1 ib2: symbolic_simu f1 f2 ib1 ib2 -> - forall ctx t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> + forall (ctx: simu_proof_context f1 f2) t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 -> exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2. Proof. unfold symbolic_simu, sstate_simu. intros SIMU ctx t s1 STEP1. - exploit (sexec_correct (bctx1 f1 ctx)); simpl; eauto. + exploit (sexec_correct (bctx1 ctx)); simpl; eauto. intros; exploit SIMU; eauto. intros (s2 & SEM1 & EQ1). - exploit (sexec_exact (bctx2 f2 ctx)); simpl; eauto. + exploit (sexec_exact (bctx2 ctx)); simpl; eauto. intros (s3 & STEP2 & EQ2). clear STEP1; eexists; split; eauto. eapply equiv_state_trans; eauto. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index b479204c..bcddcf5d 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -15,7 +15,8 @@ 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; - symbolic_simu_ok: forall pc ib1, (fn_code f1)!pc = Some ib1 -> + (* 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); }. @@ -23,6 +24,12 @@ 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 -> forall pc ib1, (fn_code f1)!pc = Some ib1 -> -- cgit From 3db0ba1f8071c35dcc432f8047cb437343ef37ce Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 2 Jun 2021 11:13:52 +0200 Subject: some advance in main liveness lemmas --- scheduling/BTL_Livecheck.v | 96 ++++++++++++++++++++++++++++++++-------------- 1 file changed, 68 insertions(+), 28 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 49660222..50719f67 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -169,11 +169,11 @@ Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := Inductive eqlive_states: state -> state -> Prop := | eqlive_states_intro - st1 st2 f sp pc rs1 rs2 m + ibf st1 st2 f sp pc rs1 rs2 m (STACKS: list_forall2 eqlive_stackframes st1 st2) - (LIVE: liveness_ok_function f): - (*(PATH: f.(fn_code)!pc = Some ibf)*) - (*(EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2):*) + (LIVE: liveness_ok_function f) + (PATH: f.(fn_code)!pc = Some ibf) + (EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2): eqlive_states (State st1 f sp pc rs1 m) (State st2 f sp pc rs2 m) | eqlive_states_call st1 st2 f args m (LIVE: liveness_ok_fundef f) @@ -193,40 +193,80 @@ Let ge := Genv.globalenv prog. Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. -Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m1 rs2 m2 ib s1 t - (STEP : iblock_step tid (Genv.globalenv prog) stk1 f sp rs1 m1 ib t s1) - (STACKS : list_forall2 eqlive_stackframes stk1 stk2) - (LIVE : liveness_ok_function f) - :exists s2 : state, - iblock_step tr_inputs (Genv.globalenv prog) stk2 f sp rs2 m2 ib t s2 /\ - eqlive_states s1 s2. +Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. + +Lemma cfgsem2fsem_ibistep_simu sp rs1 m1 rs1' m1' of ib + (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of): forall + rs2 m2, + exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of). Proof. + induction ISTEP; simpl; try_simplify_someHyps; intros. + - (* Bop *) + inversion_SOME v0; intros EVAL'; + repeat eexists. + inv EVAL'. Admitted. +Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 + (FSTEP: final_step tid ge stk1 f sp rs1 m1 fin t s1) + (STACKS: list_forall2 eqlive_stackframes stk1 stk2) + : exists s2, final_step tr_inputs ge stk2 f sp rs2 m2 fin t s2 + /\ eqlive_states s1 s2. +Proof. + destruct FSTEP. + - (* Bgoto *) + (*eexists; split; simpl ; econstructor; eauto.*) +Admitted. + +Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: + iblock_step tid (Genv.globalenv prog) stk1 f sp rs1 m ibf.(entry) t s1 -> + list_forall2 eqlive_stackframes stk1 stk2 -> + eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> + liveness_ok_function f -> + (fn_code f) ! pc = Some ibf -> + exists s2 : state, + iblock_step tr_inputs (Genv.globalenv prog) stk2 f sp rs2 m ibf.(entry) t s2 /\ + eqlive_states s1 s2. +Proof. + intros STEP STACKS EQLIVE LIVE PC. + unfold liveness_ok_function, liveness_checker in LIVE. + destruct list_iblock_checker eqn:LIBC in LIVE; try discriminate. + clear LIVE. + destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP). + exploit cfgsem2fsem_ibistep_simu; eauto. + intros (rs2' & m2' & ISTEP2). + rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP. + exploit cfgsem2fsem_finalstep_simu; eauto. + intros (s2 & FSTEP2 & STATES). clear FSTEP. + unfold iblock_step. repeat eexists; eauto. +Qed. + + Local Hint Constructors step: core. -Lemma cfgsem2fsem_step_simu s1 s1' t s2 - (STEP : step tid (Genv.globalenv prog) s1 t s1') - (STATES : eqlive_states s1 s2) - :exists s2' : state, +Lemma cfgsem2fsem_step_simu s1 s1' t s2: + step tid (Genv.globalenv prog) s1 t s1' -> + eqlive_states s1 s2 -> + exists s2' : state, step tr_inputs (Genv.globalenv prog) s2 t s2' /\ eqlive_states s1' s2'. Proof. - destruct STEP; inv STATES. - - (* iblock *) + destruct 1 as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; intros STATES; + try (inv STATES; inv LIVE; eexists; split; econstructor; eauto; fail). + - inv STATES; simplify_someHyps. + intros. exploit cfgsem2fsem_ibstep_simu; eauto. - intros (s2 & STEP2 & STATES2). + intros (s2 & STEP2 & EQUIV2). eexists; split; eauto. - - (* internal call *) - inv LIVE. - eexists; split; repeat econstructor; eauto. - - (* external_call *) - eexists; split; repeat econstructor; eauto. - - (* return *) - inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst. - inv STF2. - eexists; split; econstructor; eauto. -Qed. + - inv STATES; inv LIVE. + eexists; split; econstructor; eauto. all : admit. (* TODO gourdinl *) + - inv STATES. + inversion STACKS as [|s1 st1 s' s2 STACK STACKS']; subst; clear STACKS. + inv STACK. + exists (State s2 f sp pc (rs2 # res <- vres) m); split. + * apply exec_return. + * eapply eqlive_states_intro; eauto. +Admitted. Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). Proof. -- cgit From 37c8b2a474bb07890f21db0535b5886bec4e4e6d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 2 Jun 2021 14:48:59 +0200 Subject: fix the definition of [symbolic_simu] in BTL_SEtheory --- scheduling/BTL_SEsimuref.v | 116 ++--------------- scheduling/BTL_SEtheory.v | 307 +++++++++++++++++++++++++++++++++++---------- 2 files changed, 251 insertions(+), 172 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 65052310..e9ae11e0 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -1,9 +1,7 @@ (** Refinement of BTL_SEtheory data-structures in order to introduce (and prove correct) a lower-level specification of the simulation test. - TODO: not yet with hash-consing ? Or "fake" hash-consing only ? - - Ceci est un "bac à sable". + Ceci est un "bac à sable". TODO: A REVOIR COMPLETEMENT. Introduire "fake" hash-consing - On introduit une représentation plus concrète pour les types d'état symbolique [sistate] et [sstate]. - Etant donné une spécification intuitive "*_simu" pour tester la simulation sur cette représentation des états symboliques, @@ -17,106 +15,6 @@ Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL BTL OptionMonad BTL_SEtheory. -(** * Preservation properties *) - -(* TODO: inherited from RTLpath. Use simu_proof_context + bctx2 + bctx1 instead of all the hypotheses/variables below ? *) - -Section SymbValPreserved. - -Variable ge ge': BTL.genv. - -Hypothesis symbols_preserved_BTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. - -Hypothesis senv_preserved_BTL: Senv.equiv ge ge'. - -Lemma senv_find_symbol_preserved id: - Senv.find_symbol ge id = Senv.find_symbol ge' id. -Proof. - destruct senv_preserved_BTL as (A & B & C). congruence. -Qed. - -Lemma senv_symbol_address_preserved id ofs: - Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs. -Proof. - unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. - reflexivity. -Qed. - -Variable stk stk': list stackframe. -Variable f f': function. -Variable sp: val. -Variable rs0: regset. -Variable m0: mem. - -Lemma eval_sval_preserved sv: - eval_sval (Bctx ge stk f sp rs0 m0) sv = eval_sval (Bctx ge' stk' f' sp rs0 m0) sv. -Proof. - induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv) - (P1 := fun sm => eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. - erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. - rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. - rewrite IHsv1; auto. -Qed. - -Lemma seval_builtin_arg_preserved m: forall bs varg, - seval_builtin_arg (Bctx ge stk f sp rs0 m0) m bs varg -> - seval_builtin_arg (Bctx ge' stk' f' sp rs0 m0) m bs varg. -Proof. - induction 1; simpl. - all: try (constructor; auto). - - rewrite <- eval_sval_preserved. assumption. - - rewrite <- senv_symbol_address_preserved. assumption. - - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. -Qed. - -Lemma seval_builtin_args_preserved m lbs vargs: - seval_builtin_args (Bctx ge stk f sp rs0 m0) m lbs vargs -> - seval_builtin_args (Bctx ge' stk' f' sp rs0 m0) m lbs vargs. -Proof. - induction 1; constructor; eauto. - eapply seval_builtin_arg_preserved; auto. -Qed. - -Lemma list_sval_eval_preserved lsv: - eval_list_sval (Bctx ge stk f sp rs0 m0) lsv = eval_list_sval (Bctx ge' stk' f' sp rs0 m0) lsv. -Proof. - induction lsv; simpl; auto. - rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. - rewrite IHlsv; auto. -Qed. - -Lemma smem_eval_preserved sm: - eval_smem (Bctx ge stk f sp rs0 m0) sm = eval_smem (Bctx ge' stk' f' sp rs0 m0) sm. -Proof. - induction sm; simpl; auto. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing _ sp _ _); auto. - rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. - rewrite eval_sval_preserved; auto. -Qed. - -Lemma seval_condition_preserved cond lsv sm: - seval_condition (Bctx ge stk f sp rs0 m0) cond lsv sm = seval_condition (Bctx ge' stk' f' sp rs0 m0) cond lsv sm. -Proof. - unfold seval_condition. - rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - rewrite smem_eval_preserved; auto. -Qed. - -End SymbValPreserved. - (** * ... *) Record sis_ok ctx (sis: sistate): Prop := { @@ -187,8 +85,8 @@ Lemma ris_simu_ok_preserv f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2): Proof. intros SIMU OK; econstructor; eauto. - erewrite <- SIMU_MEM; eauto. - unfold bctx2; erewrite smem_eval_preserved; eauto. - - unfold bctx2; intros; erewrite eval_sval_preserved; eauto. + erewrite <- smem_eval_preserved; eauto. + - intros; erewrite <- eval_sval_preserved; eauto. Qed. Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2 rs m: @@ -199,7 +97,6 @@ Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2 sem_sistate (bctx2 ctx) sis2 rs m. Proof. intros RIS REF1 REF2 SEM. - (* destruct REF1. destruct REF2. *) exploit sem_sok; eauto. erewrite OK_EQUIV; eauto. intros ROK1. @@ -209,13 +106,15 @@ Proof. destruct SEM as (PRE & SMEM & SREG). unfold sem_sistate; intuition eauto. + erewrite <- MEM_EQ, <- SIMU_MEM; eauto. - unfold bctx2; erewrite smem_eval_preserved; eauto. + erewrite <- smem_eval_preserved; eauto. erewrite MEM_EQ; eauto. + erewrite <- REG_EQ, <- SIMU_REG; eauto. - unfold bctx2; erewrite eval_sval_preserved; eauto. + erewrite <- eval_sval_preserved; eauto. erewrite REG_EQ; eauto. Qed. +(* TODO: + Definition option_refines ctx (orsv: option sval) (osv: option sval) := match orsv, osv with | Some rsv, Some sv => eval_sval ctx rsv = eval_sval ctx sv @@ -357,3 +256,4 @@ Proof. *) Admitted. +*) \ No newline at end of file diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 72789121..c437846e 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -16,7 +16,7 @@ Require Import RTL BTL OptionMonad. Record iblock_exec_context := Bctx { cge: BTL.genv; - cstk: list stackframe; + (* cstk: list stackframe; *) (* having the stack here does seem not a good idea *) cf: function; csp: val; crs0: regset; @@ -400,39 +400,39 @@ Definition sfind_function ctx (svos : sval + ident): option fundef := Import ListNotations. Local Open Scope list_scope. -Inductive sem_sfval ctx: sfval -> regset -> mem -> trace -> state -> Prop := +Inductive sem_sfval ctx stk: sfval -> regset -> mem -> trace -> state -> Prop := | exec_Sgoto pc rs m: - sem_sfval ctx (Sgoto pc) rs m E0 (State (cstk ctx) (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] None rs) m) - | exec_Sreturn stk osv rs m m' v: - (csp ctx) = (Vptr stk Ptrofs.zero) -> - Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> + sem_sfval ctx stk (Sgoto pc) rs m E0 (State stk (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] None rs) m) + | exec_Sreturn pstk osv rs m m' v: + (csp ctx) = (Vptr pstk Ptrofs.zero) -> + Mem.free m pstk 0 (cf ctx).(fn_stacksize) = Some m' -> match osv with Some sv => eval_sval ctx sv | None => Some Vundef end = Some v -> - sem_sfval ctx (Sreturn osv) rs m - E0 (Returnstate (cstk ctx) v m') + sem_sfval ctx stk (Sreturn osv) rs m + E0 (Returnstate stk v m') | exec_Scall rs m sig svos lsv args res pc fd: sfind_function ctx svos = Some fd -> funsig fd = sig -> eval_list_sval ctx lsv = Some args -> - sem_sfval ctx (Scall sig svos lsv res pc) rs m - E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] (Some res) rs)::cstk ctx) fd args m) - | exec_Stailcall stk rs m sig svos args fd m' lsv: + sem_sfval ctx stk (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] (Some res) rs)::stk) fd args m) + | exec_Stailcall pstk rs m sig svos args fd m' lsv: sfind_function ctx svos = Some fd -> funsig fd = sig -> - (csp ctx) = Vptr stk Ptrofs.zero -> - Mem.free m stk 0 (cf ctx).(fn_stacksize) = Some m' -> + (csp ctx) = Vptr pstk Ptrofs.zero -> + Mem.free m pstk 0 (cf ctx).(fn_stacksize) = Some m' -> eval_list_sval ctx lsv = Some args -> - sem_sfval ctx (Stailcall sig svos lsv) rs m - E0 (Callstate (cstk ctx) fd args m') + sem_sfval ctx stk (Stailcall sig svos lsv) rs m + E0 (Callstate stk fd args m') | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: seval_builtin_args ctx m sargs vargs -> external_call ef (cge ctx) vargs m t vres m' -> - sem_sfval ctx (Sbuiltin ef sargs res pc) rs m - t (State (cstk ctx) (cf ctx) (csp ctx) pc (regmap_setres res vres (tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs)) m') + sem_sfval ctx stk (Sbuiltin ef sargs res pc) rs m + t (State stk (cf ctx) (csp ctx) pc (regmap_setres res vres (tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs)) m') | exec_Sjumptable sv tbl pc' n rs m: eval_sval ctx sv = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - sem_sfval ctx (Sjumptable sv tbl) rs m - E0 (State (cstk ctx) (cf ctx) (csp ctx) pc' (tr_inputs ctx.(cf) tbl None rs) m) + sem_sfval ctx stk (Sjumptable sv tbl) rs m + E0 (State stk (cf ctx) (csp ctx) pc' (tr_inputs ctx.(cf) tbl None rs) m) . (* Syntax and Semantics of symbolic internal states *) @@ -470,10 +470,10 @@ Definition sexec_final_sfv (i: final) (sis: sistate): sfval := Local Hint Constructors sem_sfval: core. -Lemma sexec_final_svf_correct ctx i sis t rs m s: +Lemma sexec_final_svf_correct ctx stk i sis t rs m s: sem_sistate ctx sis rs m -> - final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s -> - sem_sfval ctx (sexec_final_sfv i sis) rs m t s. + final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s -> + sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s. Proof. intros (PRE&MEM®). destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto. @@ -493,10 +493,10 @@ Qed. Local Hint Constructors final_step: core. Local Hint Resolve seval_builtin_args_exact: core. -Lemma sexec_final_svf_exact ctx i sis t rs m s: +Lemma sexec_final_svf_exact ctx stk i sis t rs m s: sem_sistate ctx sis rs m -> - sem_sfval ctx (sexec_final_sfv i sis) rs m t s - -> final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs m i t s. + sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s + -> final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s. Proof. intros (PRE&MEM®). destruct i; simpl; intros LAST; inv LAST; eauto. @@ -705,19 +705,61 @@ Inductive sstate := | Sabort . +(* outcome of a symbolic execution path *) +Record soutcome := sout { + so_sis: sistate; + so_svf: sfval; +}. + +Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := + match st with + | Sfinal sis sfv => Some (sout sis sfv) + | Scond cond args sm ifso ifnot => + SOME b <- seval_condition ctx cond args sm IN + run_sem_isstate ctx (if b then ifso else ifnot) + | _ => None + end. + (* transition (t,s) produced by a sstate in initial context ctx *) -Inductive sem_sstate ctx t s: sstate -> Prop := +Inductive sem_sstate ctx stk t s: sstate -> Prop := | sem_Sfinal sis sfv rs m (SIS: sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m) - (SFV: sem_sfval ctx sfv rs m t s) - : sem_sstate ctx t s (Sfinal sis sfv) + (SFV: sem_sfval ctx stk sfv rs m t s) + : sem_sstate ctx stk t s (Sfinal sis sfv) | sem_Scond b cond args sm ifso ifnot (SEVAL: seval_condition ctx cond args sm = Some b) - (SELECT: sem_sstate ctx t s (if b then ifso else ifnot)) - : sem_sstate ctx t s (Scond cond args sm ifso ifnot) + (SELECT: sem_sstate ctx stk t s (if b then ifso else ifnot)) + : sem_sstate ctx stk t s (Scond cond args sm ifso ifnot) (* NB: Sabort: fails to produce a transition *) . +Lemma sem_sstate_run ctx stk st t s: + sem_sstate ctx stk t s st -> + exists sis sfv rs m, + run_sem_isstate ctx st = Some (sout sis sfv) + /\ sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m + /\ sem_sfval ctx stk sfv rs m t s + . +Proof. + induction 1; simpl; try_simplify_someHyps; do 4 eexists; intuition eauto. +Qed. + +Local Hint Resolve sem_Sfinal: core. + +Lemma run_sem_sstate ctx st sis sfv: + run_sem_isstate ctx st = Some (sout sis sfv) -> + forall rs m stk s t, + sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m -> + sem_sfval ctx stk sfv rs m t s -> + sem_sstate ctx stk t s st + . +Proof. + induction st; simpl; try_simplify_someHyps. + autodestruct; intros; econstructor; eauto. + autodestruct; eauto. +Qed. + + (** * Idée de l'execution symbolique en Continuation Passing Style [k] ci-dessous est la continuation (c-a-d. la suite de la construction de l'arbre qu'on va appliquer dans chaque branche) @@ -755,14 +797,14 @@ Local Hint Constructors sem_sstate: core. Local Hint Resolve sexec_op_correct sexec_final_svf_correct tr_sis_regs_correct_aux tr_sis_regs_correct sexec_load_TRAP_correct sexec_store_correct sis_init_correct: core. -Lemma sexec_rec_correct ctx t s ib rs m rs1 m1 ofin +Lemma sexec_rec_correct ctx stk t s ib rs m rs1 m1 ofin (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): forall sis k (SIS: sem_sistate ctx sis rs m) (CONT: match ofin with - | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx t s (k sis') - | Some fin => final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs1 m1 fin t s + | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx stk t s (k sis') + | Some fin => final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs1 m1 fin t s end), - sem_sstate ctx t s (sexec_rec (cf ctx) ib sis k). + sem_sstate ctx stk t s (sexec_rec (cf ctx) ib sis k). Proof. induction ISTEP; simpl; try autodestruct; eauto. (* final value *) @@ -780,9 +822,9 @@ Qed. (* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) (sexec is a correct over-approximation) *) -Theorem sexec_correct ctx ib t s: - iblock_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> - sem_sstate ctx t s (sexec (cf ctx) ib). +Theorem sexec_correct ctx stk ib t s: + iblock_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s -> + sem_sstate ctx stk t s (sexec (cf ctx) ib). Proof. destruct 1 as (rs' & m' & fin & ISTEP & FSTEP). eapply sexec_rec_correct; simpl; eauto. @@ -812,10 +854,10 @@ Qed. Local Hint Constructors equiv_stackframe list_forall2: core. Local Hint Resolve regmap_setres_eq equiv_stack_refl equiv_stack_refl: core. -Lemma sem_sfval_equiv rs1 rs2 ctx sfv m t s: - sem_sfval ctx sfv rs1 m t s -> +Lemma sem_sfval_equiv rs1 rs2 ctx stk sfv m t s: + sem_sfval ctx stk sfv rs1 m t s -> (forall r, (str_regs (cf ctx) sfv rs1)#r = (str_regs (cf ctx) sfv rs2)#r) -> - exists s', sem_sfval ctx sfv rs2 m t s' /\ equiv_state s s'. + exists s', sem_sfval ctx stk sfv rs2 m t s' /\ equiv_state s s'. Proof. unfold str_regs. destruct 1; simpl in *; intros; subst; eexists; split; econstructor; eauto; try congruence. @@ -877,9 +919,9 @@ Qed. Local Hint Resolve sexec_op_preserves_abort sexec_load_preserves_abort sexec_store_preserves_abort sem_sistate_tr_sis_exclude_abort: core. -Lemma sexec_exclude_abort ctx ib t s1: forall sis k - (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) - (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) +Lemma sexec_exclude_abort ctx stk ib t s1: forall sis k + (SEXEC: sem_sstate ctx stk t s1 (sexec_rec (cf ctx) ib sis k)) + (CONT: forall sis', sem_sstate ctx stk t s1 (k sis') -> (abort_sistate ctx sis') -> False) (ABORT: abort_sistate ctx sis), False. Proof. @@ -952,16 +994,16 @@ Qed. Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort sexec_final_svf_exact: core. -Lemma sexec_rec_exact ctx ib t s1: forall sis k - (SEXEC: sem_sstate ctx t s1 (sexec_rec (cf ctx) ib sis k)) +Lemma sexec_rec_exact ctx stk ib t s1: forall sis k + (SEXEC: sem_sstate ctx stk t s1 (sexec_rec (cf ctx) ib sis k)) rs m (SIS: sem_sistate ctx sis rs m) - (CONT: forall sis', sem_sstate ctx t s1 (k sis') -> (abort_sistate ctx sis') -> False) + (CONT: forall sis', sem_sstate ctx stk t s1 (k sis') -> (abort_sistate ctx sis') -> False) , match iblock_istep_run (cge ctx) (csp ctx) ib rs m with | Some (out rs' m' (Some fin)) => - exists s2, final_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 - | Some (out rs' m' None) => exists sis', (sem_sstate ctx t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m') + exists s2, final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2 + | Some (out rs' m' None) => exists sis', (sem_sstate ctx stk t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m') | None => False end. Proof. @@ -1004,9 +1046,9 @@ Qed. (* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution (sexec is exact). *) -Theorem sexec_exact ctx ib t s1: - sem_sstate ctx t s1 (sexec (cf ctx) ib) -> - exists s2, iblock_step tr_inputs (cge ctx) (cstk ctx) (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 +Theorem sexec_exact ctx stk ib t s1: + sem_sstate ctx stk t s1 (sexec (cf ctx) ib) -> + exists s2, iblock_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2 /\ equiv_state s1 s2. Proof. intros; exploit sexec_rec_exact; eauto. @@ -1047,9 +1089,6 @@ Record simu_proof_context {f1 f2: BTL.function} := Sctx { sge1: BTL.genv; sge2: BTL.genv; sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; - sstk1: list BTL.stackframe; - sstk2: list BTL.stackframe; - sstk_equiv: list_forall2 equiv_stackframe sstk1 sstk2; (* inputs_equiv: equiv_input_regs f1 f2;*) (* TODO a-t-on besoin de ça ? *) ssp: val; srs0: regset; @@ -1057,24 +1096,69 @@ Record simu_proof_context {f1 f2: BTL.function} := Sctx { }. Arguments simu_proof_context: clear implicits. -Definition bctx1 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge1) ctx.(sstk1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). -Definition bctx2 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge2) ctx.(sstk2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx1 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0). +Definition bctx2 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0). + +(* NOTE: we need to mix semantical simulation and syntactic definition on [sfval] in order to abstract + the [match_states] of BTL_Schedulerproof. + Indeed, the [match_states] involves [match_function] in [match_stackframe]. + And, here, we aim to define a notion of simulation for defining [match_function]. -(* TODO: A REVOIR. L'approche suivie ne marche pas !!! + A syntactic definition of the simulation on [sfval] avoids the circularity issue. -le [equiv_state] ci-dessous n'est pas assez général - => il faut un [match_state] qui va dépendre de [BTL_Scheduler.match_function] (dans le [match_stackframe]). +*) -Or, le [sstate_simu] qu'on cherche à définir ici, c'était pour définir ce [match_function]: circularité ! +Inductive optsv_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (option sval) -> (option sval) -> Prop := + | Ssome_simu sv1 sv2: + eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2 + -> optsv_simu ctx (Some sv1) (Some sv2) + | Snone_simu: optsv_simu ctx None None + . -Une solution: définir le match_function à l'aide d'un [iblock_step_simu] ? +Inductive svident_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (sval + ident) -> (sval + ident) -> Prop := + | Sleft_simu sv1 sv2: + eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2 + -> svident_simu ctx (inl sv1) (inl sv2) + | Sright_simu id1 id2: + id1 = id2 + -> svident_simu ctx (inr id1) (inr id2) + . -Pas sûr que ça marche, on aura aussi la dépendance circulaire entre [simu_proof_context] et [match_function] ! +Definition bargs_simu {f1 f2: function} (ctx: simu_proof_context f1 f2) (args1 args2: list (builtin_arg sval)): Prop := + eval_list_builtin_sval (bctx1 ctx) args1 = eval_list_builtin_sval (bctx2 ctx) args2. + +Inductive sfv_simu {f1 f2} (ctx: simu_proof_context f1 f2): sfval -> sfval -> Prop := + | Sgoto_simu pc: sfv_simu ctx (Sgoto pc) (Sgoto pc) + | Scall_simu sig ros1 ros2 args1 args2 r pc + (SVID: svident_simu ctx ros1 ros2) + (ARGS:eval_list_sval (bctx1 ctx) args1 = eval_list_sval (bctx2 ctx) args2) + :sfv_simu ctx (Scall sig ros1 args1 r pc) (Scall sig ros2 args2 r pc) + | Stailcall_simu sig ros1 ros2 args1 args2 + (SVID: svident_simu ctx ros1 ros2) + (ARGS:eval_list_sval (bctx1 ctx) args1 = eval_list_sval (bctx2 ctx) args2) + :sfv_simu ctx (Stailcall sig ros1 args1) (Stailcall sig ros2 args2) + | Sbuiltin_simu ef lba1 lba2 br pc + (BARGS: bargs_simu ctx lba1 lba2) + :sfv_simu ctx (Sbuiltin ef lba1 br pc) (Sbuiltin ef lba2 br pc) + | Sjumptable_simu sv1 sv2 lpc + (VAL: eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2) + :sfv_simu ctx (Sjumptable sv1 lpc) (Sjumptable sv2 lpc) + | simu_Sreturn osv1 osv2 + (OPT:optsv_simu ctx osv1 osv2) + :sfv_simu ctx (Sreturn osv1) (Sreturn osv2) +. -Autre solution: définir [sfval_simu] syntaxiquement comme dans [RTLpath_SEtheory] pour éviter le problème de circularité ! +Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop := + forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> + exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) /\ sfv_simu ctx sfv1 sfv2 + . + +Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2). +(* REM. L'approche suivie initialement ne marche pas !!! *) +(* Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2: sstate) := forall t s1, sem_sstate (bctx1 ctx) t s1 st1 -> exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2. @@ -1095,4 +1179,99 @@ Proof. intros (s3 & STEP2 & EQ2). clear STEP1; eexists; split; eauto. eapply equiv_state_trans; eauto. -Qed. \ No newline at end of file +Qed. +*) + +(** * Preservation properties *) + +Section SymbValPreserved. + +Variable f1 f2: function. + +Hypothesis ctx: simu_proof_context f1 f2. +Local Hint Resolve sge_match: core. + +Lemma eval_sval_preserved sv: + eval_sval (bctx1 ctx) sv = eval_sval (bctx2 ctx) sv. +Proof. + induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv) + (P1 := fun sm => eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite eval_addressing_preserved; eauto. + destruct (eval_addressing _ _ _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. + erewrite eval_addressing_preserved; eauto. + destruct (eval_addressing _ _ _ _); auto. + rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma list_sval_eval_preserved lsv: + eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv. +Proof. + induction lsv; simpl; auto. + rewrite eval_sval_preserved. destruct (eval_sval _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma smem_eval_preserved sm: + eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + erewrite eval_addressing_preserved; eauto. + destruct (eval_addressing _ _ _ _); auto. + rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto. + rewrite eval_sval_preserved; auto. +Qed. + +Lemma seval_condition_preserved cond lsv sm: + seval_condition (bctx1 ctx) cond lsv sm = seval_condition (bctx2 ctx) cond lsv sm. +Proof. + unfold seval_condition. + rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + +Hypothesis senv_preserved_BTL: Senv.equiv (sge1 ctx) (sge2 ctx). + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol (sge1 ctx) id = Senv.find_symbol (sge2 ctx) id. +Proof. + destruct senv_preserved_BTL as (A & B & C). congruence. +Qed. + +Lemma senv_symbol_address_preserved id ofs: + Senv.symbol_address (sge1 ctx) id ofs = Senv.symbol_address (sge2 ctx) id ofs. +Proof. + unfold Senv.symbol_address. rewrite senv_find_symbol_preserved. + reflexivity. +Qed. + +Lemma seval_builtin_arg_preserved m: forall bs varg, + seval_builtin_arg (bctx1 ctx) m bs varg -> + seval_builtin_arg (bctx2 ctx) m bs varg. +Proof. + induction 1; simpl. + all: try (constructor; auto). + - rewrite <- eval_sval_preserved. assumption. + - rewrite <- senv_symbol_address_preserved. assumption. + - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. +Qed. + +Lemma seval_builtin_args_preserved m lbs vargs: + seval_builtin_args (bctx1 ctx) m lbs vargs -> + seval_builtin_args (bctx2 ctx) m lbs vargs. +Proof. + induction 1; constructor; eauto. + eapply seval_builtin_arg_preserved; auto. +Qed. + +End SymbValPreserved. + -- cgit From 79cca6a66cbad09f298e0d3b69974c47a8327fc0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 2 Jun 2021 14:58:35 +0200 Subject: minor renaming --- scheduling/BTL_SEtheory.v | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index c437846e..7f00e46f 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -470,7 +470,7 @@ Definition sexec_final_sfv (i: final) (sis: sistate): sfval := Local Hint Constructors sem_sfval: core. -Lemma sexec_final_svf_correct ctx stk i sis t rs m s: +Lemma sexec_final_sfv_correct ctx stk i sis t rs m s: sem_sistate ctx sis rs m -> final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s -> sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s. @@ -493,7 +493,7 @@ Qed. Local Hint Constructors final_step: core. Local Hint Resolve seval_builtin_args_exact: core. -Lemma sexec_final_svf_exact ctx stk i sis t rs m s: +Lemma sexec_final_sfv_exact ctx stk i sis t rs m s: sem_sistate ctx sis rs m -> sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s -> final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s. @@ -677,8 +677,8 @@ Proof. econstructor; simpl; intuition eauto || congruence. Qed. -Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f (svf: sfval): A := - match svf with +Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f (sfv: sfval): A := + match sfv with | Sgoto pc => tr f [pc] None | Scall _ _ _ res pc => tr f [pc] (Some res) | Stailcall _ _ args => tr f [] None @@ -707,8 +707,8 @@ Inductive sstate := (* outcome of a symbolic execution path *) Record soutcome := sout { - so_sis: sistate; - so_svf: sfval; + _sis: sistate; + _sfv: sfval; }. Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := @@ -717,7 +717,7 @@ Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := | Scond cond args sm ifso ifnot => SOME b <- seval_condition ctx cond args sm IN run_sem_isstate ctx (if b then ifso else ifnot) - | _ => None + | Sabort => None end. (* transition (t,s) produced by a sstate in initial context ctx *) @@ -794,7 +794,7 @@ Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := Definition sexec f ib := sexec_rec f ib sis_init (fun _ => Sabort). Local Hint Constructors sem_sstate: core. -Local Hint Resolve sexec_op_correct sexec_final_svf_correct tr_sis_regs_correct_aux tr_sis_regs_correct +Local Hint Resolve sexec_op_correct sexec_final_sfv_correct tr_sis_regs_correct_aux tr_sis_regs_correct sexec_load_TRAP_correct sexec_store_correct sis_init_correct: core. Lemma sexec_rec_correct ctx stk t s ib rs m rs1 m1 ofin @@ -992,7 +992,7 @@ Proof. intros; rewrite REG; autodestruct; try_simplify_someHyps. Qed. -Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort sexec_final_svf_exact: core. +Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort sexec_final_sfv_exact: core. Lemma sexec_rec_exact ctx stk ib t s1: forall sis k (SEXEC: sem_sstate ctx stk t s1 (sexec_rec (cf ctx) ib sis k)) -- cgit From 1586662d37caefd065f54bd693c0594109b3db36 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 2 Jun 2021 17:21:17 +0200 Subject: lemmas on entrypoint, tr_inputs, eqlive_regs, ... --- scheduling/BTL_Livecheck.v | 133 ++++++++++++++++++++++++++++++++++++++------- 1 file changed, 112 insertions(+), 21 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 50719f67..82efff44 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -120,24 +120,48 @@ Proof. simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto). Qed. + +Definition liveness_checker_bool (f: BTL.function): bool := + f.(fn_code)!(f.(fn_entrypoint)) &&& list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)). + Definition liveness_checker (f: BTL.function): res unit := - match list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)) with + match liveness_checker_bool f with | true => OK tt | false => Error (msg "BTL_Livecheck: liveness_checker failed") end. +Lemma decomp_liveness_checker f: + liveness_checker f = OK tt -> + exists ibf, f.(fn_code)!(f.(fn_entrypoint)) = Some ibf /\ + list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)) = true. +Proof. + intros LIVE; unfold liveness_checker in LIVE. + destruct liveness_checker_bool eqn:EQL; try congruence. + clear LIVE. unfold liveness_checker_bool in EQL. + rewrite lazy_and_Some_true in EQL; destruct EQL as [[ibf ENTRY] LIST]. + eexists; split; eauto. +Qed. + Lemma liveness_checker_correct f n ibf: liveness_checker f = OK tt -> f.(fn_code)!n = Some ibf -> iblock_checker f.(fn_code) ibf.(entry) ibf.(input_regs) = Some tt. Proof. - unfold liveness_checker. - destruct list_iblock_checker eqn:EQL; try congruence. - intros _ P. exploit list_iblock_checker_correct; eauto. + intros LIVE PC. + apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]]. + exploit list_iblock_checker_correct; eauto. - eapply PTree.elements_correct; eauto. - simpl; auto. Qed. +Lemma liveness_checker_entrypoint f: + liveness_checker f = OK tt -> + f.(fn_code)!(f.(fn_entrypoint)) <> None. +Proof. + intros LIVE; apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]]. + unfold not; intros CONTRA. congruence. +Qed. + Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt. Inductive liveness_ok_fundef: fundef -> Prop := @@ -195,16 +219,79 @@ Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f - Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. -Lemma cfgsem2fsem_ibistep_simu sp rs1 m1 rs1' m1' of ib - (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of): forall - rs2 m2, - exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of). +(*Lemma tr_inputs_eqlive f rs1 rs2 tbl or r + (REGS: forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) + :eqlive_reg (tr_inputs f tbl or rs1) (tid f tbl or rs2). +Proof. + unfold tid; rewrite tr_inputs_get. + autodestruct. + Qed.*) + +Remark is_none_true {A} (fin: option A): + is_none fin = true -> + fin = None. +Proof. + unfold is_none; destruct fin; congruence. +Qed. + +Local Hint Resolve Regset.mem_2 Regset.subset_2: core. +Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. +Proof. + destruct b1; simpl; intuition. +Qed. + +Lemma list_mem_correct (rl: list reg) (alive: Regset.t): + list_mem rl alive = true -> forall r, List.In r rl -> ext alive r. +Proof. + induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto. +Qed. + +Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args. +Proof. + induction args; simpl; auto. + intros EQLIVE ALIVE; rewrite IHargs; auto. + unfold eqlive_reg in EQLIVE. + rewrite EQLIVE; auto. +Qed. + +Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args. Proof. - induction ISTEP; simpl; try_simplify_someHyps; intros. + intros; eapply eqlive_reg_list; eauto. + intros; eapply list_mem_correct; eauto. +Qed. + +Lemma cfgsem2fsem_ibistep_simu sp f: forall rs1 m rs1' m' rs2 of ibf res, + iblock_istep ge sp rs1 m ibf.(entry) rs1' m' of -> + (*iblock_checker (fn_code f) ibf.(entry) ibf.(input_regs) = Some tt ->*) + eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> + body_checker (fn_code f) ibf.(entry) ibf.(input_regs) = Some res -> +(*LIST : list_iblock_checker (fn_code f) (PTree.elements (fn_code f)) = true*) + exists rs2', iblock_istep_run ge sp ibf.(entry) rs2 m = Some (out rs2' m' of). +Proof. + induction 1; simpl; repeat inversion_ASSERT; intros; try_simplify_someHyps. - (* Bop *) - inversion_SOME v0; intros EVAL'; - repeat eexists. - inv EVAL'. + intros. + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + (* + (*- [> BF <]*) + (*intros. inv H1; eexists; reflexivity.*) + - (* Bseq stop *) + admit. (* + inversion_SOME res0. + inversion_ASSERT; intros. + exploit IHiblock_istep. unfold iblock_checker. + rewrite H1. inv H0. + inversion_SOME out1. + try_simplify_someHyps.*) + - inversion_SOME res0. + inversion_ASSERT; intros. + admit. + (* TODO The first induction hyp should not be based on iblock_checker *) + (* exploit IHiblock_istep1. unfold iblock_checker. + rewrite H2. apply is_none_true in H1. rewrite H1. *) + - (* Bcond *) + admit.*) Admitted. Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 @@ -215,7 +302,7 @@ Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 Proof. destruct FSTEP. - (* Bgoto *) - (*eexists; split; simpl ; econstructor; eauto.*) + eexists; split; simpl. econstructor; eauto. Admitted. Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: @@ -229,19 +316,21 @@ Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: eqlive_states s1 s2. Proof. intros STEP STACKS EQLIVE LIVE PC. - unfold liveness_ok_function, liveness_checker in LIVE. - destruct list_iblock_checker eqn:LIBC in LIVE; try discriminate. - clear LIVE. - destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP). + unfold liveness_ok_function in LIVE. + apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]]. + eapply PTree.elements_correct in PC as PCIN. + eapply list_iblock_checker_correct in LIST as IBC; eauto. + unfold iblock_checker in IBC. generalize IBC; clear IBC. + inversion_SOME res; inversion_SOME fin; intros RES BODY FINAL. + destruct STEP as (rs1' & m1' & fin' & ISTEP & FSTEP). exploit cfgsem2fsem_ibistep_simu; eauto. - intros (rs2' & m2' & ISTEP2). + intros (rs2' & ISTEP2). rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP. exploit cfgsem2fsem_finalstep_simu; eauto. intros (s2 & FSTEP2 & STATES). clear FSTEP. unfold iblock_step. repeat eexists; eauto. Qed. - Local Hint Constructors step: core. Lemma cfgsem2fsem_step_simu s1 s1' t s2: @@ -259,14 +348,16 @@ Proof. intros (s2 & STEP2 & EQUIV2). eexists; split; eauto. - inv STATES; inv LIVE. - eexists; split; econstructor; eauto. all : admit. (* TODO gourdinl *) + apply liveness_checker_entrypoint in H0 as ENTRY. + destruct ((fn_code f) ! (fn_entrypoint f)) eqn:EQENTRY; try congruence; eauto. + eexists; split; repeat econstructor; eauto. - inv STATES. inversion STACKS as [|s1 st1 s' s2 STACK STACKS']; subst; clear STACKS. inv STACK. exists (State s2 f sp pc (rs2 # res <- vres) m); split. * apply exec_return. * eapply eqlive_states_intro; eauto. -Admitted. +Qed. Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog). Proof. -- cgit From 0d856574d11ccccab397e007d43437980e07aeac Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 3 Jun 2021 08:31:20 +0200 Subject: progress in BTL_SEsimuref --- scheduling/BTL_SEsimuref.v | 136 +++++++++++++++++++-------------------------- scheduling/BTL_SEtheory.v | 85 +++++++++++++++++----------- 2 files changed, 112 insertions(+), 109 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index e9ae11e0..c322a9a5 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -89,14 +89,13 @@ Proof. - intros; erewrite <- eval_sval_preserved; eauto. Qed. -Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2 rs m: +Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2: ris_simu ris1 ris2 -> ris_refines (bctx1 ctx) ris1 sis1 -> ris_refines (bctx2 ctx) ris2 sis2 -> - sem_sistate (bctx1 ctx) sis1 rs m -> - sem_sistate (bctx2 ctx) sis2 rs m. + sistate_simu ctx sis1 sis2. Proof. - intros RIS REF1 REF2 SEM. + intros RIS REF1 REF2 rs m SEM. exploit sem_sok; eauto. erewrite OK_EQUIV; eauto. intros ROK1. @@ -113,62 +112,76 @@ Proof. erewrite REG_EQ; eauto. Qed. -(* TODO: - -Definition option_refines ctx (orsv: option sval) (osv: option sval) := - match orsv, osv with - | Some rsv, Some sv => eval_sval ctx rsv = eval_sval ctx sv - | None, None => True - | _, _ => False - end. +Inductive optrsv_refines ctx: (option sval) -> (option sval) -> Prop := + | RefSome rsv sv + (REF:eval_sval ctx rsv = eval_sval ctx sv) + :optrsv_refines ctx (Some rsv) (Some sv) + | RefNone: optrsv_refines ctx None None + . -Definition sum_refines ctx (rsi: sval + ident) (si: sval + ident) := - match rsi, si with - | inl rsv, inl sv => eval_sval ctx rsv = eval_sval ctx sv - | inr id, inr id' => id = id' - | _, _ => False - end. +Inductive rsvident_refines ctx: (sval + ident) -> (sval + ident) -> Prop := + | RefLeft rsv sv + (REF:eval_sval ctx rsv = eval_sval ctx sv) + :rsvident_refines ctx (inl rsv) (inl sv) + | RefRight id1 id2 + (IDSIMU: id1 = id2) + :rsvident_refines ctx (inr id1) (inr id2) + . Definition bargs_refines ctx (rargs: list (builtin_arg sval)) (args: list (builtin_arg sval)): Prop := eval_list_builtin_sval ctx rargs = eval_list_builtin_sval ctx args. Inductive rfv_refines ctx: sfval -> sfval -> Prop := - | refines_Sgoto pc: rfv_refines ctx (Sgoto pc) (Sgoto pc) - | refines_Scall: forall sig rvos ros rargs args r pc - (SUM:sum_refines ctx rvos ros) + | RefGoto pc: rfv_refines ctx (Sgoto pc) (Sgoto pc) + | RefCall sig rvos ros rargs args r pc + (SV:rsvident_refines ctx rvos ros) (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args) - ,rfv_refines ctx (Scall sig rvos rargs r pc) (Scall sig ros args r pc) - | refines_Stailcall: forall sig rvos ros rargs args - (SUM:sum_refines ctx rvos ros) + :rfv_refines ctx (Scall sig rvos rargs r pc) (Scall sig ros args r pc) + | RefTailcall sig rvos ros rargs args + (SV:rsvident_refines ctx rvos ros) (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args) - ,rfv_refines ctx (Stailcall sig rvos rargs) (Stailcall sig ros args) - | refines_Sbuiltin: forall ef lbra lba br pc + :rfv_refines ctx (Stailcall sig rvos rargs) (Stailcall sig ros args) + | RefBuiltin ef lbra lba br pc (BARGS: bargs_refines ctx lbra lba) - ,rfv_refines ctx (Sbuiltin ef lbra br pc) (Sbuiltin ef lba br pc) - | refines_Sjumptable: forall rsv sv lpc + :rfv_refines ctx (Sbuiltin ef lbra br pc) (Sbuiltin ef lba br pc) + | RefJumptable rsv sv lpc (VAL: eval_sval ctx rsv = eval_sval ctx sv) - ,rfv_refines ctx (Sjumptable rsv lpc) (Sjumptable sv lpc) - | refines_Sreturn: forall orsv osv - (OPT:option_refines ctx orsv osv) - ,rfv_refines ctx (Sreturn orsv) (Sreturn osv) + :rfv_refines ctx (Sjumptable rsv lpc) (Sjumptable sv lpc) + | RefReturn orsv osv + (OPT:optrsv_refines ctx orsv osv) + :rfv_refines ctx (Sreturn orsv) (Sreturn osv) . Definition rfv_simu (rfv1 rfv2: sfval): Prop := rfv1 = rfv2. -Local Hint Constructors sem_sfval equiv_state: core. +Local Hint Resolve eval_sval_preserved list_sval_eval_preserved smem_eval_preserved eval_list_builtin_sval_preserved: core. -Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context f1 f2) sfv1 sfv2 rs m t s: +Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context f1 f2) sfv1 sfv2: rfv_simu rfv1 rfv2 -> rfv_refines (bctx1 ctx) rfv1 sfv1 -> rfv_refines (bctx2 ctx) rfv2 sfv2 -> - sem_sfval (bctx1 ctx) sfv1 rs m t s -> - exists s', sem_sfval (bctx2 ctx) sfv2 rs m t s' /\ equiv_state s s'. + sfv_simu ctx sfv1 sfv2. Proof. - unfold rfv_simu; intros X REF1 REF2 SEM. subst. - unfold bctx2; destruct REF1; inv REF2; inv SEM; simpl. - - eexists; split; eauto; simpl. - (* eapply State_equiv; eauto. NE MARCHE PAS ! *) -Admitted. + unfold rfv_simu; intros X REF1 REF2. subst. + unfold bctx2; destruct REF1; inv REF2; simpl; econstructor; eauto. + - (* call svid *) + inv SV; inv SV0; econstructor; eauto. + rewrite <- REF, <- REF0; eauto. + - (* call args *) + rewrite <- LIST, <- LIST0; eauto. + - (* taillcall svid *) + inv SV; inv SV0; econstructor; eauto. + rewrite <- REF, <- REF0; eauto. + - (* tailcall args *) + rewrite <- LIST, <- LIST0; eauto. + - (* builtin args *) + unfold bargs_refines, bargs_simu in *. + rewrite <- BARGS, <- BARGS0; eauto. + - rewrite <- VAL, <- VAL0; eauto. + - (* return *) + inv OPT; inv OPT0; econstructor; eauto. + rewrite <- REF, <- REF0; eauto. +Qed. (* refinement of (symbolic) state *) Inductive rstate := @@ -197,18 +210,18 @@ Inductive rst_simu: rstate -> rstate -> Prop := (* Comment prend on en compte le ris en cours d'execution ??? *) Inductive rst_refines ctx: (* Prop -> *) rstate -> sstate -> Prop := - | refines_Sfinal ris sis rfv sfv (*ok: Prop*) + | Reffinal ris sis rfv sfv (*ok: Prop*) (*OK: ris_ok ctx ris -> ok*) (RIS: ris_refines ctx ris sis) (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) : rst_refines ctx (*ok*) (Rfinal ris rfv) (Sfinal sis sfv) - | refines_Scond cond rargs args sm rifso rifnot ifso ifnot (*ok1 ok2: Prop*) + | Refcond cond rargs args sm rifso rifnot ifso ifnot (*ok1 ok2: Prop*) (*OK1: ok2 -> ok1*) (RCOND: (* ok2 -> *) seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm) (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx (*ok2*) rifso ifso) (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx (*ok2*) rifnot ifnot) : rst_refines ctx (*ok1*) (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot) - | refines_Sabort + | Refabort : rst_refines ctx Rabort Sabort . @@ -223,37 +236,4 @@ Proof. induction 1; simpl; auto. - (* final *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. inv REF1. inv REF2. inv SEM1. - exploit sem_sok; eauto. - rewrite OK_EQUIV; eauto. - intros RIS_OK1. - exploit (ris_simu_ok_preserv f1 f2); eauto. - intros RIS_OK2. intuition. - exploit ris_simu_correct; eauto. - exploit rvf_simu_correct; eauto. - simpl. (* - eexists; split. - + econstructor; eauto. - simpl. - clear SIS. - admit. (* TODO: condition sur les tr_inputs du simu_proof_context ! *) - (* TODO: le rfv_refines est trop sémantique ! *) - + admit. - - (* cond *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. - inv REF1. inv REF2. inv SEM1. - destruct b; simpl. - + assert (TODO1: rst_refines (bctx1 ctx) rifso1 ifso). admit. - assert (TODO2: rst_refines (bctx2 ctx) rifso2 ifso0). admit. - exploit IHrst_simu1; eauto. - intros (s2 & X1 & X2). - exists s2; split; eauto. - econstructor; eauto. - * assert (TODO3: seval_condition (bctx2 ctx) cond args0 sm0 = Some true). admit. - eauto. - * simpl; eauto. - + admit. - - (* abort *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. - inv REF1. inv SEM1. -*) - -Admitted. -*) \ No newline at end of file +Admitted. \ No newline at end of file diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 7f00e46f..557541ce 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -235,16 +235,16 @@ Proof. eapply seval_builtin_arg_exact; eauto. Qed. -Fixpoint seval_builtin_sval ctx bsv := +Fixpoint eval_builtin_sval ctx bsv := match bsv with | BA sv => SOME v <- eval_sval ctx sv IN Some (BA v) | BA_splitlong sv1 sv2 => - SOME v1 <- seval_builtin_sval ctx sv1 IN - SOME v2 <- seval_builtin_sval ctx sv2 IN + SOME v1 <- eval_builtin_sval ctx sv1 IN + SOME v2 <- eval_builtin_sval ctx sv2 IN Some (BA_splitlong v1 v2) | BA_addptr sv1 sv2 => - SOME v1 <- seval_builtin_sval ctx sv1 IN - SOME v2 <- seval_builtin_sval ctx sv2 IN + SOME v1 <- eval_builtin_sval ctx sv1 IN + SOME v2 <- eval_builtin_sval ctx sv2 IN Some (BA_addptr v1 v2) | BA_int i => Some (BA_int i) | BA_long l => Some (BA_long l) @@ -259,7 +259,7 @@ Fixpoint seval_builtin_sval ctx bsv := Fixpoint eval_list_builtin_sval ctx lbsv := match lbsv with | nil => Some nil - | bsv::lbsv => SOME v <- seval_builtin_sval ctx bsv IN + | bsv::lbsv => SOME v <- eval_builtin_sval ctx bsv IN SOME lv <- eval_list_builtin_sval ctx lbsv IN Some (v::lv) end. @@ -271,9 +271,9 @@ Proof. destruct lbs2; simpl; repeat autodestruct; congruence. Qed. -Lemma seval_builtin_sval_arg ctx bs: +Lemma eval_builtin_sval_arg ctx bs: forall ba m v, - seval_builtin_sval ctx bs = Some ba -> + eval_builtin_sval ctx bs = Some ba -> eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v -> seval_builtin_arg ctx m bs v. Proof. @@ -286,14 +286,14 @@ Proof. intros H; inversion H; subst. econstructor; auto. - intros ba m v. - destruct (seval_builtin_sval _ bs1) eqn: SV1; try congruence. - destruct (seval_builtin_sval _ bs2) eqn: SV2; try congruence. + destruct (eval_builtin_sval _ bs1) eqn: SV1; try congruence. + destruct (eval_builtin_sval _ bs2) eqn: SV2; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst. econstructor; eauto. - intros ba m v. - destruct (seval_builtin_sval _ bs1) eqn: SV1; try congruence. - destruct (seval_builtin_sval _ bs2) eqn: SV2; try congruence. + destruct (eval_builtin_sval _ bs1) eqn: SV1; try congruence. + destruct (eval_builtin_sval _ bs2) eqn: SV2; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst. econstructor; eauto. @@ -302,7 +302,7 @@ Qed. Lemma seval_builtin_arg_sval ctx m v: forall bs, seval_builtin_arg ctx m bs v -> exists ba, - seval_builtin_sval ctx bs = Some ba + eval_builtin_sval ctx bs = Some ba /\ eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v. Proof. induction 1. @@ -323,7 +323,7 @@ Proof. + constructor; assumption. Qed. -Lemma seval_builtin_sval_args ctx lbs: +Lemma eval_builtin_sval_args ctx lbs: forall lba m v, eval_list_builtin_sval ctx lbs = Some lba -> list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba v -> @@ -332,12 +332,12 @@ Proof. unfold seval_builtin_args; induction lbs; simpl; intros lba m v. - intros H; inversion H; subst; clear H. intros H; inversion H. econstructor. - - destruct (seval_builtin_sval _ _) eqn:SV; try congruence. + - destruct (eval_builtin_sval _ _) eqn:SV; try congruence. destruct (eval_list_builtin_sval _ _) eqn: SVL; try congruence. intros H; inversion H; subst; clear H. intros H; inversion H; subst; clear H. econstructor; eauto. - eapply seval_builtin_sval_arg; eauto. + eapply eval_builtin_sval_arg; eauto. Qed. Lemma seval_builtin_args_sval ctx m lv: forall lbs, @@ -357,14 +357,14 @@ Proof. + constructor; assumption. Qed. -Lemma seval_builtin_sval_correct ctx m: forall bs1 v bs2, +Lemma eval_builtin_sval_correct ctx m: forall bs1 v bs2, seval_builtin_arg ctx m bs1 v -> - (seval_builtin_sval ctx bs1) = (seval_builtin_sval ctx bs2) -> + (eval_builtin_sval ctx bs1) = (eval_builtin_sval ctx bs2) -> seval_builtin_arg ctx m bs2 v. Proof. intros. exploit seval_builtin_arg_sval; eauto. intros (ba & X1 & X2). - eapply seval_builtin_sval_arg; eauto. + eapply eval_builtin_sval_arg; eauto. congruence. Qed. @@ -375,7 +375,7 @@ Lemma eval_list_builtin_sval_correct ctx m vargs: forall lbs1, Proof. intros. exploit seval_builtin_args_sval; eauto. intros (ba & X1 & X2). - eapply seval_builtin_sval_args; eauto. + eapply eval_builtin_sval_args; eauto. congruence. Qed. @@ -1110,19 +1110,19 @@ Definition bctx2 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge2) f2 ct *) Inductive optsv_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (option sval) -> (option sval) -> Prop := - | Ssome_simu sv1 sv2: - eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2 - -> optsv_simu ctx (Some sv1) (Some sv2) + | Ssome_simu sv1 sv2 + (SIMU:eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2) + :optsv_simu ctx (Some sv1) (Some sv2) | Snone_simu: optsv_simu ctx None None . Inductive svident_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (sval + ident) -> (sval + ident) -> Prop := - | Sleft_simu sv1 sv2: - eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2 - -> svident_simu ctx (inl sv1) (inl sv2) - | Sright_simu id1 id2: - id1 = id2 - -> svident_simu ctx (inr id1) (inr id2) + | Sleft_simu sv1 sv2 + (SIMU:eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2) + :svident_simu ctx (inl sv1) (inl sv2) + | Sright_simu id1 id2 + (IDSIMU: id1 = id2) + :svident_simu ctx (inr id1) (inr id2) . Definition bargs_simu {f1 f2: function} (ctx: simu_proof_context f1 f2) (args1 args2: list (builtin_arg sval)): Prop := @@ -1149,9 +1149,14 @@ Inductive sfv_simu {f1 f2} (ctx: simu_proof_context f1 f2): sfval -> sfval -> Pr :sfv_simu ctx (Sreturn osv1) (Sreturn osv2) . +Definition sistate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (sis1 sis2:sistate): Prop := + forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sem_sistate (bctx2 ctx) sis2 rs m. + Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop := forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> - exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) /\ sfv_simu ctx sfv1 sfv2 + exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) + /\ sistate_simu ctx sis1 sis2 + /\ sfv_simu ctx sfv1 sfv2 . Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2). @@ -1182,7 +1187,7 @@ Proof. Qed. *) -(** * Preservation properties *) +(** * Preservation properties under a [simu_proof_context] *) Section SymbValPreserved. @@ -1231,6 +1236,23 @@ Proof. rewrite eval_sval_preserved; auto. Qed. +Lemma eval_builtin_sval_preserved sv: + eval_builtin_sval (bctx1 ctx) sv = eval_builtin_sval (bctx2 ctx) sv. +Proof. + induction sv; simpl; auto. + all: try (erewrite eval_sval_preserved by eauto); trivial. + all: erewrite IHsv1 by eauto; erewrite IHsv2 by eauto; reflexivity. +Qed. + +Lemma eval_list_builtin_sval_preserved lsv: + eval_list_builtin_sval (bctx1 ctx) lsv = eval_list_builtin_sval (bctx2 ctx) lsv. +Proof. + induction lsv; simpl; auto. + erewrite eval_builtin_sval_preserved by eauto. + erewrite IHlsv by eauto. + reflexivity. +Qed. + Lemma seval_condition_preserved cond lsv sm: seval_condition (bctx1 ctx) cond lsv sm = seval_condition (bctx2 ctx) cond lsv sm. Proof. @@ -1239,6 +1261,7 @@ Proof. rewrite smem_eval_preserved; auto. Qed. +(* additional preservation properties under this additional hypothesis *) Hypothesis senv_preserved_BTL: Senv.equiv (sge1 ctx) (sge2 ctx). Lemma senv_find_symbol_preserved id: -- cgit From b65a646a9ff0d16e65923d225b75942cc766bc02 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 3 Jun 2021 13:41:59 +0200 Subject: update a more general version of the liveness checker --- scheduling/BTL_Livecheck.v | 211 +++++++++++++++++++++++++-------------------- 1 file changed, 118 insertions(+), 93 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 82efff44..d5aad8b8 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -59,40 +59,34 @@ Definition final_inst_checker (btl: code) (alive: Regset.t) (fin: final): option ASSERT exit_list_checker btl alive tbl IN Some tt end. -Definition is_none {A} (oa: option A): bool := - match oa with - | Some _ => false - | None => true - end. - -Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option (Regset.t*option final) := +Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option Regset.t := match ib with | Bseq ib1 ib2 => - SOME res <- body_checker btl ib1 alive IN - ASSERT is_none (snd res) IN - body_checker btl ib2 (fst res) - | Bnop _ => Some (alive, None) + SOME alive1 <- body_checker btl ib1 alive IN + body_checker btl ib2 alive1 + | Bnop _ => Some alive | Bop _ args dest _ => ASSERT list_mem args alive IN - Some (Regset.add dest alive, None) + Some (Regset.add dest alive) | Bload _ _ _ args dest _ => ASSERT list_mem args alive IN - Some (Regset.add dest alive, None) + Some (Regset.add dest alive) | Bstore _ _ args src _ => ASSERT Regset.mem src alive IN ASSERT list_mem args alive IN - Some (alive, None) - | Bcond _ args (BF (Bgoto s) _) (Bnop None) _ => + Some alive + | Bcond _ args ib1 ib2 _ => ASSERT list_mem args alive IN - exit_checker btl alive s (alive, None) - | BF fin _ => Some (alive, Some fin) - | _ => None + SOME alive1 <- body_checker btl ib1 alive IN + SOME alive2 <- body_checker btl ib2 alive IN + Some (Regset.union alive1 alive2) + | BF fin _ => + SOME _ <- final_inst_checker btl alive fin IN + Some Regset.empty end. -Definition iblock_checker (btl: code) (ib: iblock) (alive: Regset.t) : option unit := - SOME res <- body_checker btl ib alive IN - SOME fin <- snd res IN - final_inst_checker btl (fst res) fin. +Definition iblock_checker (btl: code) (ib: iblock) (alive: Regset.t): option unit := + SOME _ <- body_checker btl ib alive IN Some tt. Fixpoint list_iblock_checker (btl: code) (l: list (node*iblock_info)): bool := match l with @@ -113,8 +107,9 @@ Proof. destruct x; auto. Qed. -Lemma list_iblock_checker_correct btl l: - list_iblock_checker btl l = true -> forall e, List.In e l -> iblock_checker btl (snd e).(entry) (snd e).(input_regs) = Some tt. +Lemma list_iblock_checker_correct btl l: + list_iblock_checker btl l = true -> + forall e, List.In e l -> iblock_checker btl (snd e).(entry) (snd e).(input_regs) = Some tt. Proof. intros CHECKER e H; induction l as [|(n & ibf) l]; intuition. simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto). @@ -181,9 +176,53 @@ Proof. * eapply Regset.add_2; auto. Qed. +Local Hint Resolve Regset.mem_2 Regset.subset_2: core. +Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. +Proof. + destruct b1; simpl; intuition. +Qed. + +Lemma list_mem_correct (rl: list reg) (alive: Regset.t): + list_mem rl alive = true -> forall r, List.In r rl -> ext alive r. +Proof. + induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto. +Qed. + Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := forall r, (alive r) -> rs1#r = rs2#r. +Lemma eqlive_reg_empty rs1 rs2: eqlive_reg (ext Regset.empty) rs1 rs2. +Proof. + unfold eqlive_reg; intros. inv H. +Qed. + +Lemma eqlive_reg_update (alive: Regset.elt -> Prop) rs1 rs2 r v: eqlive_reg (fun r1 => r1 <> r /\ alive r1) rs1 rs2 -> eqlive_reg alive (rs1 # r <- v) (rs2 # r <- v). +Proof. + unfold eqlive_reg; intros EQLIVE r0 ALIVE. + destruct (Pos.eq_dec r r0) as [H|H]. + - subst. rewrite! Regmap.gss. auto. + - rewrite! Regmap.gso; auto. +Qed. + +Lemma eqlive_reg_monotonic (alive1 alive2: Regset.elt -> Prop) rs1 rs2: eqlive_reg alive2 rs1 rs2 -> (forall r, alive1 r -> alive2 r) -> eqlive_reg alive1 rs1 rs2. +Proof. + unfold eqlive_reg; intuition. +Qed. + +Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args. +Proof. + induction args; simpl; auto. + intros EQLIVE ALIVE; rewrite IHargs; auto. + unfold eqlive_reg in EQLIVE. + rewrite EQLIVE; auto. +Qed. + +Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args. +Proof. + intros; eapply eqlive_reg_list; eauto. + intros; eapply list_mem_correct; eauto. +Qed. + Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := | eqlive_stackframes_intro ibf res f sp pc rs1 rs2 (LIVE: liveness_ok_function f) @@ -218,6 +257,7 @@ Let ge := Genv.globalenv prog. Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. +Local Hint Resolve eqlive_reg_empty: core. (*Lemma tr_inputs_eqlive f rs1 rs2 tbl or r (REGS: forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) @@ -227,82 +267,67 @@ Proof. autodestruct. Qed.*) -Remark is_none_true {A} (fin: option A): - is_none fin = true -> - fin = None. -Proof. - unfold is_none; destruct fin; congruence. -Qed. - -Local Hint Resolve Regset.mem_2 Regset.subset_2: core. -Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. -Proof. - destruct b1; simpl; intuition. -Qed. - -Lemma list_mem_correct (rl: list reg) (alive: Regset.t): - list_mem rl alive = true -> forall r, List.In r rl -> ext alive r. -Proof. - induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto. -Qed. - -Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args. -Proof. - induction args; simpl; auto. - intros EQLIVE ALIVE; rewrite IHargs; auto. - unfold eqlive_reg in EQLIVE. - rewrite EQLIVE; auto. -Qed. - -Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args. -Proof. - intros; eapply eqlive_reg_list; eauto. - intros; eapply list_mem_correct; eauto. -Qed. - -Lemma cfgsem2fsem_ibistep_simu sp f: forall rs1 m rs1' m' rs2 of ibf res, - iblock_istep ge sp rs1 m ibf.(entry) rs1' m' of -> +Lemma cfgsem2fsem_ibistep_simu sp f: forall rs1 m rs1' m' of ib, + iblock_istep ge sp rs1 m ib rs1' m' of -> (*iblock_checker (fn_code f) ibf.(entry) ibf.(input_regs) = Some tt ->*) - eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> - body_checker (fn_code f) ibf.(entry) ibf.(input_regs) = Some res -> + forall alive1 alive2 rs2, eqlive_reg (ext alive1) rs1 rs2 -> + body_checker (fn_code f) ib alive1 = Some alive2 -> (*LIST : list_iblock_checker (fn_code f) (PTree.elements (fn_code f)) = true*) - exists rs2', iblock_istep_run ge sp ibf.(entry) rs2 m = Some (out rs2' m' of). + exists rs2', iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' of) + /\ eqlive_reg (ext alive2) rs1' rs2'. Proof. - induction 1; simpl; repeat inversion_ASSERT; intros; try_simplify_someHyps. + induction 1; simpl; try_simplify_someHyps; + repeat inversion_ASSERT; intros. + - (* BF *) + generalize H0; clear H0; inversion_SOME x; try_simplify_someHyps. - (* Bop *) - intros. erewrite <- eqlive_reg_listmem; eauto. try_simplify_someHyps. - (* - (*- [> BF <]*) - (*intros. inv H1; eexists; reflexivity.*) + repeat econstructor. + apply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + - (* Bload *) + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps; intros. + rewrite LOAD; eauto. + repeat econstructor. + apply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + - (* Bstore *) + erewrite <- eqlive_reg_listmem; eauto. + rewrite <- (H src); auto. + try_simplify_someHyps; intros. + rewrite STORE; eauto. - (* Bseq stop *) - admit. (* - inversion_SOME res0. - inversion_ASSERT; intros. - exploit IHiblock_istep. unfold iblock_checker. - rewrite H1. inv H0. - inversion_SOME out1. - try_simplify_someHyps.*) - - inversion_SOME res0. - inversion_ASSERT; intros. - admit. - (* TODO The first induction hyp should not be based on iblock_checker *) - (* exploit IHiblock_istep1. unfold iblock_checker. - rewrite H2. apply is_none_true in H1. rewrite H1. *) + generalize H1; clear H1. + inversion_SOME aliveMid. intros BDY1 BDY2. + (*rewrite iblock_istep_run_equiv in H.*) + exploit IHiblock_istep; eauto. + intros (rs2' & ISTEP1 & REGS). rewrite ISTEP1; simpl. + repeat econstructor. admit. + - (* Bseq continue *) + generalize H2; clear H2. + inversion_SOME aliveMid; intros BDY1 BDY2. + exploit IHiblock_istep1; eauto. + intros (rs2' & ISTEP1 & REGS1). rewrite ISTEP1; simpl. + eapply IHiblock_istep2; eauto. - (* Bcond *) - admit.*) + generalize H2; clear H2. + inversion_SOME aliveSo; inversion_SOME aliveNot; intros BDY1 BDY2 UNION. + inv UNION. erewrite <- eqlive_reg_listmem; eauto. + admit. Admitted. -Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2 - (FSTEP: final_step tid ge stk1 f sp rs1 m1 fin t s1) - (STACKS: list_forall2 eqlive_stackframes stk1 stk2) - : exists s2, final_step tr_inputs ge stk2 f sp rs2 m2 fin t s2 - /\ eqlive_states s1 s2. +Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2: + final_step tid ge stk1 f sp rs1 m1 fin t s1 -> + list_forall2 eqlive_stackframes stk1 stk2 -> + exists s2, final_step tr_inputs ge stk2 f sp rs2 m2 fin t s2 + /\ eqlive_states s1 s2. Proof. - destruct FSTEP. - - (* Bgoto *) - eexists; split; simpl. econstructor; eauto. Admitted. Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: @@ -321,11 +346,11 @@ Proof. eapply PTree.elements_correct in PC as PCIN. eapply list_iblock_checker_correct in LIST as IBC; eauto. unfold iblock_checker in IBC. generalize IBC; clear IBC. - inversion_SOME res; inversion_SOME fin; intros RES BODY FINAL. - destruct STEP as (rs1' & m1' & fin' & ISTEP & FSTEP). + inversion_SOME x; intros BODY _. + destruct STEP as (rs1' & m1' & fin' & ISTEP1 & FSTEP). exploit cfgsem2fsem_ibistep_simu; eauto. - intros (rs2' & ISTEP2). - rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP. + intros (rs2' & ISTEP2 & REGS). + rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP1. exploit cfgsem2fsem_finalstep_simu; eauto. intros (s2 & FSTEP2 & STATES). clear FSTEP. unfold iblock_step. repeat eexists; eauto. -- cgit From 3348c88e1d61d109a8c0388ec421ac6bf17a5c6b Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 4 Jun 2021 09:07:45 +0200 Subject: progress in BTL_SEsimuref --- scheduling/BTL_SEsimuref.v | 56 ++++++++++++++++++++++++++++++++++++---------- scheduling/BTL_SEtheory.v | 34 +++++++++------------------- scheduling/BTL_Scheduler.v | 22 ++++++++++++++++++ 3 files changed, 77 insertions(+), 35 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index c322a9a5..a4da4291 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -191,14 +191,14 @@ Inductive rstate := . Inductive rst_simu: rstate -> rstate -> Prop := - | Rfinal_simu ris1 ris2 rfv1 rfv2: - ris_simu ris1 ris2 -> - rfv_simu rfv1 rfv2 -> - rst_simu (Rfinal ris1 rfv1) (Rfinal ris2 rfv2) - | Rcond_simu cond rargs rifso1 rifnot1 rifso2 rifnot2: - rst_simu rifso1 rifso2 -> - rst_simu rifnot1 rifnot2 -> - rst_simu (Rcond cond rargs rifso1 rifnot1) (Rcond cond rargs rifso2 rifnot2) + | Rfinal_simu ris1 ris2 rfv1 rfv2 + (RIS: ris_simu ris1 ris2) + (RFV: rfv_simu rfv1 rfv2) + : rst_simu (Rfinal ris1 rfv1) (Rfinal ris2 rfv2) + | Rcond_simu cond rargs rifso1 rifnot1 rifso2 rifnot2 + (IFSO: rst_simu rifso1 rifso2) + (IFNOT: rst_simu rifnot1 rifnot2) + : rst_simu (Rcond cond rargs rifso1 rifnot1) (Rcond cond rargs rifso2 rifnot2) | Rabort_simu: rst_simu Rabort Rabort (* TODO: extension à voir dans un second temps ! | Rcond_skip cond rargs rifso1 rifnot1 rst: @@ -225,6 +225,8 @@ Inductive rst_refines ctx: (* Prop -> *) rstate -> sstate -> Prop := : rst_refines ctx Rabort Sabort . +Local Hint Resolve ris_simu_correct rvf_simu_correct: core. + Lemma rst_simu_correct rst1 rst2: rst_simu rst1 rst2 -> forall f1 f2 (ctx: simu_proof_context f1 f2) st1 st2 (*ok1 ok2*), @@ -233,7 +235,37 @@ Lemma rst_simu_correct rst1 rst2: (* ok1 -> ok2 -> *) sstate_simu ctx st1 st2. Proof. - induction 1; simpl; auto. - - (* final *) intros f1 f2 ctx st1 st2 REF1 REF2 t s1 SEM1. - inv REF1. inv REF2. inv SEM1. -Admitted. \ No newline at end of file + induction 1; simpl; intros f1 f2 ctx st1 st2 REF1 REF2 sis1 svf1 SEM1; + inv REF1; inv REF2; simpl in *; inv SEM1; auto. + - (* final_simu *) + do 2 eexists; intuition; eauto. + exploit sem_sok; eauto. + erewrite OK_EQUIV; eauto. + intros ROK1. + exploit ris_simu_ok_preserv; eauto. + - (* cond_simu *) + destruct (seval_condition (bctx1 ctx) cond args sm) eqn: SCOND; try congruence. + generalize RCOND0. + erewrite <- seval_condition_preserved, RCOND by eauto. + intros SCOND0; rewrite <- SCOND0 in RCOND0. + erewrite <- SCOND0. + destruct b; simpl. + * exploit IHrst_simu1; eauto. + * exploit IHrst_simu2; eauto. +Qed. + + +(** TODO: could be useful ? +Lemma seval_condition_valid_preserv ctx cond args sm b + (VALID_PRESERV: forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) + :seval_condition ctx cond args sm = Some b -> + seval_condition ctx cond args Sinit = Some b. +Proof. + unfold seval_condition; autodestruct; simpl; try congruence. + intros EVAL_LIST. + autodestruct; try congruence. + intros MEM COND. rewrite <- COND. + eapply cond_valid_pointer_eq; intros. + exploit VALID_PRESERV; eauto. +Qed. +*) \ No newline at end of file diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 557541ce..9ab64d7d 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1064,32 +1064,10 @@ Qed. (** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *) -(* -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)). - -Lemma equiv_input_regs_union f1 f2: - equiv_input_regs f1 f2 -> forall tbl, union_inputs f1 tbl = union_inputs f2 tbl. -Proof. - intros (EQNone & EQSome). induction tbl as [|pc l']; simpl; auto. - generalize (EQNone pc) (EQSome pc); clear EQNone EQSome; intros EQN EQS. - do 2 autodestruct; intuition; try_simplify_someHyps. - intros; exploit EQS; eauto; clear EQS. congruence. -Qed. - -Lemma equiv_input_regs_pre f1 f2 tbl or: - equiv_input_regs f1 f2 -> pre_inputs f1 tbl or = pre_inputs f2 tbl or. -Proof. - intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto. -Qed. -*) - Record simu_proof_context {f1 f2: BTL.function} := Sctx { sge1: BTL.genv; sge2: BTL.genv; sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s; - (* inputs_equiv: equiv_input_regs f1 f2;*) (* TODO a-t-on besoin de ça ? *) ssp: val; srs0: regset; sm0: mem @@ -1156,7 +1134,7 @@ Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) /\ sistate_simu ctx sis1 sis2 - /\ sfv_simu ctx sfv1 sfv2 + /\ (forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sfv_simu ctx sfv1 sfv2) . Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2). @@ -1261,6 +1239,16 @@ Proof. rewrite smem_eval_preserved; auto. Qed. +(* TODO: useless ? +Lemma run_sem_isstate_preserved sis: + run_sem_isstate (bctx1 ctx) sis = run_sem_isstate (bctx2 ctx) sis. +Proof. + induction sis; simpl; eauto. + erewrite seval_condition_preserved. + repeat (autodestruct; auto). +Qed. +*) + (* additional preservation properties under this additional hypothesis *) Hypothesis senv_preserved_BTL: Senv.equiv (sge1 ctx) (sge2 ctx). diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index bcddcf5d..ec83b3c1 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -9,6 +9,28 @@ 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)). + +Lemma equiv_input_regs_union f1 f2: + equiv_input_regs f1 f2 -> forall tbl, union_inputs f1 tbl = union_inputs f2 tbl. +Proof. + intros (EQNone & EQSome). induction tbl as [|pc l']; simpl; auto. + generalize (EQNone pc) (EQSome pc); clear EQNone EQSome; intros EQN EQS. + do 2 autodestruct; intuition; try_simplify_someHyps. + intros; exploit EQS; eauto; clear EQS. congruence. +Qed. + +Lemma equiv_input_regs_pre f1 f2 tbl or: + equiv_input_regs f1 f2 -> pre_inputs f1 tbl or = pre_inputs f2 tbl or. +Proof. + intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto. +Qed. +*) + + (* a specification of the verification to do on each function *) Record match_function (f1 f2: BTL.function): Prop := { preserv_fnsig: fn_sig f1 = fn_sig f2; -- cgit From 640cdb752b812f2e5301acbfc63c2d4fa798de06 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 4 Jun 2021 18:47:40 +0200 Subject: Some advance in Liveness lemmas --- scheduling/BTL_Livecheck.v | 209 ++++++++++++++++++++++++++++++--------------- 1 file changed, 138 insertions(+), 71 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index d5aad8b8..95b6a155 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -2,7 +2,6 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep Op Registers OptionMonad. Require Import Errors RTL BTL BTLmatchRTL. -(** TODO: adapt stuff RTLpathLivegen *) Local Open Scope lazy_bool_scope. @@ -59,32 +58,44 @@ Definition final_inst_checker (btl: code) (alive: Regset.t) (fin: final): option ASSERT exit_list_checker btl alive tbl IN Some tt end. -Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option Regset.t := +(* This definition is the meet (infimum) subset of alive registers, + used for conditions by the below checker. + A None argument represents the neutral element for intersection. *) +Definition meet (o1 o2: option Regset.t): option Regset.t := + match o1, o2 with + | None, _ => o2 + | _, None => o1 + | Some alive1, Some alive2 => Some (Regset.inter alive1 alive2) + end. + +Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option (option Regset.t) := match ib with | Bseq ib1 ib2 => - SOME alive1 <- body_checker btl ib1 alive IN + SOME oalive1 <- body_checker btl ib1 alive IN + SOME alive1 <- oalive1 IN body_checker btl ib2 alive1 - | Bnop _ => Some alive + | Bnop _ => Some (Some alive) | Bop _ args dest _ => ASSERT list_mem args alive IN - Some (Regset.add dest alive) + Some (Some (Regset.add dest alive)) | Bload _ _ _ args dest _ => ASSERT list_mem args alive IN - Some (Regset.add dest alive) + Some (Some (Regset.add dest alive)) | Bstore _ _ args src _ => ASSERT Regset.mem src alive IN ASSERT list_mem args alive IN - Some alive + Some (Some alive) | Bcond _ args ib1 ib2 _ => ASSERT list_mem args alive IN - SOME alive1 <- body_checker btl ib1 alive IN - SOME alive2 <- body_checker btl ib2 alive IN - Some (Regset.union alive1 alive2) + SOME oalive1 <- body_checker btl ib1 alive IN + SOME oalive2 <- body_checker btl ib2 alive IN + Some (meet oalive1 oalive2) | BF fin _ => SOME _ <- final_inst_checker btl alive fin IN - Some Regset.empty + Some None end. +(* This definition simply convert the result in an option unit *) Definition iblock_checker (btl: code) (ib: iblock) (alive: Regset.t): option unit := SOME _ <- body_checker btl ib alive IN Some tt. @@ -115,7 +126,6 @@ Proof. simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto). Qed. - Definition liveness_checker_bool (f: BTL.function): bool := f.(fn_code)!(f.(fn_entrypoint)) &&& list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)). @@ -163,10 +173,25 @@ Inductive liveness_ok_fundef: fundef -> Prop := | liveness_ok_Internal f: liveness_ok_function f -> liveness_ok_fundef (Internal f) | liveness_ok_External ef: liveness_ok_fundef (External ef). -(** TODO: adapt stuff RTLpathLivegenproof *) Local Notation ext alive := (fun r => Regset.In r alive). +Definition ext_opt (oalive: option Regset.t): Regset.elt -> Prop := + match oalive with + | Some alive => ext alive + | None => fun _ => True + end. + +Lemma ext_opt_meet: forall r oalive1 oalive2, + ext_opt (meet oalive1 oalive2) r -> + ext_opt oalive1 r /\ ext_opt oalive2 r. +Proof. + intros. destruct oalive1, oalive2; + simpl in *; intuition. + eapply Regset.inter_1; eauto. + eapply Regset.inter_2; eauto. +Qed. + Lemma regset_add_spec live r1 r2: Regset.In r1 (Regset.add r2 live) <-> (r1 = r2 \/ Regset.In r1 live). Proof. destruct (Pos.eq_dec r1 r2). @@ -177,6 +202,7 @@ Proof. Qed. Local Hint Resolve Regset.mem_2 Regset.subset_2: core. + Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. Proof. destruct b1; simpl; intuition. @@ -191,11 +217,6 @@ Qed. Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := forall r, (alive r) -> rs1#r = rs2#r. -Lemma eqlive_reg_empty rs1 rs2: eqlive_reg (ext Regset.empty) rs1 rs2. -Proof. - unfold eqlive_reg; intros. inv H. -Qed. - Lemma eqlive_reg_update (alive: Regset.elt -> Prop) rs1 rs2 r v: eqlive_reg (fun r1 => r1 <> r /\ alive r1) rs1 rs2 -> eqlive_reg alive (rs1 # r <- v) (rs2 # r <- v). Proof. unfold eqlive_reg; intros EQLIVE r0 ALIVE. @@ -246,8 +267,6 @@ Inductive eqlive_states: state -> state -> Prop := (STACKS: list_forall2 eqlive_stackframes st1 st2): eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). -(* continue to adapt stuff RTLpathLivegenproof *) - Section FSEM_SIMULATES_CFGSEM. Variable prog: BTL.program. @@ -257,7 +276,6 @@ Let ge := Genv.globalenv prog. Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. -Local Hint Resolve eqlive_reg_empty: core. (*Lemma tr_inputs_eqlive f rs1 rs2 tbl or r (REGS: forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) @@ -267,22 +285,39 @@ Proof. autodestruct. Qed.*) -Lemma cfgsem2fsem_ibistep_simu sp f: forall rs1 m rs1' m' of ib, - iblock_istep ge sp rs1 m ib rs1' m' of -> - (*iblock_checker (fn_code f) ibf.(entry) ibf.(input_regs) = Some tt ->*) - forall alive1 alive2 rs2, eqlive_reg (ext alive1) rs1 rs2 -> - body_checker (fn_code f) ib alive1 = Some alive2 -> -(*LIST : list_iblock_checker (fn_code f) (PTree.elements (fn_code f)) = true*) - exists rs2', iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' of) - /\ eqlive_reg (ext alive2) rs1' rs2'. +Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m fin t s1 stk2 rs2 + (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s1) + (LIVE: liveness_ok_function f) + (STACKS: list_forall2 eqlive_stackframes stk1 stk2) + (*body_checker (fn_code f) (entry ibf) (input_regs ibf) = Some x ->*) + (*eqlive_reg (ext x) rs1' rs2' ->*) + :exists s2, + final_step tr_inputs ge stk2 f sp rs2 m fin t s2 + /\ eqlive_states s1 s2. Proof. - induction 1; simpl; try_simplify_someHyps; - repeat inversion_ASSERT; intros. - - (* BF *) - generalize H0; clear H0; inversion_SOME x; try_simplify_someHyps. + destruct FSTEP. + - (* Bgoto *) + econstructor. econstructor. + econstructor. econstructor. + eauto. eauto. + repeat econstructor; eauto. +Admitted. + +Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m' + (ISTEP: iblock_istep ge sp rs1 m ib rs1' m' None) + alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2) + (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2)), + exists rs2', + iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' None) + /\ eqlive_reg (ext_opt oalive2) rs1' rs2'. +Proof. + induction ib; intros; try_simplify_someHyps; + repeat inversion_ASSERT; intros; inv ISTEP. + - (* Bnop *) + inv BDY; eauto. - (* Bop *) erewrite <- eqlive_reg_listmem; eauto. - try_simplify_someHyps. + try_simplify_someHyps; intros. repeat econstructor. apply eqlive_reg_update. eapply eqlive_reg_monotonic; eauto. @@ -299,36 +334,65 @@ Proof. intuition. - (* Bstore *) erewrite <- eqlive_reg_listmem; eauto. - rewrite <- (H src); auto. + rewrite <- (REGS src); auto. try_simplify_someHyps; intros. rewrite STORE; eauto. - - (* Bseq stop *) - generalize H1; clear H1. - inversion_SOME aliveMid. intros BDY1 BDY2. - (*rewrite iblock_istep_run_equiv in H.*) - exploit IHiblock_istep; eauto. - intros (rs2' & ISTEP1 & REGS). rewrite ISTEP1; simpl. - repeat econstructor. admit. - (* Bseq continue *) - generalize H2; clear H2. - inversion_SOME aliveMid; intros BDY1 BDY2. - exploit IHiblock_istep1; eauto. + destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate. + generalize BDY; clear BDY. + inversion_SOME aliveMid; intros OALIVE BDY2. inv OALIVE. + exploit IHib1; eauto. intros (rs2' & ISTEP1 & REGS1). rewrite ISTEP1; simpl. - eapply IHiblock_istep2; eauto. + eapply IHib2; eauto. - (* Bcond *) - generalize H2; clear H2. - inversion_SOME aliveSo; inversion_SOME aliveNot; intros BDY1 BDY2 UNION. - inv UNION. erewrite <- eqlive_reg_listmem; eauto. - admit. -Admitted. + generalize BDY; clear BDY. + inversion_SOME oaliveSo; inversion_SOME oaliveNot; intros BDY1 BDY2 JOIN. + erewrite <- eqlive_reg_listmem; eauto. + rewrite EVAL. + destruct b; [ exploit IHib1; eauto | exploit IHib2; eauto]. + all: + intros (rs2' & ISTEP1 & REGS1); + econstructor; split; eauto; inv JOIN; + eapply eqlive_reg_monotonic; eauto; + intros r EXTM; apply ext_opt_meet in EXTM; intuition. +Qed. -Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2: - final_step tid ge stk1 f sp rs1 m1 fin t s1 -> - list_forall2 eqlive_stackframes stk1 stk2 -> - exists s2, final_step tr_inputs ge stk2 f sp rs2 m2 fin t s2 - /\ eqlive_states s1 s2. +Lemma cfgsem2fsem_ibistep_simu_Some sp f stk1 stk2 ib: forall s t rs1 m rs1' m' fin + (ISTEP: iblock_istep ge sp rs1 m ib rs1' m' (Some fin)) + (FSTEP: final_step tid ge stk1 f sp rs1' m' fin t s) + alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2) + (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2)) + (LIVE: liveness_ok_function f) + (STACKS: list_forall2 eqlive_stackframes stk1 stk2), + exists rs2' s', + iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' (Some fin)) + /\ final_step tr_inputs ge stk2 f sp rs2' m' fin t s' + /\ eqlive_states s s'. Proof. -Admitted. + induction ib; simpl; try_simplify_someHyps; + repeat inversion_ASSERT; intros; inv ISTEP. + - (* BF *) + exploit cfgsem2fsem_finalstep_simu; eauto. + intros (s2 & FSTEP' & STATES). eauto. + - (* Bseq stop *) + destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate. + generalize BDY; clear BDY. + inversion_SOME aliveMid. intros OALIVE BDY2. inv OALIVE. + exploit IHib1; eauto. intros (rs2' & s' & ISTEP1 & FSTEP1 & STATES). + rewrite ISTEP1; simpl. + do 2 eexists; intuition eauto. + - (* Bseq continue *) + destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate. + generalize BDY; clear BDY. + inversion_SOME aliveMid; intros OALIVE BDY2. inv OALIVE. + exploit cfgsem2fsem_ibistep_simu_None; eauto. + intros (rs2' & ISTEP1 & REGS'). rewrite ISTEP1; simpl; eauto. + - (* Bcond *) + generalize BDY; clear BDY. + inversion_SOME oaliveSo; inversion_SOME oaliveNot; intros BDY1 BDY2 JOIN. + erewrite <- eqlive_reg_listmem; eauto. rewrite EVAL. + destruct b; eauto. +Qed. Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: iblock_step tid (Genv.globalenv prog) stk1 f sp rs1 m ibf.(entry) t s1 -> @@ -341,18 +405,17 @@ Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t: eqlive_states s1 s2. Proof. intros STEP STACKS EQLIVE LIVE PC. - unfold liveness_ok_function in LIVE. - apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]]. + assert (CHECKER: liveness_ok_function f) by auto. + unfold liveness_ok_function in CHECKER. + apply decomp_liveness_checker in CHECKER; destruct CHECKER as [ibf' [ENTRY LIST]]. eapply PTree.elements_correct in PC as PCIN. eapply list_iblock_checker_correct in LIST as IBC; eauto. unfold iblock_checker in IBC. generalize IBC; clear IBC. - inversion_SOME x; intros BODY _. - destruct STEP as (rs1' & m1' & fin' & ISTEP1 & FSTEP). - exploit cfgsem2fsem_ibistep_simu; eauto. - intros (rs2' & ISTEP2 & REGS). - rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP1. - exploit cfgsem2fsem_finalstep_simu; eauto. - intros (s2 & FSTEP2 & STATES). clear FSTEP. + inversion_SOME alive; intros BODY _. + destruct STEP as (rs1' & m1' & fin' & ISTEP & FSTEP). + exploit cfgsem2fsem_ibistep_simu_Some; eauto. + intros (rs2' & s' & ISTEP' & FSTEP' & REGS). + rewrite <- iblock_istep_run_equiv in ISTEP'. clear ISTEP. unfold iblock_step. repeat eexists; eauto. Qed. @@ -365,18 +428,22 @@ Lemma cfgsem2fsem_step_simu s1 s1' t s2: step tr_inputs (Genv.globalenv prog) s2 t s2' /\ eqlive_states s1' s2'. Proof. - destruct 1 as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; intros STATES; - try (inv STATES; inv LIVE; eexists; split; econstructor; eauto; fail). - - inv STATES; simplify_someHyps. + destruct 1 as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; intros STATES. + - (* iblock *) + inv STATES; simplify_someHyps. intros. exploit cfgsem2fsem_ibstep_simu; eauto. intros (s2 & STEP2 & EQUIV2). eexists; split; eauto. - - inv STATES; inv LIVE. + - (* function internal *) + inv STATES; inv LIVE. apply liveness_checker_entrypoint in H0 as ENTRY. destruct ((fn_code f) ! (fn_entrypoint f)) eqn:EQENTRY; try congruence; eauto. eexists; split; repeat econstructor; eauto. - - inv STATES. + - (* function external *) + inv STATES; inv LIVE; eexists; split; econstructor; eauto. + - (* return *) + inv STATES. inversion STACKS as [|s1 st1 s' s2 STACK STACKS']; subst; clear STACKS. inv STACK. exists (State s2 f sp pc (rs2 # res <- vres) m); split. -- cgit From 49f759ebbb3eb569c456a9dbe6affd165f3fc8b5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 4 Jun 2021 19:08:44 +0200 Subject: starting refinement of symbolic execution --- scheduling/BTL_SEsimuref.v | 109 +++++++++++++++++++++++++++++++++++++++++++-- scheduling/BTL_SEtheory.v | 18 ++++---- 2 files changed, 115 insertions(+), 12 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index a4da4291..6d74ccf9 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -1,13 +1,18 @@ (** Refinement of BTL_SEtheory data-structures in order to introduce (and prove correct) a lower-level specification of the simulation test. - Ceci est un "bac à sable". TODO: A REVOIR COMPLETEMENT. Introduire "fake" hash-consing + Ceci est un "bac à sable". - On introduit une représentation plus concrète pour les types d'état symbolique [sistate] et [sstate]. - Etant donné une spécification intuitive "*_simu" pour tester la simulation sur cette représentation des états symboliques, on essaye déjà de trouver les bonnes notions de raffinement "*_refines" qui permette de prouver les lemmes "*_simu_correct". - Il faudra ensuite vérifier que l'exécution symbolique préserve ces relations de raffinement ! + TODO: A REVOIR COMPLETEMENT. + - essayer de découper chaque relation de raffinement en une fonction "projection/abstraction" (à la "hsval_proj" à renommer en "hsval_abstract") + et une équivalence sur la représentation abstraite. + - introduire "fake" hash-consed values (à renommer en "refined" plutôt que "fake"). + *) Require Import Coqlib Maps Floats. @@ -15,7 +20,7 @@ Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL BTL OptionMonad BTL_SEtheory. -(** * ... *) +(** * Refinement of data-structures and of the specification of the simulation test *) Record sis_ok ctx (sis: sistate): Prop := { OK_PRE: (sis.(si_pre) ctx); @@ -254,6 +259,103 @@ Proof. * exploit IHrst_simu2; eauto. Qed. +(** * Refinement of the symbolic execution *) + +Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core. + +Lemma eval_list_sval_refpreserv ctx args ris sis: + ris_refines ctx ris sis -> + ris_ok ctx ris -> + eval_list_sval ctx (list_sval_inj (map ris args)) = + eval_list_sval ctx (list_sval_inj (map sis args)). +Proof. + intros REF OK. + induction args; simpl; eauto. + intros; erewrite REG_EQ, IHargs; eauto. +Qed. + +Local Hint Resolve eval_list_sval_refpreserv: core. + +Lemma eval_builtin_sval_refpreserv ctx arg ris sis: + ris_refines ctx ris sis -> + ris_ok ctx ris -> + eval_builtin_sval ctx (builtin_arg_map ris arg) = eval_builtin_sval ctx (builtin_arg_map sis arg). +Proof. + intros REF OK; induction arg; simpl; eauto. + + erewrite REG_EQ; eauto. + + erewrite IHarg1, IHarg2; eauto. + + erewrite IHarg1, IHarg2; eauto. +Qed. + +Lemma bargs_refpreserv ctx args ris sis: + ris_refines ctx ris sis -> + ris_ok ctx ris -> + bargs_refines ctx (map (builtin_arg_map ris) args) (map (builtin_arg_map sis) args). +Proof. + unfold bargs_refines. intros REF OK. + induction args; simpl; eauto. + erewrite eval_builtin_sval_refpreserv, IHargs; eauto. +Qed. + +Local Hint Resolve bargs_refpreserv: core. + +Lemma exec_final_refpreserv ctx i ris sis: + ris_refines ctx ris sis -> + ris_ok ctx ris -> + rfv_refines ctx (sexec_final_sfv i ris) (sexec_final_sfv i sis). +Proof. + destruct i; simpl; unfold sum_left_map; try autodestruct; eauto. +Qed. + +Definition ris_init: ristate := {| ris_smem:= Sinit; ris_input_init:=true; ris_ok_sval := nil; ris_sreg := PTree.empty _ |}. + +Lemma ris_init_correct ctx: + ris_refines ctx ris_init sis_init. +Proof. + unfold ris_init, sis_init; econstructor; simpl in *; eauto. + + split; destruct 1; econstructor; simpl in *; eauto. + congruence. + + destruct 1; simpl in *. unfold ris_sreg_get; simpl. + intros; rewrite PTree.gempty; eauto. + + try_simplify_someHyps. +Qed. + +Definition rset_smem rm (ris:ristate): ristate := + {| ris_smem := rm; + ris_input_init := ris.(ris_input_init); + ris_ok_sval := ris.(ris_ok_sval); + ris_sreg:= ris.(ris_sreg) + |}. + +Lemma sis_ok_set_mem ctx sm sis: + sis_ok ctx (set_smem sm sis) + <-> (sis_ok ctx sis /\ eval_smem ctx sm <> None). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. + intuition eauto. +Qed. + +Lemma ris_ok_set_mem ctx rm (ris: ristate): + (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> + ris_ok ctx (rset_smem rm ris) + <-> (ris_ok ctx ris /\ eval_smem ctx rm <> None). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. +Qed. + +Lemma rset_mem_refpreserv ctx rm sm ris sis: + (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> + (forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) -> + ris_refines ctx ris sis -> + (ris_ok ctx ris -> eval_smem ctx rm = eval_smem ctx sm) -> + ris_refines ctx (rset_smem rm ris) (set_smem sm sis). +Proof. + destruct 3; intros. + econstructor; eauto. + + rewrite sis_ok_set_mem, ris_ok_set_mem; intuition congruence. + + rewrite ris_ok_set_mem; intuition eauto. + + rewrite ris_ok_set_mem; intuition eauto. +Qed. (** TODO: could be useful ? Lemma seval_condition_valid_preserv ctx cond args sm b @@ -268,4 +370,5 @@ Proof. eapply cond_valid_pointer_eq; intros. exploit VALID_PRESERV; eauto. Qed. -*) \ No newline at end of file +*) + diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 9ab64d7d..4b4571e4 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -437,7 +437,7 @@ Inductive sem_sfval ctx stk: sfval -> regset -> mem -> trace -> state -> Prop := (* Syntax and Semantics of symbolic internal states *) (* [si_pre] is a precondition on initial context *) -Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg: reg -> sval; si_smem: smem }. +Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg:> reg -> sval; si_smem: smem }. (* Predicate on which (rs, m) is a possible final state after evaluating [st] on ((crs0 ctx), (cm0 ctx)) *) Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop := @@ -446,25 +446,25 @@ Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop := /\ forall (r:reg), eval_sval ctx (sis.(si_sreg) r) = Some (rs#r). (** * Symbolic execution of final step *) -Definition sexec_final_sfv (i: final) (sis: sistate): sfval := +Definition sexec_final_sfv (i: final) (sreg: reg -> sval): sfval := match i with | Bgoto pc => Sgoto pc | Bcall sig ros args res pc => - let svos := sum_left_map sis.(si_sreg) ros in - let sargs := list_sval_inj (List.map sis.(si_sreg) args) in + let svos := sum_left_map sreg ros in + let sargs := list_sval_inj (List.map sreg args) in Scall sig svos sargs res pc | Btailcall sig ros args => - let svos := sum_left_map sis.(si_sreg) ros in - let sargs := list_sval_inj (List.map sis.(si_sreg) args) in + let svos := sum_left_map sreg ros in + let sargs := list_sval_inj (List.map sreg args) in Stailcall sig svos sargs | Bbuiltin ef args res pc => - let sargs := List.map (builtin_arg_map sis.(si_sreg)) args in + let sargs := List.map (builtin_arg_map sreg) args in Sbuiltin ef sargs res pc | Breturn or => - let sor := SOME r <- or IN Some (sis.(si_sreg) r) in + let sor := SOME r <- or IN Some (sreg r) in Sreturn sor | Bjumptable reg tbl => - let sv := sis.(si_sreg) reg in + let sv := sreg reg in Sjumptable sv tbl end. -- cgit From 4527f1503e7e7c76d12cfac10e0e7e719b0578a6 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 4 Jun 2021 19:40:11 +0200 Subject: preparation and starting final lemma --- scheduling/BTL_Livecheck.v | 121 +++++++++++++++++++++++++++++++++++++-------- 1 file changed, 100 insertions(+), 21 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 95b6a155..8d613ef8 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -26,33 +26,33 @@ Definition reg_builtin_res (res: builtin_res reg) (alive: Regset.t): Regset.t := | _ => alive end. -Definition exit_checker {A} (btl: code) (alive: Regset.t) (s: node) (v: A): option A := +Definition exit_checker (btl: code) (alive: Regset.t) (s: node): option unit := SOME next <- btl!s IN ASSERT Regset.subset next.(input_regs) alive IN - Some v. + Some tt. Fixpoint exit_list_checker (btl: code) (alive: Regset.t) (l: list node): bool := match l with | nil => true - | s :: l' => exit_checker btl alive s true &&& exit_list_checker btl alive l' + | s :: l' => exit_checker btl alive s &&& exit_list_checker btl alive l' end. Definition final_inst_checker (btl: code) (alive: Regset.t) (fin: final): option unit := match fin with | Bgoto s => - exit_checker btl alive s tt + exit_checker btl alive s | Breturn oreg => ASSERT reg_option_mem oreg alive IN Some tt | Bcall _ ros args res s => ASSERT list_mem args alive IN ASSERT reg_sum_mem ros alive IN - exit_checker btl (Regset.add res alive) s tt + exit_checker btl (Regset.add res alive) s | Btailcall _ ros args => ASSERT list_mem args alive IN ASSERT reg_sum_mem ros alive IN Some tt | Bbuiltin _ args res s => ASSERT list_mem (params_of_builtin_args args) alive IN - exit_checker btl (reg_builtin_res res alive) s tt + exit_checker btl (reg_builtin_res res alive) s | Bjumptable arg tbl => ASSERT Regset.mem arg alive IN ASSERT exit_list_checker btl alive tbl IN Some tt @@ -273,34 +273,110 @@ Variable prog: BTL.program. Let ge := Genv.globalenv prog. -Hypothesis all_liveness_checker: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. +Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. - -(*Lemma tr_inputs_eqlive f rs1 rs2 tbl or r - (REGS: forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) +(* +Lemma tr_inputs_eqlive f rs1 rs2 tbl or r + forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) :eqlive_reg (tr_inputs f tbl or rs1) (tid f tbl or rs2). +REGS' : forall v : val, + eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v +forall v : val, +eqlive_reg (ext (input_regs ibf)) + (tid f tbl or rs1) # res <- v + (tr_inputs f tbl or rs2) # res <- v + + Proof. unfold tid; rewrite tr_inputs_get. autodestruct. - Qed.*) +Qed. + *) +Lemma find_funct_liveness_ok v fd: + Genv.find_funct ge v = Some fd -> liveness_ok_fundef fd. +Proof. + unfold Genv.find_funct. + destruct v; try congruence. + destruct (Integers.Ptrofs.eq_dec _ _); try congruence. + eapply all_fundef_liveness_ok; eauto. +Qed. + +Lemma find_function_liveness_ok ros rs f: + find_function ge ros rs = Some f -> liveness_ok_fundef f. +Proof. + destruct ros as [r|i]; simpl. + - intros; eapply find_funct_liveness_ok; eauto. + - destruct (Genv.find_symbol ge i); try congruence. + eapply all_fundef_liveness_ok; eauto. +Qed. -Lemma cfgsem2fsem_finalstep_simu stk1 f sp rs1 m fin t s1 stk2 rs2 - (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s1) +Lemma find_function_eqlive alive ros rs1 rs2: + eqlive_reg (ext alive) rs1 rs2 -> + reg_sum_mem ros alive = true -> + find_function ge ros rs1 = find_function ge ros rs2. +Proof. + intros EQLIVE. + destruct ros; simpl; auto. + intros H; erewrite (EQLIVE r); eauto. +Qed. + +Lemma exit_checker_eqlive_ext1 (btl: code) (alive: Regset.t) (pc: node) r rs1 rs2: + exit_checker btl (Regset.add r alive) pc = Some tt -> + eqlive_reg (ext alive) rs1 rs2 -> + exists ibf, btl!pc = Some ibf /\ (forall v, eqlive_reg (ext ibf.(input_regs)) (rs1 # r <- v) (rs2 # r <- v)). +Proof. + unfold exit_checker. + inversion_SOME next. + inversion_ASSERT. try_simplify_someHyps. + repeat (econstructor; eauto). + intros; eapply eqlive_reg_update; eauto. + eapply eqlive_reg_monotonic; eauto. + intros r0 [X1 X2]; exploit Regset.subset_2; eauto. + rewrite regset_add_spec. intuition subst. +Qed. + +Lemma cfgsem2fsem_finalstep_simu sp f stk1 stk2 s t fin alive rs1 m rs2 + (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s) (LIVE: liveness_ok_function f) + (REGS: eqlive_reg (ext alive) rs1 rs2) + (FCHK: final_inst_checker (fn_code f) alive fin = Some tt) (STACKS: list_forall2 eqlive_stackframes stk1 stk2) - (*body_checker (fn_code f) (entry ibf) (input_regs ibf) = Some x ->*) - (*eqlive_reg (ext x) rs1' rs2' ->*) - :exists s2, - final_step tr_inputs ge stk2 f sp rs2 m fin t s2 - /\ eqlive_states s1 s2. + :exists s', + final_step tr_inputs ge stk2 f sp rs2 m fin t s' + /\ eqlive_states s s'. Proof. - destruct FSTEP. + destruct FSTEP; try_simplify_someHyps; repeat inversion_ASSERT; intros. - (* Bgoto *) econstructor. econstructor. econstructor. econstructor. eauto. eauto. repeat econstructor; eauto. + all: admit. + - (* Breturn *) + eexists; split. econstructor; eauto. + destruct or; simpl in *; + try erewrite (REGS r); eauto. + - (* Bcall *) + exploit exit_checker_eqlive_ext1; eauto. + intros (ibf & PC & REGS'). + eexists; split. + + econstructor; eauto. + erewrite <- find_function_eqlive; eauto. + + erewrite eqlive_reg_listmem; eauto. + eapply eqlive_states_call; eauto. + eapply find_function_liveness_ok; eauto. + repeat (econstructor; eauto). + admit. + - (* Btailcall *) + eexists; split. + + econstructor; eauto. + erewrite <- find_function_eqlive; eauto. + + erewrite eqlive_reg_listmem; eauto. + eapply eqlive_states_call; eauto. + eapply find_function_liveness_ok; eauto. + - (* Bbuiltin *) + Admitted. Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m' @@ -363,6 +439,7 @@ Lemma cfgsem2fsem_ibistep_simu_Some sp f stk1 stk2 ib: forall s t rs1 m rs1' m' alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2) (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2)) (LIVE: liveness_ok_function f) + (*(PC : (fn_code f) ! pc = Some ibf)*) (STACKS: list_forall2 eqlive_stackframes stk1 stk2), exists rs2' s', iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' (Some fin)) @@ -372,8 +449,10 @@ Proof. induction ib; simpl; try_simplify_someHyps; repeat inversion_ASSERT; intros; inv ISTEP. - (* BF *) - exploit cfgsem2fsem_finalstep_simu; eauto. - intros (s2 & FSTEP' & STATES). eauto. + generalize BDY; clear BDY. + inversion_SOME x; try_simplify_someHyps; intros FCHK. + destruct x; exploit cfgsem2fsem_finalstep_simu; eauto. + intros (s2 & FSTEP' & STATES); eauto. - (* Bseq stop *) destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate. generalize BDY; clear BDY. -- cgit From 322192e944a099c004dc8e71df0b1da975ec680c Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 6 Jun 2021 09:34:36 +0200 Subject: IDEE pour la STRENGTH-REDUCTION --- scheduling/BTLroadmap.md | 44 +++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 43 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 954143c1..bfbc1e03 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -346,13 +346,55 @@ Au final, on obtient deux BDD identiques sur cet exemple (sans explosion combina Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. Case-study: support of strength-reduction. -**IDEAS** +**PRELIMINARY IDEAS** - En pratique, il faudrait affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. Une solution possible: calquer Bjumptable sur Bcond (c-a-d autoriser les Bjumptable en milieu de blocs). On pourrait étendre la prédiction de branchement par profiling aux jumptables (par exemple, avoir des superblocks avec des jumptables au milieu). Un autre intérêt: simplifier (un peu) le support pour des formes SSA partielles (cf. ci-dessous). Ceci dit, ça compliquerait les choses à plein d'endroits (type Coq [iblock] mutuellement inductif avec [list_iblock] pour avoir les bons principes d'inductions, etc) pour des gains minimes en pratiques ? - Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). +**EXAMPLE OF STRENGTH REDUCTION** + +On veut passer du code C1: + + init: // inputs: int *t, int n, int s + int i=0; + loop: // inputs: int *t, int n, int s, int i + if (i >= n) goto exit; + s += t[i]; + i += 1; + goto loop; + exit: // inputs: int *t, int s + +au code C2: + + init: // inputs: int *t, int n, int s + int *ti = t; + int *tn = t+n; + loop: // inputs: int *t, int s, int *ti, int *tn + if (ti >= tn) goto exit; + s += *ti; + ti += 4; + goto loop; + exit; // inputs: int *t, int s + +Pour donner la correspondance entre les variables des 2 blocs, pour chaque entrée "pc", on introduit une "fonction" de C1@pc.(inputs) -> C2@pc.(inputs). + +Typiquement, cette fonction est codable comme un map associant une valeur symbolique sur les C1@pc.(inputs) aux variables de C2@pc.(inputs). Exemple: + + init: // map vide (identité) + loop: + ti = t+i + tn = t+n + exit: // map vide (identité) + +Si on note `TRANSFER` cette fonction, alors la vérification par +exécution symbolique que le bloc `ib1` est simulé par `ib2` modulo +`TRANSFER` sur l'entrée `pc` se ramène à montrer `ib1[TRANSFER]` avec +`TRANSFER(pc); ib2` pour la simulation symbolique usuelle. + + Ci-dessus `ib1[TRANSFER]` dit qu'on a appliqué `TRANSFER(pc')` sur chacune des sorties `pc'` de `ib1`. + ## Support of SSA-optimizations Minimum feature: extends BTL with "register renamings" at exits. This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. -- cgit From f855402a10148047e449073ae6c32d269275cdf4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 6 Jun 2021 20:32:28 +0200 Subject: m --- scheduling/BTLroadmap.md | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index bfbc1e03..9dd21be9 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -393,7 +393,9 @@ exécution symbolique que le bloc `ib1` est simulé par `ib2` modulo `TRANSFER` sur l'entrée `pc` se ramène à montrer `ib1[TRANSFER]` avec `TRANSFER(pc); ib2` pour la simulation symbolique usuelle. - Ci-dessus `ib1[TRANSFER]` dit qu'on a appliqué `TRANSFER(pc')` sur chacune des sorties `pc'` de `ib1`. +Ci-dessus `ib1[TRANSFER]` dit qu'on a appliqué `TRANSFER(pc')` sur chacune des sorties `pc'` de `ib1`. + +**REM** pour que ce soit correct, il faut sans doute vérifier une condition du style `ok(ib1) => ok(ib1[TRANSFER])`... ## Support of SSA-optimizations -- cgit From 9e3940d183cb08cb9ed4b257f6dbe7a5aa05e9b6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 7 Jun 2021 09:32:29 +0200 Subject: avancement ref de l'exe symbolique --- scheduling/BTL_SEsimuref.v | 321 ++++++++++++++++++++++++++++++++++++++++++--- scheduling/BTL_SEtheory.v | 11 +- 2 files changed, 306 insertions(+), 26 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 6d74ccf9..fdefe205 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -9,8 +9,6 @@ - Il faudra ensuite vérifier que l'exécution symbolique préserve ces relations de raffinement ! TODO: A REVOIR COMPLETEMENT. - - essayer de découper chaque relation de raffinement en une fonction "projection/abstraction" (à la "hsval_proj" à renommer en "hsval_abstract") - et une équivalence sur la représentation abstraite. - introduire "fake" hash-consed values (à renommer en "refined" plutôt que "fake"). *) @@ -48,7 +46,7 @@ Record ristate := 1) we store a list of sval evaluations 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval *) ris_input_init: bool; - ris_ok_sval: list sval; + ok_rsval: list sval; ris_sreg:> PTree.t sval }. @@ -61,7 +59,7 @@ Coercion ris_sreg_get: ristate >-> Funclass. Record ris_ok ctx (ris: ristate) : Prop := { OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; - OK_RREG: forall sv, List.In sv (ris_ok_sval ris) -> eval_sval ctx sv <> None + OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None }. Local Hint Resolve OK_RMEM OK_RREG: core. Local Hint Constructors ris_ok: core. @@ -77,7 +75,7 @@ Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ VALID_PRESERV: core. Local Hint Constructors ris_refines: core. Record ris_simu ris1 ris2: Prop := { - SIMU_FAILS: forall sv, List.In sv ris2.(ris_ok_sval) -> List.In sv ris1.(ris_ok_sval); + SIMU_FAILS: forall sv, List.In sv ris2.(ok_rsval) -> List.In sv ris1.(ok_rsval); SIMU_MEM: ris1.(ris_smem) = ris2.(ris_smem); SIMU_REG: forall r, ris_sreg_get ris1 r = ris_sreg_get ris2 r }. @@ -307,7 +305,7 @@ Proof. destruct i; simpl; unfold sum_left_map; try autodestruct; eauto. Qed. -Definition ris_init: ristate := {| ris_smem:= Sinit; ris_input_init:=true; ris_ok_sval := nil; ris_sreg := PTree.empty _ |}. +Definition ris_init: ristate := {| ris_smem:= Sinit; ris_input_init:=true; ok_rsval := nil; ris_sreg := PTree.empty _ |}. Lemma ris_init_correct ctx: ris_refines ctx ris_init sis_init. @@ -323,11 +321,11 @@ Qed. Definition rset_smem rm (ris:ristate): ristate := {| ris_smem := rm; ris_input_init := ris.(ris_input_init); - ris_ok_sval := ris.(ris_ok_sval); + ok_rsval := ris.(ok_rsval); ris_sreg:= ris.(ris_sreg) |}. -Lemma sis_ok_set_mem ctx sm sis: +Lemma ok_set_mem ctx sm sis: sis_ok ctx (set_smem sm sis) <-> (sis_ok ctx sis /\ eval_smem ctx sm <> None). Proof. @@ -335,7 +333,7 @@ Proof. intuition eauto. Qed. -Lemma ris_ok_set_mem ctx rm (ris: ristate): +Lemma ok_rset_mem ctx rm (ris: ristate): (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> ris_ok ctx (rset_smem rm ris) <-> (ris_ok ctx ris /\ eval_smem ctx rm <> None). @@ -343,7 +341,7 @@ Proof. split; destruct 1; econstructor; simpl in *; eauto. Qed. -Lemma rset_mem_refpreserv ctx rm sm ris sis: +Lemma rset_mem_correct ctx rm sm ris sis: (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> (forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) -> ris_refines ctx ris sis -> @@ -352,23 +350,308 @@ Lemma rset_mem_refpreserv ctx rm sm ris sis: Proof. destruct 3; intros. econstructor; eauto. - + rewrite sis_ok_set_mem, ris_ok_set_mem; intuition congruence. - + rewrite ris_ok_set_mem; intuition eauto. - + rewrite ris_ok_set_mem; intuition eauto. + + rewrite ok_set_mem, ok_rset_mem; intuition congruence. + + rewrite ok_rset_mem; intuition eauto. + + rewrite ok_rset_mem; intuition eauto. Qed. -(** TODO: could be useful ? -Lemma seval_condition_valid_preserv ctx cond args sm b +Definition rexec_store chunk addr args src ris: ristate := + let args := list_sval_inj (List.map (ris_sreg_get ris) args) in + let src := ris_sreg_get ris src in + let rm := Sstore ris.(ris_smem) chunk addr args src in + rset_smem rm ris. + +Lemma rexec_store_correct ctx chunk addr args src ris sis: + ris_refines ctx ris sis -> + ris_refines ctx (rexec_store chunk addr args src ris) (sexec_store chunk addr args src sis). +Proof. + intros REF; eapply rset_mem_correct; simpl; eauto. + + intros X; rewrite X. repeat autodestruct; eauto. + + intros m b ofs; repeat autodestruct. + intros; erewrite <- Mem.storev_preserv_valid; [| eauto]. + eauto. + + intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ, REG_EQ; eauto. +Qed. + +(* TODO: reintroduire le "root_apply" ? *) + +Definition rset_sreg r rsv (ris:ristate): ristate := + {| ris_smem := ris.(ris_smem); + ris_input_init := ris.(ris_input_init); + ok_rsval := rsv::ris.(ok_rsval); (* TODO: A CHANGER ? *) + ris_sreg:= PTree.set r rsv ris.(ris_sreg) (* TODO: A CHANGER *) + |}. + +Lemma ok_set_sreg ctx r sv sis: + sis_ok ctx (set_sreg r sv sis) + <-> (sis_ok ctx sis /\ eval_sval ctx sv <> None). +Proof. + unfold set_sreg; split. + + intros [(SVAL & PRE) SMEM SREG]; simpl in *; split. + - econstructor; eauto. + intros r0; generalize (SREG r0); destruct (Pos.eq_dec r r0); try congruence. + - generalize (SREG r); destruct (Pos.eq_dec r r); try congruence. + + intros ([PRE SMEM SREG] & SVAL). + econstructor; simpl; eauto. + intros r0; destruct (Pos.eq_dec r r0); try congruence. +Qed. + +Lemma ok_rset_sreg ctx r rsv ris: + ris_ok ctx (rset_sreg r rsv ris) + <-> (ris_ok ctx ris /\ eval_sval ctx rsv <> None). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. + intuition subst; eauto. + exploit OK_RREG; eauto. +Qed. + +Lemma rset_reg_correct ctx r rsv sv ris sis: + ris_refines ctx ris sis -> + (ris_ok ctx ris -> eval_sval ctx rsv = eval_sval ctx sv) -> + ris_refines ctx (rset_sreg r rsv ris) (set_sreg r sv sis). +Proof. + destruct 1; intros. + econstructor; eauto. + + rewrite ok_set_sreg, ok_rset_sreg; intuition congruence. + + rewrite ok_rset_sreg; intuition eauto. + + rewrite ok_rset_sreg. intros; unfold rset_sreg, set_sreg, ris_sreg_get; simpl. intuition eauto. + destruct (Pos.eq_dec _ _). + * subst; rewrite PTree.gss; eauto. + * rewrite PTree.gso; eauto. +Qed. + +Definition rexec_op op args dst (ris:ristate): ristate := + let args := list_sval_inj (List.map ris args) in + rset_sreg dst (Sop op args Sinit) ris. + +Lemma rexec_op_correct ctx op args dst ris sis: + ris_refines ctx ris sis -> + ris_refines ctx (rexec_op op args dst ris) (sexec_op op args dst sis). +Proof. + intros REF; eapply rset_reg_correct; simpl; eauto. + intros OK; erewrite eval_list_sval_refpreserv; eauto. + do 2 autodestruct; auto. + + intros. erewrite <- op_valid_pointer_eq; eauto. + + erewrite <- MEM_EQ; eauto. + intros; exploit OK_RMEM; eauto. destruct 1. +Qed. + +Definition rexec_load trap chunk addr args dst (ris:ristate): ristate := + let args := list_sval_inj (List.map ris args) in + rset_sreg dst (Sload ris.(ris_smem) trap chunk addr args) ris. + +Lemma rexec_load_correct ctx trap chunk addr args dst ris sis: + ris_refines ctx ris sis -> + ris_refines ctx (rexec_load trap chunk addr args dst ris) (sexec_load trap chunk addr args dst sis). +Proof. + intros REF; eapply rset_reg_correct; simpl; eauto. + intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ; eauto. +Qed. + +Lemma seval_condition_valid_preserv ctx cond args sm + (OK: eval_smem ctx sm <> None) (VALID_PRESERV: forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) - :seval_condition ctx cond args sm = Some b -> - seval_condition ctx cond args Sinit = Some b. + :seval_condition ctx cond args sm = seval_condition ctx cond args Sinit. Proof. unfold seval_condition; autodestruct; simpl; try congruence. intros EVAL_LIST. autodestruct; try congruence. - intros MEM COND. rewrite <- COND. + intros. eapply cond_valid_pointer_eq; intros. exploit VALID_PRESERV; eauto. Qed. -*) + +Lemma seval_condition_refpreserv ctx cond args ris sis: + ris_refines ctx ris sis -> + ris_ok ctx ris -> + seval_condition ctx cond (list_sval_inj (map ris args)) Sinit = + seval_condition ctx cond (list_sval_inj (map sis args)) Sinit. +Proof. + intros; unfold seval_condition. + erewrite eval_list_sval_refpreserv; eauto. +Qed. + + +(* transfer *) + +Definition rseto_sreg r rsv (ris:ristate): ristate := + {| ris_smem := ris.(ris_smem); + ris_input_init := ris.(ris_input_init); + ok_rsval := ris.(ok_rsval); + ris_sreg:= PTree.set r rsv ris.(ris_sreg) (* TODO: A CHANGER *) + |}. + +Lemma ok_rseto_sreg ctx r rsv ris: + ris_ok ctx (rseto_sreg r rsv ris) + <-> (ris_ok ctx ris). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. +Qed. + +Lemma rseto_reg_correct ctx r rsv sv ris sis: + ris_refines ctx ris sis -> + (ris_ok ctx ris -> eval_sval ctx rsv <> None) -> + (ris_ok ctx ris -> eval_sval ctx rsv = eval_sval ctx sv) -> + ris_refines ctx (rseto_sreg r rsv ris) (set_sreg r sv sis). +Proof. + destruct 1; intros. + econstructor; eauto. + + rewrite ok_set_sreg, ok_rseto_sreg; intuition congruence. + + rewrite ok_rseto_sreg; intuition eauto. + + rewrite ok_rseto_sreg. intros; unfold rseto_sreg, set_sreg, ris_sreg_get; simpl. intuition eauto. + destruct (Pos.eq_dec _ _). + * subst; rewrite PTree.gss; eauto. + * rewrite PTree.gso; eauto. +Qed. + +Fixpoint transfer_ris (inputs: list reg) (ris:ristate): ristate := + match inputs with + | nil => {| ris_smem := ris.(ris_smem); + ris_input_init := false; + ok_rsval := ris.(ok_rsval); + ris_sreg:= PTree.empty _ + |} + | r::l => rseto_sreg r (ris_sreg_get ris r) (transfer_ris l ris) + end. + +Definition transfer_sis (inputs: list reg) (sis:sistate): sistate := + {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); + si_sreg := transfer_sreg inputs sis; + si_smem := sis.(si_smem) |}. + +Lemma ok_transfer_sis ctx inputs sis: + sis_ok ctx (transfer_sis inputs sis) + <-> (sis_ok ctx sis). +Proof. + unfold transfer_sis. induction inputs as [|r l]; simpl. + + split; destruct 1; econstructor; simpl in *; intuition eauto. congruence. + + split. + * destruct 1; econstructor; simpl in *; intuition eauto. + * intros X; generalize X. rewrite <- IHl in X; clear IHl. + intros [PRE SMEM SREG]. + econstructor; simpl; eauto. + intros r0; destruct (Pos.eq_dec r r0); try congruence. + intros H; eapply OK_SREG; eauto. +Qed. + +Lemma ok_transfer_ris ctx inputs ris: + ris_ok ctx (transfer_ris inputs ris) + <-> (ris_ok ctx ris). +Proof. + induction inputs as [|r l]; simpl. + + split; destruct 1; econstructor; simpl in *; intuition eauto. + + rewrite ok_rseto_sreg. auto. +Qed. + +Lemma transfer_ris_correct ctx inputs ris sis: + ris_refines ctx ris sis -> + ris_refines ctx (transfer_ris inputs ris) (transfer_sis inputs sis). +Proof. + destruct 1; intros. + induction inputs as [|r l]. + + econstructor; eauto. + * erewrite ok_transfer_sis, ok_transfer_ris; eauto. + * erewrite ok_transfer_ris; eauto. + * erewrite ok_transfer_ris; simpl; unfold ris_sreg_get; simpl; eauto. + intros; rewrite PTree.gempty. simpl; auto. + + econstructor; eauto. + * erewrite ok_transfer_sis, ok_transfer_ris; eauto. + * erewrite ok_transfer_ris; simpl. + intros; erewrite MEM_EQ. 2: eauto. + - unfold transfer_sis; simpl; eauto. + - rewrite ok_transfer_ris; simpl; eauto. + * erewrite ok_transfer_ris; simpl. + intros H r0. + erewrite REG_EQ. 2: eapply rseto_reg_correct; eauto. + - unfold set_sreg; simpl; auto. + destruct (Pos.eq_dec _ _); simpl; auto. + - intros. rewrite REG_EQ0; auto. apply OK_SREG; tauto. + - rewrite ok_rseto_sreg, ok_transfer_ris. auto. +Qed. + +Definition alt_tr_sis := poly_tr (fun f tbl or => transfer_sis (Regset.elements (pre_inputs f tbl or))). + +Lemma tr_sis_alt_def f fi sis: + alt_tr_sis f fi sis = tr_sis f fi sis. +Proof. + unfold tr_sis, str_inputs. destruct fi; simpl; auto. +Qed. + +Definition tr_ris := poly_tr (fun f tbl or => transfer_ris (Regset.elements (pre_inputs f tbl or))). + +Local Hint Resolve transfer_ris_correct ok_transfer_ris: core. +Local Opaque transfer_ris. + +Lemma ok_tr_ris ctx f fi ris: + ris_ok ctx (tr_ris f fi ris) + <-> (ris_ok ctx ris). +Proof. + destruct fi; simpl; eauto. +Qed. + +Lemma ok_tr_ris_imp ctx f fi ris: + ris_ok ctx (tr_ris f fi ris) + -> (ris_ok ctx ris). +Proof. + rewrite ok_tr_ris; auto. +Qed. + + +Lemma tr_ris_correct ctx f fi ris sis: + ris_refines ctx ris sis -> + ris_refines ctx (tr_ris f fi ris) (tr_sis f fi sis). +Proof. + intros REF. rewrite <- tr_sis_alt_def. + destruct fi; simpl; eauto. +Qed. + +Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := + match ib with + | BF fin _ => Rfinal (tr_ris f fin ris) (sexec_final_sfv fin ris) + (* basic instructions *) + | Bnop _ => k ris + | Bop op args res _ => k (rexec_op op args res ris) + | Bload TRAP chunk addr args dst _ => k (rexec_load TRAP chunk addr args dst ris) + | Bload NOTRAP chunk addr args dst _ => Rabort + | Bstore chunk addr args src _ => k (rexec_store chunk addr args src ris) + (* composed instructions *) + | Bseq ib1 ib2 => + rexec_rec f ib1 ris (fun ris2 => rexec_rec f ib2 ris2 k) + | Bcond cond args ifso ifnot _ => + let args := list_sval_inj (List.map ris args) in + let ifso := rexec_rec f ifso ris k in + let ifnot := rexec_rec f ifnot ris k in + Rcond cond args ifso ifnot + end + . + +Local Hint Constructors rst_refines: core. +Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ok_tr_ris_imp + rexec_op_correct rexec_load_correct rexec_store_correct: core. + +Lemma rexec_rec_correct ctx f ib: + forall ris sis rk k + (CONT: forall rs ss, ris_refines ctx rs ss -> rst_refines ctx (rk rs) (k ss)) + (REF: ris_refines ctx ris sis) + (*OK: ris_ok ctx ris*) + , rst_refines ctx (rexec_rec f ib ris rk) (sexec_rec f ib sis k). +Proof. + induction ib; simpl; eauto. + - (* load *) autodestruct; eauto. + - intros. + econstructor; eauto. + symmetry. + erewrite seval_condition_valid_preserv; eauto. +Admitted. + +Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). + +Lemma rexec_correct ctx f ib: + rst_refines ctx (rexec f ib) (sexec f ib). +Proof. + eapply rexec_rec_correct; eauto. + (* econstructor; simpl; auto. congruence. *) +Qed. + + diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 4b4571e4..464bb0f0 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -630,7 +630,7 @@ Qed. Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := match inputs with | nil => fun r => Sundef - | r1::l => fun r => if Pos.eq_dec r r1 then sreg r1 else transfer_sreg l sreg r + | r1::l => fun r => if Pos.eq_dec r1 r then sreg r1 else transfer_sreg l sreg r end. Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (pre_inputs f tbl or)). @@ -655,15 +655,12 @@ Definition tr_sis f (fi: final) (sis: sistate) := si_sreg := poly_tr str_inputs f fi sis.(si_sreg); si_smem := sis.(si_smem) |}. -Definition sreg_eval ctx (sis: sistate) (r: reg): option val := - eval_sval ctx (sis.(si_sreg) r). - Lemma tr_sis_regs_correct_aux ctx fin sis rs m: sem_sistate ctx sis rs m -> - (forall r, sreg_eval ctx (tr_sis (cf ctx) fin sis) r = Some (tr_regs (cf ctx) fin rs) # r). + (forall r, eval_sval ctx (tr_sis (cf ctx) fin sis r) = Some (tr_regs (cf ctx) fin rs) # r). Proof. Local Opaque str_inputs. - unfold sreg_eval; simpl. destruct 1 as (_ & _ & REG). + simpl. destruct 1 as (_ & _ & REG). destruct fin; simpl; eauto. Qed. @@ -842,7 +839,7 @@ Lemma sem_sistate_tr_sis_determ ctx sis rs1 m1 fi rs2 m2: Proof. intros H1 H2. lapply (tr_sis_regs_correct_aux ctx fi sis rs1 m1); eauto. - unfold sreg_eval; intros X. + intros X. destruct H1 as (_&MEM1®1). destruct H2 as (_&MEM2®2); simpl in *. intuition try congruence. -- cgit From c918205205fd0883850308c9598ab2e221bdff7b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 7 Jun 2021 10:48:54 +0200 Subject: advance for tr_input proof --- scheduling/BTL_Livecheck.v | 195 +++++++++++++++++++++++++++++++++++++++------ 1 file changed, 171 insertions(+), 24 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 8d613ef8..3882639c 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -276,23 +276,50 @@ Let ge := Genv.globalenv prog. Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f. Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. -(* -Lemma tr_inputs_eqlive f rs1 rs2 tbl or r - forall r, eqlive_reg (ext (input_regs ibf)) rs1 rs2) - :eqlive_reg (tr_inputs f tbl or rs1) (tid f tbl or rs2). -REGS' : forall v : val, - eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v -forall v : val, -eqlive_reg (ext (input_regs ibf)) - (tid f tbl or rs1) # res <- v - (tr_inputs f tbl or rs2) # res <- v - +Lemma tr_inputs_eqlive_None f pc ibf rs1 rs2: + (fn_code f) ! pc = Some ibf -> + eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> + eqlive_reg (ext (input_regs ibf)) (tid f (pc :: nil) None rs1) + (tr_inputs f (pc :: nil) None rs2). +Proof. + intros PC REGS. + unfold eqlive_reg. intros r INR. + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC. + exploit Regset.union_3. eapply INR. + intros INRU. eapply Regset.mem_1 in INRU. + erewrite INRU; eauto. +Qed. + +Lemma eqlive_reg_update_gso alive rs1 rs2 res r: forall v : val, + eqlive_reg (ext alive) rs1 # res <- v rs2 # res <- v -> + res <> r -> Regset.In r alive -> + rs1 # r = rs2 # r. +Proof. + intros v REGS NRES INR. unfold eqlive_reg in REGS. + specialize REGS with r. apply REGS in INR. + rewrite !Regmap.gso in INR; auto. +Qed. + +Lemma tr_inputs_eqlive_update f pc ibf rs1 rs2 res: + (fn_code f) ! pc = Some ibf -> + forall v : val, + eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v -> + eqlive_reg (ext (input_regs ibf)) + (tid f (pc :: nil) (Some res) rs1) # res <- v + (tr_inputs f (pc :: nil) (Some res) rs2) # res <- v. Proof. - unfold tid; rewrite tr_inputs_get. - autodestruct. + intros PC v REGS. apply eqlive_reg_update. + unfold eqlive_reg. intros r (NRES & INR). + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC. assert (NRES': res <> r) by auto. + clear NRES. exploit Regset.union_3. eapply INR. + intros INRU. exploit Regset.remove_2; eauto. + intros INRU_RES. eapply Regset.mem_1 in INRU_RES. + erewrite INRU_RES. eapply eqlive_reg_update_gso; eauto. Qed. - *) + Lemma find_funct_liveness_ok v fd: Genv.find_funct ge v = Some fd -> liveness_ok_fundef fd. Proof. @@ -321,7 +348,34 @@ Proof. intros H; erewrite (EQLIVE r); eauto. Qed. -Lemma exit_checker_eqlive_ext1 (btl: code) (alive: Regset.t) (pc: node) r rs1 rs2: +Lemma exit_checker_eqlive (btl: code) (alive: Regset.t) (pc: node) rs1 rs2: + exit_checker btl alive pc = Some tt -> + eqlive_reg (ext alive) rs1 rs2 -> + exists ibf, btl!pc = Some ibf /\ eqlive_reg (ext ibf.(input_regs)) rs1 rs2. +Proof. + unfold exit_checker. + inversion_SOME next. + inversion_ASSERT. try_simplify_someHyps. + repeat (econstructor; eauto). + intros; eapply eqlive_reg_monotonic; eauto. + intros; exploit Regset.subset_2; eauto. +Qed. + +Lemma exit_list_checker_eqlive (btl: code) (alive: Regset.t) (tbl: list node) rs1 rs2 pc: forall n, + exit_list_checker btl alive tbl = true -> + eqlive_reg (ext alive) rs1 rs2 -> + list_nth_z tbl n = Some pc -> + exists ibf, btl!pc = Some ibf /\ eqlive_reg (ext ibf.(input_regs)) rs1 rs2. +Proof. + induction tbl; simpl. + - intros; try congruence. + - intros n; rewrite lazy_and_Some_tt_true; destruct (zeq n 0) eqn: Hn. + * try_simplify_someHyps; intuition. + exploit exit_checker_eqlive; eauto. + * intuition. eapply IHtbl; eauto. +Qed. + +Lemma exit_checker_eqlive_update (btl: code) (alive: Regset.t) (pc: node) r rs1 rs2: exit_checker btl (Regset.add r alive) pc = Some tt -> eqlive_reg (ext alive) rs1 rs2 -> exists ibf, btl!pc = Some ibf /\ (forall v, eqlive_reg (ext ibf.(input_regs)) (rs1 # r <- v) (rs2 # r <- v)). @@ -336,6 +390,56 @@ Proof. rewrite regset_add_spec. intuition subst. Qed. +Lemma exit_checker_eqlive_builtin_res (btl: code) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg): + exit_checker btl (reg_builtin_res res alive) pc = Some tt -> + eqlive_reg (ext alive) rs1 rs2 -> + exists ibf, btl!pc = Some ibf /\ (forall vres, eqlive_reg (ext ibf.(input_regs)) (regmap_setres res vres rs1) (regmap_setres res vres rs2)). +Proof. + destruct res; simpl. + - intros; exploit exit_checker_eqlive_update; eauto. + - intros; exploit exit_checker_eqlive; eauto. + intros (ibf & PC & REGS). + eexists; intuition eauto. + - intros; exploit exit_checker_eqlive; eauto. + intros (ibf & PC & REGS). + eexists; intuition eauto. +Qed. + +Local Hint Resolve in_or_app: local. +Lemma eqlive_eval_builtin_args alive rs1 rs2 sp m args vargs: + eqlive_reg alive rs1 rs2 -> + Events.eval_builtin_args ge (fun r => rs1 # r) sp m args vargs -> + (forall r, List.In r (params_of_builtin_args args) -> alive r) -> + Events.eval_builtin_args ge (fun r => rs2 # r) sp m args vargs. +Proof. + unfold Events.eval_builtin_args. + intros EQLIVE; induction 1 as [|a1 al b1 bl EVAL1 EVALL]; simpl. + { econstructor; eauto. } + intro X. + assert (X1: eqlive_reg (fun r => In r (params_of_builtin_arg a1)) rs1 rs2). + { eapply eqlive_reg_monotonic; eauto with local. } + lapply IHEVALL; eauto with local. + clear X IHEVALL; intro X. econstructor; eauto. + generalize X1; clear EVALL X1 X. + induction EVAL1; simpl; try (econstructor; eauto; fail). + - intros X1; erewrite X1; [ econstructor; eauto | eauto ]. + - intros; econstructor. + + eapply IHEVAL1_1; eauto. + eapply eqlive_reg_monotonic; eauto. + simpl; intros; eauto with local. + + eapply IHEVAL1_2; eauto. + eapply eqlive_reg_monotonic; eauto. + simpl; intros; eauto with local. + - intros; econstructor. + + eapply IHEVAL1_1; eauto. + eapply eqlive_reg_monotonic; eauto. + simpl; intros; eauto with local. + + eapply IHEVAL1_2; eauto. + eapply eqlive_reg_monotonic; eauto. + simpl; intros; eauto with local. +Qed. + +Local Hint Resolve tr_inputs_eqlive_None tr_inputs_eqlive_update: core. Lemma cfgsem2fsem_finalstep_simu sp f stk1 stk2 s t fin alive rs1 m rs2 (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s) (LIVE: liveness_ok_function f) @@ -348,17 +452,17 @@ Lemma cfgsem2fsem_finalstep_simu sp f stk1 stk2 s t fin alive rs1 m rs2 Proof. destruct FSTEP; try_simplify_someHyps; repeat inversion_ASSERT; intros. - (* Bgoto *) - econstructor. econstructor. - econstructor. econstructor. - eauto. eauto. - repeat econstructor; eauto. - all: admit. + eexists; split. + + econstructor; eauto. + + exploit exit_checker_eqlive; eauto. + intros (ibf & PC & REGS'). + econstructor; eauto. - (* Breturn *) eexists; split. econstructor; eauto. destruct or; simpl in *; try erewrite (REGS r); eauto. - (* Bcall *) - exploit exit_checker_eqlive_ext1; eauto. + exploit exit_checker_eqlive_update; eauto. intros (ibf & PC & REGS'). eexists; split. + econstructor; eauto. @@ -366,8 +470,6 @@ Proof. + erewrite eqlive_reg_listmem; eauto. eapply eqlive_states_call; eauto. eapply find_function_liveness_ok; eauto. - repeat (econstructor; eauto). - admit. - (* Btailcall *) eexists; split. + econstructor; eauto. @@ -376,7 +478,52 @@ Proof. eapply eqlive_states_call; eauto. eapply find_function_liveness_ok; eauto. - (* Bbuiltin *) - + exploit exit_checker_eqlive_builtin_res; eauto. + intros (ibf & PC & REGS'). + eexists; split. + + econstructor; eauto. + eapply eqlive_eval_builtin_args; eauto. + intros; eapply list_mem_correct; eauto. + + repeat (econstructor; simpl; eauto). + unfold regmap_setres. destruct res; simpl in *; eauto. + - (* Bjumptable *) + exploit exit_list_checker_eqlive; eauto. + intros (ibf & PC & REGS'). + eexists; split. + + econstructor; eauto. + erewrite <- REGS; eauto. + + repeat (econstructor; simpl; eauto). + +(* + generalize dependent ibf. + generalize dependent n. + generalize dependent alive. + generalize dependent tbl. + induction tbl; simpl; try congruence. + intros alive REGS. rewrite lazy_and_Some_tt_true. + intros (EXIT & EXIT_REM) INARG n ARG. autodestruct. + * try_simplify_someHyps; intros. + + unfold eqlive_reg. intros r INR. + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC. + exploit Regset.union_3. eapply INR. + intros INRU. eapply Regset.mem_1 in INRU. + erewrite INRU; eauto. + + + * try_simplify_someHyps; intros. + exploit IHtbl. eapply REGS. + eauto. eauto. eauto. eauto. + exploit exit_checker_eqlive; eauto. + intros (ibf' & PC' & REGS''). + unfold eqlive_reg. intros r INR. + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC'. + exploit Regset.union_3. eapply INR. + intros INRU. eapply Regset.mem_1 in INRU. + erewrite INRU; eauto. + admit.*) Admitted. Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m' -- cgit From 644f1a3208f9a4e013928e042257763313ac2b0d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 7 Jun 2021 12:17:53 +0200 Subject: ugly complete version of liveness proof (I will clean in next commits) --- scheduling/BTL_Livecheck.v | 74 +++++++++++++++++++++++++--------------------- 1 file changed, 40 insertions(+), 34 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 3882639c..a92f24c3 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -277,11 +277,11 @@ Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr ge b = Some f Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. -Lemma tr_inputs_eqlive_None f pc ibf rs1 rs2: +Lemma tr_inputs_eqlive_None f pc tbl ibf rs1 rs2: (fn_code f) ! pc = Some ibf -> eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> - eqlive_reg (ext (input_regs ibf)) (tid f (pc :: nil) None rs1) - (tr_inputs f (pc :: nil) None rs2). + eqlive_reg (ext (input_regs ibf)) (tid f (pc :: tbl) None rs1) + (tr_inputs f (pc :: tbl) None rs2). Proof. intros PC REGS. unfold eqlive_reg. intros r INR. @@ -439,6 +439,36 @@ Proof. simpl; intros; eauto with local. Qed. +Lemma tr_inputs_eqlive_list_None tbl: forall f pc n alive ibf rs1 rs2, + (*rs1 # arg = Vint n ->*) + eqlive_reg (ext alive) rs1 rs2 -> + exit_list_checker (fn_code f) alive tbl = true -> +(*H1 : Regset.mem arg alive = true*) + list_nth_z tbl n = Some pc -> + (fn_code f) ! pc = Some ibf -> + eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> + eqlive_reg (ext (input_regs ibf)) (tid f tbl None rs1) + (tr_inputs f tbl None rs2). +Proof. + induction tbl as [| pc' tbl IHtbl]; try_simplify_someHyps. + autodestruct; try_simplify_someHyps. + - intros; eapply tr_inputs_eqlive_None; eauto. + - rewrite lazy_and_Some_tt_true in H0. + destruct H0 as [EXIT EXIT_REM]. + intros LIST PC. unfold eqlive_reg. + intros r INR. + exploit (IHtbl f pc (Z.pred n) alive ibf rs1 rs2); eauto. + unfold tid. rewrite !tr_inputs_get. + exploit exit_checker_eqlive; eauto. + intros (ibf' & PC' & REGS). + simpl; rewrite PC'. autodestruct. + + intro INRU. apply Regset.mem_2 in INRU. + intros EQR. eapply Regset.union_2 in INRU. + eapply Regset.mem_1 in INRU. erewrite INRU; auto. + + intros. autodestruct. + rewrite (H3 r); auto. +Qed. + Local Hint Resolve tr_inputs_eqlive_None tr_inputs_eqlive_update: core. Lemma cfgsem2fsem_finalstep_simu sp f stk1 stk2 s t fin alive rs1 m rs2 (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s) @@ -494,37 +524,13 @@ Proof. erewrite <- REGS; eauto. + repeat (econstructor; simpl; eauto). -(* - generalize dependent ibf. - generalize dependent n. - generalize dependent alive. - generalize dependent tbl. - induction tbl; simpl; try congruence. - intros alive REGS. rewrite lazy_and_Some_tt_true. - intros (EXIT & EXIT_REM) INARG n ARG. autodestruct. - * try_simplify_someHyps; intros. - - unfold eqlive_reg. intros r INR. - unfold tid. rewrite tr_inputs_get. - simpl. rewrite PC. - exploit Regset.union_3. eapply INR. - intros INRU. eapply Regset.mem_1 in INRU. - erewrite INRU; eauto. - - - * try_simplify_someHyps; intros. - exploit IHtbl. eapply REGS. - eauto. eauto. eauto. eauto. - exploit exit_checker_eqlive; eauto. - intros (ibf' & PC' & REGS''). - unfold eqlive_reg. intros r INR. - unfold tid. rewrite tr_inputs_get. - simpl. rewrite PC'. - exploit Regset.union_3. eapply INR. - intros INRU. eapply Regset.mem_1 in INRU. - erewrite INRU; eauto. - admit.*) -Admitted. + eapply tr_inputs_eqlive_list_None. + eapply REGS. + eauto. + eauto. + eauto. + eauto. +Qed. Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m' (ISTEP: iblock_istep ge sp rs1 m ib rs1' m' None) -- cgit From cb167a462bd4e26ae31694ecf48425ad257fb0f3 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 7 Jun 2021 12:33:49 +0200 Subject: End of liveness proof --- scheduling/BTL_Livecheck.v | 99 +++++++++++++++++++++------------------------- 1 file changed, 45 insertions(+), 54 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index a92f24c3..25a1c545 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -277,21 +277,6 @@ Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr ge b = Some f Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core. -Lemma tr_inputs_eqlive_None f pc tbl ibf rs1 rs2: - (fn_code f) ! pc = Some ibf -> - eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> - eqlive_reg (ext (input_regs ibf)) (tid f (pc :: tbl) None rs1) - (tr_inputs f (pc :: tbl) None rs2). -Proof. - intros PC REGS. - unfold eqlive_reg. intros r INR. - unfold tid. rewrite tr_inputs_get. - simpl. rewrite PC. - exploit Regset.union_3. eapply INR. - intros INRU. eapply Regset.mem_1 in INRU. - erewrite INRU; eauto. -Qed. - Lemma eqlive_reg_update_gso alive rs1 rs2 res r: forall v : val, eqlive_reg (ext alive) rs1 # res <- v rs2 # res <- v -> res <> r -> Regset.In r alive -> @@ -302,24 +287,6 @@ Proof. rewrite !Regmap.gso in INR; auto. Qed. -Lemma tr_inputs_eqlive_update f pc ibf rs1 rs2 res: - (fn_code f) ! pc = Some ibf -> - forall v : val, - eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v -> - eqlive_reg (ext (input_regs ibf)) - (tid f (pc :: nil) (Some res) rs1) # res <- v - (tr_inputs f (pc :: nil) (Some res) rs2) # res <- v. -Proof. - intros PC v REGS. apply eqlive_reg_update. - unfold eqlive_reg. intros r (NRES & INR). - unfold tid. rewrite tr_inputs_get. - simpl. rewrite PC. assert (NRES': res <> r) by auto. - clear NRES. exploit Regset.union_3. eapply INR. - intros INRU. exploit Regset.remove_2; eauto. - intros INRU_RES. eapply Regset.mem_1 in INRU_RES. - erewrite INRU_RES. eapply eqlive_reg_update_gso; eauto. -Qed. - Lemma find_funct_liveness_ok v fd: Genv.find_funct ge v = Some fd -> liveness_ok_fundef fd. Proof. @@ -439,34 +406,63 @@ Proof. simpl; intros; eauto with local. Qed. -Lemma tr_inputs_eqlive_list_None tbl: forall f pc n alive ibf rs1 rs2, - (*rs1 # arg = Vint n ->*) - eqlive_reg (ext alive) rs1 rs2 -> - exit_list_checker (fn_code f) alive tbl = true -> -(*H1 : Regset.mem arg alive = true*) - list_nth_z tbl n = Some pc -> - (fn_code f) ! pc = Some ibf -> - eqlive_reg (ext (input_regs ibf)) rs1 rs2 -> +Lemma tr_inputs_eqlive_None f pc tbl ibf rs1 rs2 + (PC: (fn_code f) ! pc = Some ibf) + (REGS: eqlive_reg (ext (input_regs ibf)) rs1 rs2) + :eqlive_reg (ext (input_regs ibf)) (tid f (pc :: tbl) None rs1) + (tr_inputs f (pc :: tbl) None rs2). +Proof. + unfold eqlive_reg. intros r INR. + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC. + exploit Regset.union_3. eapply INR. + intros INRU. eapply Regset.mem_1 in INRU. + erewrite INRU; eauto. +Qed. + +Lemma tr_inputs_eqlive_list_None tbl: forall f pc n alive ibf rs1 rs2 + (REGS1: eqlive_reg (ext alive) rs1 rs2) + (EXIT_LIST: exit_list_checker (fn_code f) alive tbl = true) + (LIST: list_nth_z tbl n = Some pc) + (PC: (fn_code f) ! pc = Some ibf) + (REGS2: eqlive_reg (ext (input_regs ibf)) rs1 rs2), eqlive_reg (ext (input_regs ibf)) (tid f tbl None rs1) (tr_inputs f tbl None rs2). Proof. induction tbl as [| pc' tbl IHtbl]; try_simplify_someHyps. autodestruct; try_simplify_someHyps. - intros; eapply tr_inputs_eqlive_None; eauto. - - rewrite lazy_and_Some_tt_true in H0. - destruct H0 as [EXIT EXIT_REM]. - intros LIST PC. unfold eqlive_reg. - intros r INR. + - rewrite lazy_and_Some_tt_true in EXIT_LIST. + destruct EXIT_LIST as [EXIT EXIT_REM]. + intros. unfold eqlive_reg. intros r INR. exploit (IHtbl f pc (Z.pred n) alive ibf rs1 rs2); eauto. unfold tid. rewrite !tr_inputs_get. exploit exit_checker_eqlive; eauto. - intros (ibf' & PC' & REGS). + intros (ibf' & PC' & REGS3). simpl; rewrite PC'. autodestruct. + intro INRU. apply Regset.mem_2 in INRU. intros EQR. eapply Regset.union_2 in INRU. eapply Regset.mem_1 in INRU. erewrite INRU; auto. + intros. autodestruct. - rewrite (H3 r); auto. + rewrite (REGS2 r); auto. +Qed. + +Lemma tr_inputs_eqlive_update f pc ibf rs1 rs2 res + (PC: (fn_code f) ! pc = Some ibf) + :forall (v: val) + (REGS: eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v), + eqlive_reg (ext (input_regs ibf)) + (tid f (pc :: nil) (Some res) rs1) # res <- v + (tr_inputs f (pc :: nil) (Some res) rs2) # res <- v. +Proof. + intros. apply eqlive_reg_update. + unfold eqlive_reg. intros r (NRES & INR). + unfold tid. rewrite tr_inputs_get. + simpl. rewrite PC. assert (NRES': res <> r) by auto. + clear NRES. exploit Regset.union_3. eapply INR. + intros INRU. exploit Regset.remove_2; eauto. + intros INRU_RES. eapply Regset.mem_1 in INRU_RES. + erewrite INRU_RES. eapply eqlive_reg_update_gso; eauto. Qed. Local Hint Resolve tr_inputs_eqlive_None tr_inputs_eqlive_update: core. @@ -523,13 +519,8 @@ Proof. + econstructor; eauto. erewrite <- REGS; eauto. + repeat (econstructor; simpl; eauto). - - eapply tr_inputs_eqlive_list_None. - eapply REGS. - eauto. - eauto. - eauto. - eauto. + apply (tr_inputs_eqlive_list_None tbl f pc' (Int.unsigned n) alive ibf rs1 rs2); + auto. Qed. Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m' -- cgit From a7bb27d5125fd2232fa38fe9605c31766a96fa2e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 8 Jun 2021 02:23:24 +0200 Subject: preuve de rexec_rec_correct via les notions history/histOK --- scheduling/BTL_SEsimuref.v | 292 ++++++++++++++++++++++++++++++++++++++------- 1 file changed, 251 insertions(+), 41 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index fdefe205..4f472e3e 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -212,18 +212,17 @@ Inductive rst_simu: rstate -> rstate -> Prop := . (* Comment prend on en compte le ris en cours d'execution ??? *) -Inductive rst_refines ctx: (* Prop -> *) rstate -> sstate -> Prop := - | Reffinal ris sis rfv sfv (*ok: Prop*) - (*OK: ris_ok ctx ris -> ok*) + +Inductive rst_refines ctx: rstate -> sstate -> Prop := + | Reffinal ris sis rfv sfv (RIS: ris_refines ctx ris sis) (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) - : rst_refines ctx (*ok*) (Rfinal ris rfv) (Sfinal sis sfv) - | Refcond cond rargs args sm rifso rifnot ifso ifnot (*ok1 ok2: Prop*) - (*OK1: ok2 -> ok1*) - (RCOND: (* ok2 -> *) seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm) - (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx (*ok2*) rifso ifso) - (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx (*ok2*) rifnot ifnot) - : rst_refines ctx (*ok1*) (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot) + : rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv) + | Refcond cond rargs args sm rifso rifnot ifso ifnot + (RCOND: seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm) + (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx rifso ifso) + (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx rifnot ifnot) + : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot) | Refabort : rst_refines ctx Rabort Sabort . @@ -232,10 +231,9 @@ Local Hint Resolve ris_simu_correct rvf_simu_correct: core. Lemma rst_simu_correct rst1 rst2: rst_simu rst1 rst2 -> - forall f1 f2 (ctx: simu_proof_context f1 f2) st1 st2 (*ok1 ok2*), - rst_refines (*ok1*) (bctx1 ctx) rst1 st1 -> - rst_refines (*ok2*) (bctx2 ctx) rst2 st2 -> - (* ok1 -> ok2 -> *) + forall f1 f2 (ctx: simu_proof_context f1 f2) st1 st2, + rst_refines (bctx1 ctx) rst1 st1 -> + rst_refines (bctx2 ctx) rst2 st2 -> sstate_simu ctx st1 st2. Proof. induction 1; simpl; intros f1 f2 ctx st1 st2 REF1 REF2 sis1 svf1 SEM1; @@ -257,6 +255,34 @@ Proof. * exploit IHrst_simu2; eauto. Qed. + +(* TODO: useless ? +Record routcome := rout { + r_sis: ristate; + r_sfv: sfval; +}. + +Local Open Scope option_monad_scope. + +Fixpoint run_sem_irstate ctx (rst:rstate): option routcome := + match rst with + | Rfinal ris sfv => Some (rout ris sfv) + | Rcond cond args ifso ifnot => + SOME b <- seval_condition ctx cond args Sinit IN + run_sem_irstate ctx (if b then ifso else ifnot) + | Rabort => None + end. + +(* Non: pas assez de "matching" syntaxique entre rst et st pour rst_simu_correct +Definition rst_refines ctx rst st:= + (run_sem_isstate ctx st=None <-> run_sem_irstate ctx rst=None) + /\ (forall ris rfv sis sfv, run_sem_irstate ctx rst = Some (rout ris rfv) -> run_sem_isstate ctx st = Some (sout sis sfv) -> + (ris_refines ctx ris sis) /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv)). + +*) +*) + + (** * Refinement of the symbolic execution *) Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core. @@ -582,29 +608,197 @@ Definition tr_ris := poly_tr (fun f tbl or => transfer_ris (Regset.elements (pre Local Hint Resolve transfer_ris_correct ok_transfer_ris: core. Local Opaque transfer_ris. -Lemma ok_tr_ris ctx f fi ris: - ris_ok ctx (tr_ris f fi ris) +Lemma ok_tr_ris ctx fi ris: + ris_ok ctx (tr_ris (cf ctx) fi ris) <-> (ris_ok ctx ris). Proof. destruct fi; simpl; eauto. Qed. -Lemma ok_tr_ris_imp ctx f fi ris: - ris_ok ctx (tr_ris f fi ris) +Lemma ok_tr_ris_imp ctx fi ris: + ris_ok ctx (tr_ris (cf ctx) fi ris) -> (ris_ok ctx ris). Proof. rewrite ok_tr_ris; auto. Qed. -Lemma tr_ris_correct ctx f fi ris sis: +Lemma tr_ris_correct ctx fi ris sis: ris_refines ctx ris sis -> - ris_refines ctx (tr_ris f fi ris) (tr_sis f fi sis). + ris_refines ctx (tr_ris (cf ctx) fi ris) (tr_sis (cf ctx) fi sis). Proof. intros REF. rewrite <- tr_sis_alt_def. destruct fi; simpl; eauto. Qed. +(** TODO: CHANTIER **) + +Local Open Scope option_monad_scope. + +(* TODO: est-ce qu'un type inductif serait plus simple ? Pas sûr à cause des raisonnements par continuation ? + +Un avantage du fixpoint, c'est que ça pourrait permettre de prouver les propriétés de history de façon plus incrémentale/modulaire. +*) + +Fixpoint history ctx ib sis (k:sistate -> option (list sistate)): option (list sistate) := + match ib with + | BF fin _ => Some (tr_sis (cf ctx) fin sis::nil) + (* basic instructions *) + | Bnop _ => k sis + | Bop op args res _ => k (sexec_op op args res sis) + | Bload TRAP chunk addr args dst _ => k (sexec_load TRAP chunk addr args dst sis) + | Bload NOTRAP chunk addr args dst _ => None + | Bstore chunk addr args src _ => k (sexec_store chunk addr args src sis) + (* composed instructions *) + | Bseq ib1 ib2 => + history ctx ib1 sis (fun sis2 => history ctx ib2 sis2 k) + | Bcond cond args ifso ifnot _ => + let args := list_sval_inj (List.map sis args) in + SOME b <- seval_condition ctx cond args sis.(si_smem) IN + SOME oks <- history ctx (if b then ifso else ifnot) sis k IN + Some (sis::oks) + end + . + +Inductive nested ctx: sistate -> list sistate -> Type := + ns_nil sis: nested ctx sis nil +| ns_cons sis1 sis2 lsis: + (sis_ok ctx sis2 -> sis_ok ctx sis1) -> + nested ctx sis2 lsis -> + nested ctx sis1 (sis2::lsis) + . +Local Hint Constructors nested: core. + +(* TODO: A REVOIR ! +Lemma nested_monotonic ok1 oks: + nested ok1 oks -> + forall (ok2: Prop), (ok1 -> ok2) -> + nested ok2 oks. +Proof. + induction 1; simpl; eauto. +Qed. + +Lemma history_nested ctx ib: + forall k + (CONT: forall sis hsis, (k sis) = Some hsis -> nested (sis_ok ctx sis) hsis) + sis hsis + (RET: history ctx ib sis k = Some hsis), + nested (sis_ok ctx sis) hsis. +Proof. + induction ib; simpl; intros; try_simplify_someHyps. + + econstructor; eauto. admit. + + intros; eapply nested_monotonic; eauto. + admit. + + destruct trap; try_simplify_someHyps. + intros; eapply nested_monotonic; eauto. + admit. + + intros; eapply nested_monotonic; eauto. + admit. + + intros; eapply nested_monotonic. 2: eauto. + eapply IHib1. 2: eauto. + simpl; intros; eapply nested_monotonic; eauto. + + repeat autodestruct; intros; try_simplify_someHyps. +Admitted. + +Inductive is_last {A} (x:A): list A -> Prop := + is_last_refl: is_last x (x::nil) + | is_last_cons a l (NXT: is_last x l): is_last x (a::l) + . + +Lemma nested_last_1 P oks (LAST: is_last P oks): forall ok + (N: nested ok oks), + P -> ok. +Proof. + induction LAST; simpl; intros. + + inv N; auto. + + inv N; eauto. + eapply IHLAST; eauto. + eapply nested_monotonic; eauto. +Qed. + +Lemma nested_last_all P oks (LAST: is_last P oks): forall ok + (N: nested ok oks), + P -> forall Q, List.In Q oks -> Q. +Proof. + induction LAST; simpl; intros. + + intuition subst. inv N; auto. + + inv N. + intuition subst. + * eapply nested_last_1; eauto. + * eapply IHLAST; eauto. +Qed. +Local Hint Constructors is_last: core. + +Lemma run_sem_isstate_okcnd_last ctx ib: forall ks kok + (CONT: forall sis1 sis2 sfv, run_sem_isstate ctx (ks sis1) = Some (sout sis2 sfv) -> exists hsis, kok sis1 = Some hsis /\ is_last (sis_ok ctx sis2) hsis) + sis1 sis2 sfv + (RUN: run_sem_isstate ctx (sexec_rec (cf ctx) ib sis1 ks) = Some (sout sis2 sfv)) + ,exists hsis, history ctx ib sis1 kok = Some hsis /\ is_last (sis_ok ctx sis2) hsis. +Proof. + induction ib; simpl; intros; try_simplify_someHyps. + + destruct trap; simpl; try_simplify_someHyps. + + (* seq *) + intros; eapply IHib1; eauto. + simpl; intros; eapply IHib2; eauto. + + (* cond *) + autodestruct. intros; destruct b. + - exploit IHib1; eauto. + intros (hsis & OK & LAST). + rewrite OK; simpl. + eexists; split; eauto. + - exploit IHib2; eauto. + intros (hsis & OK & LAST). + rewrite OK; simpl. + eexists; split; eauto. +Qed. +*) + +Inductive histOK ctx: list sistate -> sstate -> Prop := + | Final_ok sis sfv + : histOK ctx (sis::nil) (Sfinal sis sfv) + | Cond_ok b sis hsis cond args ifso ifnot + (OK: (sis_ok ctx sis)) + (EVAL: seval_condition ctx cond args sis.(si_smem) = Some b) + (REC: histOK ctx hsis (if b then ifso else ifnot)) + : histOK ctx (sis::hsis) (Scond cond args sis.(si_smem) ifso ifnot) + . +Local Hint Constructors histOK: core. + +(* TODO: pour prouver ça, faut généraliser avec nested + is_last en hypothèse ? + +Lemma run_sem_isstate_okentails ctx ib: forall ks kok + (CONT: forall sis st, ks sis = st -> run_sem_isstate ctx st <> None -> exists hsis, kok sis = Some hsis /\ histOK ctx hsis st) + sis st + (EXEC: sexec_rec (cf ctx) ib sis ks = st) + (OK: run_sem_isstate ctx st <> None) + ,exists hsis, history ctx ib sis kok = Some hsis /\ histOK ctx hsis st. +Proof. + induction ib; simpl; intros; subst; try_simplify_someHyps. + + destruct trap; simpl in *; try_simplify_someHyps. congruence. + + (* seq *) + intros; eapply IHib1; eauto. + simpl; intros; eapply IHib2; eauto. + + (* cond *) + simpl in *. + autodestruct. + intros; destruct b. + - exploit IHib1; eauto. + intros (hsis & OK1 & X). + rewrite OK1; simpl. + eexists; split; eauto. + econstructor; eauto. + admit. + - exploit IHib2; eauto. + intros (hsis & OK2 & LAST). + rewrite OK2; simpl. + eexists; split; eauto. + econstructor; eauto. + admit. +Admitted. +*) + +(** RAFFINEMENT EXEC SYMBOLIQUE **) + Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := match ib with | BF fin _ => Rfinal (tr_ris f fin ris) (sexec_final_sfv fin ris) @@ -625,33 +819,49 @@ Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := end . -Local Hint Constructors rst_refines: core. Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ok_tr_ris_imp rexec_op_correct rexec_load_correct rexec_store_correct: core. -Lemma rexec_rec_correct ctx f ib: - forall ris sis rk k - (CONT: forall rs ss, ris_refines ctx rs ss -> rst_refines ctx (rk rs) (k ss)) +Local Hint Constructors rst_refines: core. + +Lemma rexec_rec_correct ctx ib: + forall kok rk k + (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> histOK ctx hsis st -> rst_refines ctx (rk ris) (k sis)) + ris sis hsis st (REF: ris_refines ctx ris sis) - (*OK: ris_ok ctx ris*) - , rst_refines ctx (rexec_rec f ib ris rk) (sexec_rec f ib sis k). -Proof. - induction ib; simpl; eauto. - - (* load *) autodestruct; eauto. - - intros. - econstructor; eauto. - symmetry. - erewrite seval_condition_valid_preserv; eauto. -Admitted. + (EXEC: sexec_rec (cf ctx) ib sis k = st) + (OK1: history ctx ib sis kok = Some hsis) + (OK2: histOK ctx hsis st) + , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st. +Proof. + induction ib; simpl; try (intros; subst; eauto; fail). + - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. + - (* seq *) + intros; subst. + eapply IHib1; eauto. + simpl. intros; subst. eapply IHib2; eauto. + - (* cond *) + intros rk k kok CONT ris sis hsis st REF EXEC. subst. + autodestruct. + intros EQb. + destruct (history _ _ _ _) eqn: EQl; + intros; try_simplify_someHyps; intros. + inv OK2. + assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV; eauto. } + try_simplify_someHyps; intros. + generalize EVAL. + erewrite seval_condition_valid_preserv, <- seval_condition_refpreserv; eauto. + intros; + econstructor; try_simplify_someHyps. +Qed. Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). -Lemma rexec_correct ctx f ib: - rst_refines ctx (rexec f ib) (sexec f ib). +Lemma rexec_correct ctx ib hsis: + history ctx ib sis_init (fun _ : sistate => None) = Some hsis -> + histOK ctx hsis (sexec_rec (cf ctx) ib sis_init (fun _ : sistate => Sabort)) -> + rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). Proof. - eapply rexec_rec_correct; eauto. - (* econstructor; simpl; auto. congruence. *) + unfold sexec; intros; eapply rexec_rec_correct; eauto. + simpl; congruence. Qed. - - - -- cgit From a3f8b6fc7931e73ca005644124597ff6656cbce7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 8 Jun 2021 14:37:02 +0200 Subject: generalize nested --- scheduling/BTL_SEsimuref.v | 145 +++++++++++++++++++++++---------------------- 1 file changed, 74 insertions(+), 71 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 4f472e3e..5098659d 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -660,16 +660,15 @@ Fixpoint history ctx ib sis (k:sistate -> option (list sistate)): option (list s end . -Inductive nested ctx: sistate -> list sistate -> Type := - ns_nil sis: nested ctx sis nil -| ns_cons sis1 sis2 lsis: - (sis_ok ctx sis2 -> sis_ok ctx sis1) -> - nested ctx sis2 lsis -> - nested ctx sis1 (sis2::lsis) +Inductive nested: Prop -> list Prop -> Type := + ns_nil ok: nested ok nil +| ns_cons (ok1 ok2: Prop) loks: + (ok2 -> ok1) -> + nested ok2 loks -> + nested ok1 (ok2::loks) . Local Hint Constructors nested: core. -(* TODO: A REVOIR ! Lemma nested_monotonic ok1 oks: nested ok1 oks -> forall (ok2: Prop), (ok1 -> ok2) -> @@ -680,10 +679,10 @@ Qed. Lemma history_nested ctx ib: forall k - (CONT: forall sis hsis, (k sis) = Some hsis -> nested (sis_ok ctx sis) hsis) + (CONT: forall sis hsis, (k sis) = Some hsis -> nested (sis_ok ctx sis) (List.map (sis_ok ctx) hsis)) sis hsis (RET: history ctx ib sis k = Some hsis), - nested (sis_ok ctx sis) hsis. + nested (sis_ok ctx sis) (List.map (sis_ok ctx) hsis). Proof. induction ib; simpl; intros; try_simplify_someHyps. + econstructor; eauto. admit. @@ -727,75 +726,73 @@ Proof. * eapply nested_last_1; eauto. * eapply IHLAST; eauto. Qed. + Local Hint Constructors is_last: core. -Lemma run_sem_isstate_okcnd_last ctx ib: forall ks kok - (CONT: forall sis1 sis2 sfv, run_sem_isstate ctx (ks sis1) = Some (sout sis2 sfv) -> exists hsis, kok sis1 = Some hsis /\ is_last (sis_ok ctx sis2) hsis) - sis1 sis2 sfv +Lemma is_last_map {A B} (f: A -> B) x l: is_last x l -> is_last (f x) (List.map f l). +Proof. + induction 1; simpl; eauto. +Qed. + +Lemma run_sem_is_last ctx ib: forall ks kok + (CONT: forall sis1 sis2 sfv hsis, run_sem_isstate ctx (ks sis1) = Some (sout sis2 sfv) -> kok sis1 = Some hsis -> is_last sis2 hsis) + sis1 sis2 sfv hsis (RUN: run_sem_isstate ctx (sexec_rec (cf ctx) ib sis1 ks) = Some (sout sis2 sfv)) - ,exists hsis, history ctx ib sis1 kok = Some hsis /\ is_last (sis_ok ctx sis2) hsis. + (HIST: history ctx ib sis1 kok = Some hsis) + ,is_last sis2 hsis. Proof. induction ib; simpl; intros; try_simplify_someHyps. + destruct trap; simpl; try_simplify_someHyps. + (* seq *) - intros; eapply IHib1; eauto. + intros; eapply IHib1. 2-3: eauto. simpl; intros; eapply IHib2; eauto. + (* cond *) - autodestruct. intros; destruct b. - - exploit IHib1; eauto. - intros (hsis & OK & LAST). - rewrite OK; simpl. - eexists; split; eauto. - - exploit IHib2; eauto. - intros (hsis & OK & LAST). - rewrite OK; simpl. - eexists; split; eauto. + autodestruct. + destruct (history _ _ _ _) eqn: HIST; try_simplify_someHyps. + intros; econstructor. + destruct b; eauto. Qed. -*) - -Inductive histOK ctx: list sistate -> sstate -> Prop := - | Final_ok sis sfv - : histOK ctx (sis::nil) (Sfinal sis sfv) - | Cond_ok b sis hsis cond args ifso ifnot - (OK: (sis_ok ctx sis)) - (EVAL: seval_condition ctx cond args sis.(si_smem) = Some b) - (REC: histOK ctx hsis (if b then ifso else ifnot)) - : histOK ctx (sis::hsis) (Scond cond args sis.(si_smem) ifso ifnot) - . -Local Hint Constructors histOK: core. -(* TODO: pour prouver ça, faut généraliser avec nested + is_last en hypothèse ? - -Lemma run_sem_isstate_okentails ctx ib: forall ks kok - (CONT: forall sis st, ks sis = st -> run_sem_isstate ctx st <> None -> exists hsis, kok sis = Some hsis /\ histOK ctx hsis st) - sis st - (EXEC: sexec_rec (cf ctx) ib sis ks = st) - (OK: run_sem_isstate ctx st <> None) - ,exists hsis, history ctx ib sis kok = Some hsis /\ histOK ctx hsis st. +Lemma run_sem_history_None ctx ib: forall ks kok + (CONT: forall sis, kok sis = None -> run_sem_isstate ctx (ks sis) = None) + sis + (HIST: history ctx ib sis kok = None) + ,run_sem_isstate ctx (sexec_rec (cf ctx) ib sis ks) = None. Proof. - induction ib; simpl; intros; subst; try_simplify_someHyps. - + destruct trap; simpl in *; try_simplify_someHyps. congruence. + induction ib; simpl; intros; try_simplify_someHyps. + + destruct trap; simpl in *; try_simplify_someHyps. + (* seq *) - intros; eapply IHib1; eauto. + intros; eapply IHib1. 2: eauto. simpl; intros; eapply IHib2; eauto. + (* cond *) simpl in *. autodestruct. - intros; destruct b. - - exploit IHib1; eauto. - intros (hsis & OK1 & X). - rewrite OK1; simpl. - eexists; split; eauto. - econstructor; eauto. - admit. - - exploit IHib2; eauto. - intros (hsis & OK2 & LAST). - rewrite OK2; simpl. - eexists; split; eauto. - econstructor; eauto. - admit. -Admitted. -*) + destruct (history _ _ _ _) eqn: X; try congruence. + intros; destruct b; eauto. +Qed. + +(* TODO: bof, "historys_spec" un peu lourd pour l'usage qui en est fait ? *) +Inductive history_spec ctx sis1 sis2: option (list sistate) -> Prop := + history_exist hsis + (NESTED: nested (sis_ok ctx sis1) (List.map (sis_ok ctx) hsis)) + (LAST: is_last sis2 hsis) + : history_spec ctx sis1 sis2 (Some hsis) + . +Local Hint Resolve history_exist: core. + +Lemma run_sem_history ctx ib sis svf + (OK: run_sem_isstate ctx (sexec (cf ctx) ib) = Some (sout sis svf)) + :history_spec ctx sis_init sis (history ctx ib sis_init (fun _ => None)). +Proof. + unfold sexec in *; destruct (history _ _ _ _) eqn: HIST. + + econstructor. + * eapply history_nested; eauto. + simpl; try congruence. + * eapply run_sem_is_last. 2-3: eauto. + simpl; try congruence. + + exploit (run_sem_history_None ctx ib (fun _ => Sabort)); eauto. + congruence. +Qed. (** RAFFINEMENT EXEC SYMBOLIQUE **) @@ -826,12 +823,12 @@ Local Hint Constructors rst_refines: core. Lemma rexec_rec_correct ctx ib: forall kok rk k - (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> histOK ctx hsis st -> rst_refines ctx (rk ris) (k sis)) + (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> (forall sis, List.In sis hsis -> sis_ok ctx sis) -> rst_refines ctx (rk ris) (k sis)) ris sis hsis st (REF: ris_refines ctx ris sis) (EXEC: sexec_rec (cf ctx) ib sis k = st) (OK1: history ctx ib sis kok = Some hsis) - (OK2: histOK ctx hsis st) + (OK2: forall sis, List.In sis hsis -> sis_ok ctx sis) , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st. Proof. induction ib; simpl; try (intros; subst; eauto; fail). @@ -843,13 +840,12 @@ Proof. - (* cond *) intros rk k kok CONT ris sis hsis st REF EXEC. subst. autodestruct. - intros EQb. + intros COND. destruct (history _ _ _ _) eqn: EQl; intros; try_simplify_someHyps; intros. - inv OK2. + exploit (OK2 sis); auto. assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV; eauto. } - try_simplify_someHyps; intros. - generalize EVAL. + generalize COND. erewrite seval_condition_valid_preserv, <- seval_condition_refpreserv; eauto. intros; econstructor; try_simplify_someHyps. @@ -857,11 +853,18 @@ Qed. Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). -Lemma rexec_correct ctx ib hsis: - history ctx ib sis_init (fun _ : sistate => None) = Some hsis -> - histOK ctx hsis (sexec_rec (cf ctx) ib sis_init (fun _ : sistate => Sabort)) -> +Lemma rexec_correct ctx ib sis sfv: + run_sem_isstate ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> + (sis_ok ctx sis) -> rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). Proof. + intros RUN SIS; exploit run_sem_history; eauto. + intros HIST; inv HIST. unfold sexec; intros; eapply rexec_rec_correct; eauto. - simpl; congruence. + + simpl; intros; subst; eauto. + + intros; eapply nested_last_all. + - eapply is_last_map; eauto. + - eapply NESTED. + - eapply SIS. + - eapply List.in_map; eauto. Qed. -- cgit From 1b1e5cfc70899d759ed099832b8821372e024ad0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 9 Jun 2021 06:56:31 +0200 Subject: valid_pointer_preserv dans BTL_SEtheory => opportunites de *grosses* simplifications (à venir ) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- scheduling/BTL_SEsimuref.v | 79 ++++++++-------------------------------------- scheduling/BTL_SEtheory.v | 50 +++++++++++++++++------------ 2 files changed, 42 insertions(+), 87 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 5098659d..40c28934 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -67,11 +67,9 @@ Local Hint Constructors ris_ok: core. Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := { OK_EQUIV: sis_ok ctx sis <-> ris_ok ctx ris; MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem); - REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r); - (* the below invariant allows to evaluate operations in the initial memory of the path instead of the current memory *) - VALID_PRESERV: forall m b ofs, eval_smem ctx sis.(si_smem) = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs + REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r) }. -Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ VALID_PRESERV: core. +Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ: core. Local Hint Constructors ris_refines: core. Record ris_simu ris1 ris2: Prop := { @@ -218,11 +216,11 @@ Inductive rst_refines ctx: rstate -> sstate -> Prop := (RIS: ris_refines ctx ris sis) (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) : rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv) - | Refcond cond rargs args sm rifso rifnot ifso ifnot - (RCOND: seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm) - (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx rifso ifso) - (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx rifnot ifnot) - : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot) + | Refcond cond rargs args rifso rifnot ifso ifnot + (RCOND: seval_condition ctx cond rargs = seval_condition ctx cond args) + (REFso: seval_condition ctx cond rargs = Some true -> rst_refines ctx rifso ifso) + (REFnot: seval_condition ctx cond rargs = Some false -> rst_refines ctx rifnot ifnot) + : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args ifso ifnot) | Refabort : rst_refines ctx Rabort Sabort . @@ -245,7 +243,7 @@ Proof. intros ROK1. exploit ris_simu_ok_preserv; eauto. - (* cond_simu *) - destruct (seval_condition (bctx1 ctx) cond args sm) eqn: SCOND; try congruence. + destruct (seval_condition (bctx1 ctx) cond args) eqn: SCOND; try congruence. generalize RCOND0. erewrite <- seval_condition_preserved, RCOND by eauto. intros SCOND0; rewrite <- SCOND0 in RCOND0. @@ -255,34 +253,6 @@ Proof. * exploit IHrst_simu2; eauto. Qed. - -(* TODO: useless ? -Record routcome := rout { - r_sis: ristate; - r_sfv: sfval; -}. - -Local Open Scope option_monad_scope. - -Fixpoint run_sem_irstate ctx (rst:rstate): option routcome := - match rst with - | Rfinal ris sfv => Some (rout ris sfv) - | Rcond cond args ifso ifnot => - SOME b <- seval_condition ctx cond args Sinit IN - run_sem_irstate ctx (if b then ifso else ifnot) - | Rabort => None - end. - -(* Non: pas assez de "matching" syntaxique entre rst et st pour rst_simu_correct -Definition rst_refines ctx rst st:= - (run_sem_isstate ctx st=None <-> run_sem_irstate ctx rst=None) - /\ (forall ris rfv sis sfv, run_sem_irstate ctx rst = Some (rout ris rfv) -> run_sem_isstate ctx st = Some (sout sis sfv) -> - (ris_refines ctx ris sis) /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv)). - -*) -*) - - (** * Refinement of the symbolic execution *) Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core. @@ -341,7 +311,6 @@ Proof. congruence. + destruct 1; simpl in *. unfold ris_sreg_get; simpl. intros; rewrite PTree.gempty; eauto. - + try_simplify_someHyps. Qed. Definition rset_smem rm (ris:ristate): ristate := @@ -369,12 +338,11 @@ Qed. Lemma rset_mem_correct ctx rm sm ris sis: (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> - (forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) -> ris_refines ctx ris sis -> (ris_ok ctx ris -> eval_smem ctx rm = eval_smem ctx sm) -> ris_refines ctx (rset_smem rm ris) (set_smem sm sis). Proof. - destruct 3; intros. + destruct 2; intros. econstructor; eauto. + rewrite ok_set_mem, ok_rset_mem; intuition congruence. + rewrite ok_rset_mem; intuition eauto. @@ -393,9 +361,6 @@ Lemma rexec_store_correct ctx chunk addr args src ris sis: Proof. intros REF; eapply rset_mem_correct; simpl; eauto. + intros X; rewrite X. repeat autodestruct; eauto. - + intros m b ofs; repeat autodestruct. - intros; erewrite <- Mem.storev_preserv_valid; [| eauto]. - eauto. + intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ, REG_EQ; eauto. Qed. @@ -448,7 +413,7 @@ Qed. Definition rexec_op op args dst (ris:ristate): ristate := let args := list_sval_inj (List.map ris args) in - rset_sreg dst (Sop op args Sinit) ris. + rset_sreg dst (Sop op args) ris. Lemma rexec_op_correct ctx op args dst ris sis: ris_refines ctx ris sis -> @@ -456,10 +421,6 @@ Lemma rexec_op_correct ctx op args dst ris sis: Proof. intros REF; eapply rset_reg_correct; simpl; eauto. intros OK; erewrite eval_list_sval_refpreserv; eauto. - do 2 autodestruct; auto. - + intros. erewrite <- op_valid_pointer_eq; eauto. - + erewrite <- MEM_EQ; eauto. - intros; exploit OK_RMEM; eauto. destruct 1. Qed. Definition rexec_load trap chunk addr args dst (ris:ristate): ristate := @@ -474,24 +435,10 @@ Proof. intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ; eauto. Qed. -Lemma seval_condition_valid_preserv ctx cond args sm - (OK: eval_smem ctx sm <> None) - (VALID_PRESERV: forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) - :seval_condition ctx cond args sm = seval_condition ctx cond args Sinit. -Proof. - unfold seval_condition; autodestruct; simpl; try congruence. - intros EVAL_LIST. - autodestruct; try congruence. - intros. - eapply cond_valid_pointer_eq; intros. - exploit VALID_PRESERV; eauto. -Qed. - Lemma seval_condition_refpreserv ctx cond args ris sis: ris_refines ctx ris sis -> ris_ok ctx ris -> - seval_condition ctx cond (list_sval_inj (map ris args)) Sinit = - seval_condition ctx cond (list_sval_inj (map sis args)) Sinit. + seval_condition ctx cond (list_sval_inj (map ris args)) = seval_condition ctx cond (list_sval_inj (map sis args)). Proof. intros; unfold seval_condition. erewrite eval_list_sval_refpreserv; eauto. @@ -654,7 +601,7 @@ Fixpoint history ctx ib sis (k:sistate -> option (list sistate)): option (list s history ctx ib1 sis (fun sis2 => history ctx ib2 sis2 k) | Bcond cond args ifso ifnot _ => let args := list_sval_inj (List.map sis args) in - SOME b <- seval_condition ctx cond args sis.(si_smem) IN + SOME b <- seval_condition ctx cond args IN SOME oks <- history ctx (if b then ifso else ifnot) sis k IN Some (sis::oks) end @@ -846,7 +793,7 @@ Proof. exploit (OK2 sis); auto. assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV; eauto. } generalize COND. - erewrite seval_condition_valid_preserv, <- seval_condition_refpreserv; eauto. + erewrite <- seval_condition_refpreserv; eauto. intros; econstructor; try_simplify_someHyps. Qed. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 464bb0f0..97dc7fc8 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -29,7 +29,7 @@ Record iblock_exec_context := Bctx { Inductive sval := | Sundef | Sinput (r: reg) - | Sop (op:operation) (lsv: list_sval) (sm: smem) + | Sop (op:operation) (lsv: list_sval) | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) with list_sval := | Snil @@ -55,10 +55,9 @@ Fixpoint eval_sval ctx (sv: sval): option val := match sv with | Sundef => Some Vundef | Sinput r => Some ((crs0 ctx)#r) - | Sop op l sm => + | Sop op l => SOME args <- eval_list_sval ctx l IN - SOME m <- eval_smem ctx sm IN - eval_operation (cge ctx) (csp ctx) op args m + eval_operation (cge ctx) (csp ctx) op args (cm0 ctx) | Sload sm trap chunk addr lsv => match trap with | TRAP => @@ -98,6 +97,16 @@ with eval_smem ctx (sm: smem): option mem := Mem.storev chunk m a sv end. + +Lemma valid_pointer_preserv ctx sm: + forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer (cm0 ctx) b ofs = Mem.valid_pointer m b ofs. +Proof. + induction sm; simpl; intros; try_simplify_someHyps; auto. + repeat autodestruct; intros; erewrite IHsm by reflexivity. + eapply Mem.storev_preserv_valid; eauto. +Qed. +Local Hint Resolve valid_pointer_preserv: core. + Lemma eval_list_sval_inj ctx l (sreg: reg -> sval) rs: (forall r : reg, eval_sval ctx (sreg r) = Some (rs # r)) -> eval_list_sval ctx (list_sval_inj (map sreg l)) = Some (rs ## l). @@ -105,10 +114,9 @@ Proof. intros H; induction l as [|r l]; simpl; repeat autodestruct; auto. Qed. -Definition seval_condition ctx (cond: condition) (lsv: list_sval) (sm: smem) : option bool := +Definition seval_condition ctx (cond: condition) (lsv: list_sval): option bool := SOME args <- eval_list_sval ctx lsv IN - SOME m <- eval_smem ctx sm IN - eval_condition cond args m. + eval_condition cond args (cm0 ctx). (** * Auxiliary definitions on Builtins *) @@ -565,7 +573,7 @@ Qed. Definition sexec_op op args dst sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in - set_sreg dst (Sop op args sis.(si_smem)) sis. + set_sreg dst (Sop op args) sis. Lemma sexec_op_correct ctx op args dst sis rs m v (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = Some v) @@ -577,6 +585,7 @@ Proof. rewrite Regmap.gss; simpl; auto. erewrite eval_list_sval_inj; simpl; auto. try_simplify_someHyps. + intros; erewrite op_valid_pointer_eq; eauto. - intros; rewrite Regmap.gso; auto. Qed. @@ -619,11 +628,11 @@ Qed. Lemma seval_condition_eq ctx cond args sis rs m (SIS : sem_sistate ctx sis rs m) - :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) (si_smem sis) = eval_condition cond rs ## args m. + :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) = eval_condition cond rs ## args m. Proof. destruct SIS as (PRE&MEM®); unfold seval_condition; simpl. erewrite eval_list_sval_inj; simpl; auto. - erewrite MEM; auto. + eapply cond_valid_pointer_eq; eauto. Qed. (** * Compute sistate associated to final values *) @@ -698,7 +707,7 @@ Qed. (* symbolic state *) Inductive sstate := | Sfinal (sis: sistate) (sfv: sfval) - | Scond (cond: condition) (args: list_sval) (sm: smem) (ifso ifnot: sstate) + | Scond (cond: condition) (args: list_sval) (ifso ifnot: sstate) | Sabort . @@ -711,8 +720,8 @@ Record soutcome := sout { Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := match st with | Sfinal sis sfv => Some (sout sis sfv) - | Scond cond args sm ifso ifnot => - SOME b <- seval_condition ctx cond args sm IN + | Scond cond args ifso ifnot => + SOME b <- seval_condition ctx cond args IN run_sem_isstate ctx (if b then ifso else ifnot) | Sabort => None end. @@ -723,10 +732,10 @@ Inductive sem_sstate ctx stk t s: sstate -> Prop := (SIS: sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m) (SFV: sem_sfval ctx stk sfv rs m t s) : sem_sstate ctx stk t s (Sfinal sis sfv) - | sem_Scond b cond args sm ifso ifnot - (SEVAL: seval_condition ctx cond args sm = Some b) + | sem_Scond b cond args ifso ifnot + (SEVAL: seval_condition ctx cond args = Some b) (SELECT: sem_sstate ctx stk t s (if b then ifso else ifnot)) - : sem_sstate ctx stk t s (Scond cond args sm ifso ifnot) + : sem_sstate ctx stk t s (Scond cond args ifso ifnot) (* NB: Sabort: fails to produce a transition *) . @@ -784,7 +793,7 @@ Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := let args := list_sval_inj (List.map sis.(si_sreg) args) in let ifso := sexec_rec f ifso sis k in let ifnot := sexec_rec f ifnot sis k in - Scond cond args sis.(si_smem) ifso ifnot + Scond cond args ifso ifnot end . @@ -954,6 +963,7 @@ Proof. simpl. destruct SIS as (PRE&MEM®). erewrite eval_list_sval_inj; simpl; auto. try_simplify_someHyps. + intros; erewrite op_valid_pointer_eq; eauto. Qed. Lemma sexec_load_TRAP_abort ctx chunk addr args dst sis rs m @@ -1177,7 +1187,6 @@ Proof. induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv) (P1 := fun sm => eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm); simpl; auto. + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. erewrite eval_operation_preserved; eauto. + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. erewrite eval_addressing_preserved; eauto. @@ -1228,12 +1237,11 @@ Proof. reflexivity. Qed. -Lemma seval_condition_preserved cond lsv sm: - seval_condition (bctx1 ctx) cond lsv sm = seval_condition (bctx2 ctx) cond lsv sm. +Lemma seval_condition_preserved cond lsv: + seval_condition (bctx1 ctx) cond lsv = seval_condition (bctx2 ctx) cond lsv. Proof. unfold seval_condition. rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - rewrite smem_eval_preserved; auto. Qed. (* TODO: useless ? -- cgit From 492a71c60c0dbc810f50b75bdf820c2acfc88d7d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 9 Jun 2021 18:45:26 +0200 Subject: fix sstate_simu --- scheduling/BTL_SEsimuref.v | 49 +++++++++++++++++----------------------------- scheduling/BTL_SEtheory.v | 18 ++++++++++++++++- 2 files changed, 35 insertions(+), 32 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 40c28934..98587b66 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -20,21 +20,8 @@ Require Import RTL BTL OptionMonad BTL_SEtheory. (** * Refinement of data-structures and of the specification of the simulation test *) -Record sis_ok ctx (sis: sistate): Prop := { - OK_PRE: (sis.(si_pre) ctx); - OK_SMEM: eval_smem ctx sis.(si_smem) <> None; - OK_SREG: forall (r: reg), eval_sval ctx (si_sreg sis r) <> None -}. Local Hint Resolve OK_PRE OK_SMEM OK_SREG: core. -Local Hint Constructors sis_ok: core. - -Lemma sem_sok ctx sis rs m: - sem_sistate ctx sis rs m -> sis_ok ctx sis. -Proof. - unfold sem_sistate; - econstructor; - intuition congruence. -Qed. +Local Hint Constructors si_ok: core. (* NB: refinement of (symbolic) internal state *) Record ristate := @@ -65,7 +52,7 @@ Local Hint Resolve OK_RMEM OK_RREG: core. Local Hint Constructors ris_ok: core. Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := { - OK_EQUIV: sis_ok ctx sis <-> ris_ok ctx ris; + OK_EQUIV: si_ok ctx sis <-> ris_ok ctx ris; MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem); REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r) }. @@ -97,7 +84,7 @@ Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2: sistate_simu ctx sis1 sis2. Proof. intros RIS REF1 REF2 rs m SEM. - exploit sem_sok; eauto. + exploit sem_si_ok; eauto. erewrite OK_EQUIV; eauto. intros ROK1. exploit ris_simu_ok_preserv; eauto. @@ -238,7 +225,7 @@ Proof. inv REF1; inv REF2; simpl in *; inv SEM1; auto. - (* final_simu *) do 2 eexists; intuition; eauto. - exploit sem_sok; eauto. + exploit sem_si_ok; eauto. erewrite OK_EQUIV; eauto. intros ROK1. exploit ris_simu_ok_preserv; eauto. @@ -249,8 +236,8 @@ Proof. intros SCOND0; rewrite <- SCOND0 in RCOND0. erewrite <- SCOND0. destruct b; simpl. - * exploit IHrst_simu1; eauto. - * exploit IHrst_simu2; eauto. + * intros; exploit IHrst_simu1; eauto. + * intros; exploit IHrst_simu2; eauto. Qed. (** * Refinement of the symbolic execution *) @@ -321,8 +308,8 @@ Definition rset_smem rm (ris:ristate): ristate := |}. Lemma ok_set_mem ctx sm sis: - sis_ok ctx (set_smem sm sis) - <-> (sis_ok ctx sis /\ eval_smem ctx sm <> None). + si_ok ctx (set_smem sm sis) + <-> (si_ok ctx sis /\ eval_smem ctx sm <> None). Proof. split; destruct 1; econstructor; simpl in *; eauto. intuition eauto. @@ -374,8 +361,8 @@ Definition rset_sreg r rsv (ris:ristate): ristate := |}. Lemma ok_set_sreg ctx r sv sis: - sis_ok ctx (set_sreg r sv sis) - <-> (sis_ok ctx sis /\ eval_sval ctx sv <> None). + si_ok ctx (set_sreg r sv sis) + <-> (si_ok ctx sis /\ eval_sval ctx sv <> None). Proof. unfold set_sreg; split. + intros [(SVAL & PRE) SMEM SREG]; simpl in *; split. @@ -493,8 +480,8 @@ Definition transfer_sis (inputs: list reg) (sis:sistate): sistate := si_smem := sis.(si_smem) |}. Lemma ok_transfer_sis ctx inputs sis: - sis_ok ctx (transfer_sis inputs sis) - <-> (sis_ok ctx sis). + si_ok ctx (transfer_sis inputs sis) + <-> (si_ok ctx sis). Proof. unfold transfer_sis. induction inputs as [|r l]; simpl. + split; destruct 1; econstructor; simpl in *; intuition eauto. congruence. @@ -626,10 +613,10 @@ Qed. Lemma history_nested ctx ib: forall k - (CONT: forall sis hsis, (k sis) = Some hsis -> nested (sis_ok ctx sis) (List.map (sis_ok ctx) hsis)) + (CONT: forall sis hsis, (k sis) = Some hsis -> nested (si_ok ctx sis) (List.map (si_ok ctx) hsis)) sis hsis (RET: history ctx ib sis k = Some hsis), - nested (sis_ok ctx sis) (List.map (sis_ok ctx) hsis). + nested (si_ok ctx sis) (List.map (si_ok ctx) hsis). Proof. induction ib; simpl; intros; try_simplify_someHyps. + econstructor; eauto. admit. @@ -721,7 +708,7 @@ Qed. (* TODO: bof, "historys_spec" un peu lourd pour l'usage qui en est fait ? *) Inductive history_spec ctx sis1 sis2: option (list sistate) -> Prop := history_exist hsis - (NESTED: nested (sis_ok ctx sis1) (List.map (sis_ok ctx) hsis)) + (NESTED: nested (si_ok ctx sis1) (List.map (si_ok ctx) hsis)) (LAST: is_last sis2 hsis) : history_spec ctx sis1 sis2 (Some hsis) . @@ -770,12 +757,12 @@ Local Hint Constructors rst_refines: core. Lemma rexec_rec_correct ctx ib: forall kok rk k - (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> (forall sis, List.In sis hsis -> sis_ok ctx sis) -> rst_refines ctx (rk ris) (k sis)) + (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> (forall sis, List.In sis hsis -> si_ok ctx sis) -> rst_refines ctx (rk ris) (k sis)) ris sis hsis st (REF: ris_refines ctx ris sis) (EXEC: sexec_rec (cf ctx) ib sis k = st) (OK1: history ctx ib sis kok = Some hsis) - (OK2: forall sis, List.In sis hsis -> sis_ok ctx sis) + (OK2: forall sis, List.In sis hsis -> si_ok ctx sis) , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st. Proof. induction ib; simpl; try (intros; subst; eauto; fail). @@ -802,7 +789,7 @@ Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). Lemma rexec_correct ctx ib sis sfv: run_sem_isstate ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> - (sis_ok ctx sis) -> + (si_ok ctx sis) -> rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). Proof. intros RUN SIS; exploit run_sem_history; eauto. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 97dc7fc8..9074b071 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -1137,8 +1137,24 @@ Inductive sfv_simu {f1 f2} (ctx: simu_proof_context f1 f2): sfval -> sfval -> Pr Definition sistate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (sis1 sis2:sistate): Prop := forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sem_sistate (bctx2 ctx) sis2 rs m. + +Record si_ok ctx (sis: sistate): Prop := { + OK_PRE: (sis.(si_pre) ctx); + OK_SMEM: eval_smem ctx sis.(si_smem) <> None; + OK_SREG: forall (r: reg), eval_sval ctx (si_sreg sis r) <> None +}. + +Lemma sem_si_ok ctx sis rs m: + sem_sistate ctx sis rs m -> si_ok ctx sis. +Proof. + unfold sem_sistate; + econstructor; + intuition congruence. +Qed. + + Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop := - forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> + forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> si_ok (bctx1 ctx) sis1 -> exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) /\ sistate_simu ctx sis1 sis2 /\ (forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sfv_simu ctx sfv1 sfv2) -- cgit From 6c0efaa166b1acbabcdcd051c5ec389b9b562fe6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 10 Jun 2021 08:15:06 +0200 Subject: remove history --- scheduling/BTL_SEsimuref.v | 299 +++++++++++++++++---------------------------- scheduling/BTL_SEtheory.v | 18 +-- 2 files changed, 121 insertions(+), 196 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 98587b66..2ca36b11 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -8,9 +8,6 @@ on essaye déjà de trouver les bonnes notions de raffinement "*_refines" qui permette de prouver les lemmes "*_simu_correct". - Il faudra ensuite vérifier que l'exécution symbolique préserve ces relations de raffinement ! - TODO: A REVOIR COMPLETEMENT. - - introduire "fake" hash-consed values (à renommer en "refined" plutôt que "fake"). - *) Require Import Coqlib Maps Floats. @@ -18,6 +15,9 @@ Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL BTL OptionMonad BTL_SEtheory. + +Local Open Scope option_monad_scope. + (** * Refinement of data-structures and of the specification of the simulation test *) Local Hint Resolve OK_PRE OK_SMEM OK_SREG: core. @@ -45,12 +45,20 @@ Definition ris_sreg_get (ris: ristate) r: sval := Coercion ris_sreg_get: ristate >-> Funclass. Record ris_ok ctx (ris: ristate) : Prop := { - OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; - OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None + OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; + OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None }. Local Hint Resolve OK_RMEM OK_RREG: core. Local Hint Constructors ris_ok: core. +(* TODO: Is it useful ? +Definition ris_abs (ris: ristate) : sistate := {| + si_pre := fun ctx => ris_ok ctx ris; + si_sreg := ris_sreg_get ris; + si_smem := ris.(ris_smem) +|}. +*) + Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := { OK_EQUIV: si_ok ctx sis <-> ris_ok ctx ris; MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem); @@ -178,6 +186,21 @@ Inductive rstate := | Rabort . + +Record routcome := rout { + _ris: ristate; + _rfv: sfval; +}. + +Fixpoint get_routcome ctx (rst:rstate): option routcome := + match rst with + | Rfinal ris rfv => Some (rout ris rfv) + | Rcond cond args ifso ifnot => + SOME b <- seval_condition ctx cond args IN + get_routcome ctx (if b then ifso else ifnot) + | Rabort => None + end. + Inductive rst_simu: rstate -> rstate -> Prop := | Rfinal_simu ris1 ris2 rfv1 rfv2 (RIS: ris_simu ris1 ris2) @@ -196,6 +219,19 @@ Inductive rst_simu: rstate -> rstate -> Prop := *) . +Lemma rst_simu_lroutcome rst1 rst2: + rst_simu rst1 rst2 -> + forall f1 f2 (ctx: simu_proof_context f1 f2) ris1 rfv1, + get_routcome (bctx1 ctx) rst1 = Some (rout ris1 rfv1) -> + exists ris2 rfv2, get_routcome (bctx2 ctx) rst2 = Some (rout ris2 rfv2) /\ ris_simu ris1 ris2 /\ rfv_simu rfv1 rfv2. +Proof. + induction 1; simpl; intros f1 f2 ctx lris1 lrfv1 ROUT; try_simplify_someHyps. + erewrite <- seval_condition_preserved. + autodestruct. + destruct b; simpl; auto. +Qed. + + (* Comment prend on en compte le ris en cours d'execution ??? *) Inductive rst_refines ctx: rstate -> sstate -> Prop := @@ -212,6 +248,30 @@ Inductive rst_refines ctx: rstate -> sstate -> Prop := : rst_refines ctx Rabort Sabort . +Lemma rst_refines_outcome_up ctx rst st: + rst_refines ctx rst st -> + forall ris rfv, + get_routcome ctx rst = Some (rout ris rfv) -> + exists sis sfv, get_soutcome ctx st = Some (sout sis sfv) /\ ris_refines ctx ris sis /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv). +Proof. + induction 1; simpl; intros lris lrfv ROUT; try_simplify_someHyps. + rewrite RCOND. + autodestruct. + destruct b; simpl; auto. +Qed. + +Lemma rst_refines_outcome_down ctx rst st: + rst_refines ctx rst st -> + forall sis sfv, + get_soutcome ctx st = Some (sout sis sfv) -> + exists ris rfv, get_routcome ctx rst = Some (rout ris rfv) /\ ris_refines ctx ris sis /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv). +Proof. + induction 1; simpl; intros lris lrfv ROUT; try_simplify_someHyps. + rewrite RCOND. + autodestruct. + destruct b; simpl; auto. +Qed. + Local Hint Resolve ris_simu_correct rvf_simu_correct: core. Lemma rst_simu_correct rst1 rst2: @@ -248,7 +308,7 @@ Lemma eval_list_sval_refpreserv ctx args ris sis: ris_refines ctx ris sis -> ris_ok ctx ris -> eval_list_sval ctx (list_sval_inj (map ris args)) = - eval_list_sval ctx (list_sval_inj (map sis args)). + eval_list_sval ctx (list_sval_inj (map (si_sreg sis) args)). Proof. intros REF OK. induction args; simpl; eauto. @@ -565,168 +625,20 @@ Proof. destruct fi; simpl; eauto. Qed. -(** TODO: CHANTIER **) - -Local Open Scope option_monad_scope. - -(* TODO: est-ce qu'un type inductif serait plus simple ? Pas sûr à cause des raisonnements par continuation ? - -Un avantage du fixpoint, c'est que ça pourrait permettre de prouver les propriétés de history de façon plus incrémentale/modulaire. -*) - -Fixpoint history ctx ib sis (k:sistate -> option (list sistate)): option (list sistate) := - match ib with - | BF fin _ => Some (tr_sis (cf ctx) fin sis::nil) - (* basic instructions *) - | Bnop _ => k sis - | Bop op args res _ => k (sexec_op op args res sis) - | Bload TRAP chunk addr args dst _ => k (sexec_load TRAP chunk addr args dst sis) - | Bload NOTRAP chunk addr args dst _ => None - | Bstore chunk addr args src _ => k (sexec_store chunk addr args src sis) - (* composed instructions *) - | Bseq ib1 ib2 => - history ctx ib1 sis (fun sis2 => history ctx ib2 sis2 k) - | Bcond cond args ifso ifnot _ => - let args := list_sval_inj (List.map sis args) in - SOME b <- seval_condition ctx cond args IN - SOME oks <- history ctx (if b then ifso else ifnot) sis k IN - Some (sis::oks) - end - . - -Inductive nested: Prop -> list Prop -> Type := - ns_nil ok: nested ok nil -| ns_cons (ok1 ok2: Prop) loks: - (ok2 -> ok1) -> - nested ok2 loks -> - nested ok1 (ok2::loks) - . -Local Hint Constructors nested: core. - -Lemma nested_monotonic ok1 oks: - nested ok1 oks -> - forall (ok2: Prop), (ok1 -> ok2) -> - nested ok2 oks. -Proof. - induction 1; simpl; eauto. -Qed. - -Lemma history_nested ctx ib: - forall k - (CONT: forall sis hsis, (k sis) = Some hsis -> nested (si_ok ctx sis) (List.map (si_ok ctx) hsis)) - sis hsis - (RET: history ctx ib sis k = Some hsis), - nested (si_ok ctx sis) (List.map (si_ok ctx) hsis). -Proof. - induction ib; simpl; intros; try_simplify_someHyps. - + econstructor; eauto. admit. - + intros; eapply nested_monotonic; eauto. - admit. - + destruct trap; try_simplify_someHyps. - intros; eapply nested_monotonic; eauto. - admit. - + intros; eapply nested_monotonic; eauto. - admit. - + intros; eapply nested_monotonic. 2: eauto. - eapply IHib1. 2: eauto. - simpl; intros; eapply nested_monotonic; eauto. - + repeat autodestruct; intros; try_simplify_someHyps. +Lemma si_ok_tr_sis_down ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis. Admitted. -Inductive is_last {A} (x:A): list A -> Prop := - is_last_refl: is_last x (x::nil) - | is_last_cons a l (NXT: is_last x l): is_last x (a::l) - . - -Lemma nested_last_1 P oks (LAST: is_last P oks): forall ok - (N: nested ok oks), - P -> ok. -Proof. - induction LAST; simpl; intros. - + inv N; auto. - + inv N; eauto. - eapply IHLAST; eauto. - eapply nested_monotonic; eauto. -Qed. - -Lemma nested_last_all P oks (LAST: is_last P oks): forall ok - (N: nested ok oks), - P -> forall Q, List.In Q oks -> Q. -Proof. - induction LAST; simpl; intros. - + intuition subst. inv N; auto. - + inv N. - intuition subst. - * eapply nested_last_1; eauto. - * eapply IHLAST; eauto. -Qed. - -Local Hint Constructors is_last: core. +Lemma si_ok_op_down ctx op args dest sis: si_ok ctx (sexec_op op args dest sis) -> si_ok ctx sis. +Admitted. -Lemma is_last_map {A B} (f: A -> B) x l: is_last x l -> is_last (f x) (List.map f l). -Proof. - induction 1; simpl; eauto. -Qed. +Lemma si_ok_trap_down ctx chunk addr args dest sis: si_ok ctx (sexec_load TRAP chunk addr args dest sis) -> si_ok ctx sis. +Admitted. -Lemma run_sem_is_last ctx ib: forall ks kok - (CONT: forall sis1 sis2 sfv hsis, run_sem_isstate ctx (ks sis1) = Some (sout sis2 sfv) -> kok sis1 = Some hsis -> is_last sis2 hsis) - sis1 sis2 sfv hsis - (RUN: run_sem_isstate ctx (sexec_rec (cf ctx) ib sis1 ks) = Some (sout sis2 sfv)) - (HIST: history ctx ib sis1 kok = Some hsis) - ,is_last sis2 hsis. -Proof. - induction ib; simpl; intros; try_simplify_someHyps. - + destruct trap; simpl; try_simplify_someHyps. - + (* seq *) - intros; eapply IHib1. 2-3: eauto. - simpl; intros; eapply IHib2; eauto. - + (* cond *) - autodestruct. - destruct (history _ _ _ _) eqn: HIST; try_simplify_someHyps. - intros; econstructor. - destruct b; eauto. -Qed. - -Lemma run_sem_history_None ctx ib: forall ks kok - (CONT: forall sis, kok sis = None -> run_sem_isstate ctx (ks sis) = None) - sis - (HIST: history ctx ib sis kok = None) - ,run_sem_isstate ctx (sexec_rec (cf ctx) ib sis ks) = None. -Proof. - induction ib; simpl; intros; try_simplify_someHyps. - + destruct trap; simpl in *; try_simplify_someHyps. - + (* seq *) - intros; eapply IHib1. 2: eauto. - simpl; intros; eapply IHib2; eauto. - + (* cond *) - simpl in *. - autodestruct. - destruct (history _ _ _ _) eqn: X; try congruence. - intros; destruct b; eauto. -Qed. +Lemma si_ok_store_down ctx chunk addr args src sis: si_ok ctx (sexec_store chunk addr args src sis) -> si_ok ctx sis. +Admitted. -(* TODO: bof, "historys_spec" un peu lourd pour l'usage qui en est fait ? *) -Inductive history_spec ctx sis1 sis2: option (list sistate) -> Prop := - history_exist hsis - (NESTED: nested (si_ok ctx sis1) (List.map (si_ok ctx) hsis)) - (LAST: is_last sis2 hsis) - : history_spec ctx sis1 sis2 (Some hsis) - . -Local Hint Resolve history_exist: core. - -Lemma run_sem_history ctx ib sis svf - (OK: run_sem_isstate ctx (sexec (cf ctx) ib) = Some (sout sis svf)) - :history_spec ctx sis_init sis (history ctx ib sis_init (fun _ => None)). -Proof. - unfold sexec in *; destruct (history _ _ _ _) eqn: HIST. - + econstructor. - * eapply history_nested; eauto. - simpl; try congruence. - * eapply run_sem_is_last. 2-3: eauto. - simpl; try congruence. - + exploit (run_sem_history_None ctx ib (fun _ => Sabort)); eauto. - congruence. -Qed. +(* TODO: a revoir: ce n'est pas une bonne idee de mettre ça dans core => ça fait ramer "eauto" !!! *) +Local Hint Resolve si_ok_tr_sis_down si_ok_op_down si_ok_trap_down si_ok_store_down: core. (** RAFFINEMENT EXEC SYMBOLIQUE **) @@ -755,50 +667,61 @@ Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ok_tr_r Local Hint Constructors rst_refines: core. +Lemma sexec_rec_down ctx ib: + forall k + (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) + sis lsis sfv + (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv)) + (OK: si_ok ctx lsis) + ,si_ok ctx sis. +Proof. + induction ib; simpl; try (autodestruct; simpl). + 1-6: try_simplify_someHyps. + - intros. eapply IHib1. 2-3: eauto. + eapply IHib2; eauto. + - intros k CONT sis lsis sfv. + do 2 autodestruct. + + intros; eapply IHib1; eauto. + + intros; eapply IHib2; eauto. +Qed. + Lemma rexec_rec_correct ctx ib: - forall kok rk k - (CONT: forall ris sis hsis st, ris_refines ctx ris sis -> k sis = st -> kok sis = Some hsis -> (forall sis, List.In sis hsis -> si_ok ctx sis) -> rst_refines ctx (rk ris) (k sis)) - ris sis hsis st + forall rk k + (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) + (CONT: forall ris sis lsis sfv st, ris_refines ctx ris sis -> k sis = st -> get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> rst_refines ctx (rk ris) (k sis)) + ris sis lsis sfv st (REF: ris_refines ctx ris sis) (EXEC: sexec_rec (cf ctx) ib sis k = st) - (OK1: history ctx ib sis kok = Some hsis) - (OK2: forall sis, List.In sis hsis -> si_ok ctx sis) + (SOUT: get_soutcome ctx st = Some (sout lsis sfv)) + (OK: si_ok ctx lsis) , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st. Proof. induction ib; simpl; try (intros; subst; eauto; fail). - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. - (* seq *) intros; subst. - eapply IHib1; eauto. - simpl. intros; subst. eapply IHib2; eauto. + eapply IHib1. 3-6: eauto. + + simpl. eapply sexec_rec_down; eauto. + + intros; subst. eapply IHib2; eauto. - (* cond *) - intros rk k kok CONT ris sis hsis st REF EXEC. subst. + intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst. + assert (rOK: ris_ok ctx ris). { + erewrite <- OK_EQUIV. 2: eauto. + eapply sexec_rec_down with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. + } + generalize OUT; clear OUT; simpl. autodestruct. - intros COND. - destruct (history _ _ _ _) eqn: EQl; - intros; try_simplify_someHyps; intros. - exploit (OK2 sis); auto. - assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV; eauto. } - generalize COND. + intros COND; generalize COND. erewrite <- seval_condition_refpreserv; eauto. - intros; econstructor; try_simplify_someHyps. Qed. Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). Lemma rexec_correct ctx ib sis sfv: - run_sem_isstate ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> + get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> (si_ok ctx sis) -> rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). Proof. - intros RUN SIS; exploit run_sem_history; eauto. - intros HIST; inv HIST. - unfold sexec; intros; eapply rexec_rec_correct; eauto. - + simpl; intros; subst; eauto. - + intros; eapply nested_last_all. - - eapply is_last_map; eauto. - - eapply NESTED. - - eapply SIS. - - eapply List.in_map; eauto. + unfold sexec; intros; eapply rexec_rec_correct; eauto; simpl; congruence. Qed. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 9074b071..5a94b235 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -25,6 +25,8 @@ Record iblock_exec_context := Bctx { (** * Syntax and semantics of symbolic values *) +(* TODO: introduire les hash-code directement ici - avec les "fake" smart constructors qui mettent un unknown_hid ? *) + (* symbolic value *) Inductive sval := | Sundef @@ -717,12 +719,12 @@ Record soutcome := sout { _sfv: sfval; }. -Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := +Fixpoint get_soutcome ctx (st:sstate): option soutcome := match st with | Sfinal sis sfv => Some (sout sis sfv) | Scond cond args ifso ifnot => SOME b <- seval_condition ctx cond args IN - run_sem_isstate ctx (if b then ifso else ifnot) + get_soutcome ctx (if b then ifso else ifnot) | Sabort => None end. @@ -742,7 +744,7 @@ Inductive sem_sstate ctx stk t s: sstate -> Prop := Lemma sem_sstate_run ctx stk st t s: sem_sstate ctx stk t s st -> exists sis sfv rs m, - run_sem_isstate ctx st = Some (sout sis sfv) + get_soutcome ctx st = Some (sout sis sfv) /\ sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m /\ sem_sfval ctx stk sfv rs m t s . @@ -753,7 +755,7 @@ Qed. Local Hint Resolve sem_Sfinal: core. Lemma run_sem_sstate ctx st sis sfv: - run_sem_isstate ctx st = Some (sout sis sfv) -> + get_soutcome ctx st = Some (sout sis sfv) -> forall rs m stk s t, sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m -> sem_sfval ctx stk sfv rs m t s -> @@ -1154,8 +1156,8 @@ Qed. Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop := - forall sis1 sfv1, run_sem_isstate (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> si_ok (bctx1 ctx) sis1 -> - exists sis2 sfv2, run_sem_isstate (bctx2 ctx) st2 = Some (sout sis2 sfv2) + forall sis1 sfv1, get_soutcome (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> si_ok (bctx1 ctx) sis1 -> + exists sis2 sfv2, get_soutcome (bctx2 ctx) st2 = Some (sout sis2 sfv2) /\ sistate_simu ctx sis1 sis2 /\ (forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sfv_simu ctx sfv1 sfv2) . @@ -1261,8 +1263,8 @@ Proof. Qed. (* TODO: useless ? -Lemma run_sem_isstate_preserved sis: - run_sem_isstate (bctx1 ctx) sis = run_sem_isstate (bctx2 ctx) sis. +Lemma get_soutcome_preserved sis: + get_soutcome (bctx1 ctx) sis = get_soutcome (bctx2 ctx) sis. Proof. induction sis; simpl; eauto. erewrite seval_condition_preserved. -- cgit From c49cec7b43157a65283dec8bbe343293faa7d012 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 10 Jun 2021 08:48:06 +0200 Subject: fix rst_simu_correct --- scheduling/BTL_SEsimuref.v | 48 ++++++++++++++++++++-------------------------- 1 file changed, 21 insertions(+), 27 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 2ca36b11..f39da275 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -231,9 +231,6 @@ Proof. destruct b; simpl; auto. Qed. - -(* Comment prend on en compte le ris en cours d'execution ??? *) - Inductive rst_refines ctx: rstate -> sstate -> Prop := | Reffinal ris sis rfv sfv (RIS: ris_refines ctx ris sis) @@ -274,30 +271,27 @@ Qed. Local Hint Resolve ris_simu_correct rvf_simu_correct: core. -Lemma rst_simu_correct rst1 rst2: - rst_simu rst1 rst2 -> - forall f1 f2 (ctx: simu_proof_context f1 f2) st1 st2, - rst_refines (bctx1 ctx) rst1 st1 -> - rst_refines (bctx2 ctx) rst2 st2 -> - sstate_simu ctx st1 st2. -Proof. - induction 1; simpl; intros f1 f2 ctx st1 st2 REF1 REF2 sis1 svf1 SEM1; - inv REF1; inv REF2; simpl in *; inv SEM1; auto. - - (* final_simu *) - do 2 eexists; intuition; eauto. - exploit sem_si_ok; eauto. - erewrite OK_EQUIV; eauto. - intros ROK1. - exploit ris_simu_ok_preserv; eauto. - - (* cond_simu *) - destruct (seval_condition (bctx1 ctx) cond args) eqn: SCOND; try congruence. - generalize RCOND0. - erewrite <- seval_condition_preserved, RCOND by eauto. - intros SCOND0; rewrite <- SCOND0 in RCOND0. - erewrite <- SCOND0. - destruct b; simpl. - * intros; exploit IHrst_simu1; eauto. - * intros; exploit IHrst_simu2; eauto. +Lemma rst_simu_correct f1 f2 (ctx: simu_proof_context f1 f2) rst1 rst2 st1 st2 + (SIMU: rst_simu rst1 rst2) + (REF1: forall sis sfv, get_soutcome (bctx1 ctx) st1 = Some (sout sis sfv) -> si_ok (bctx1 ctx) sis -> rst_refines (bctx1 ctx) rst1 st1) + (REF2: forall ris rfv, get_routcome (bctx2 ctx) rst2 = Some (rout ris rfv) -> ris_ok (bctx2 ctx) ris -> rst_refines (bctx2 ctx) rst2 st2) + :sstate_simu ctx st1 st2. +Proof. + intros sis1 sfv1 SOUT OK1. + exploit REF1; eauto. + clear REF1; intros REF1. + exploit rst_refines_outcome_down; eauto. clear REF1 SOUT. + intros (ris1 & rfv1 & ROUT1 & REFI1 & REFF1). + rewrite OK_EQUIV in OK1; eauto. + exploit REFF1; eauto. clear REFF1; intros REFF1. + exploit rst_simu_lroutcome; eauto. + intros (ris2 & rfv2 & ROUT2 & SIMUI & SIMUF). clear ROUT1. + exploit ris_simu_ok_preserv; eauto. + clear OK1. intros OK2. + exploit REF2; eauto. clear REF2; intros REF2. + exploit rst_refines_outcome_up; eauto. + intros (sis2 & sfv2 & SOUT & REFI2 & REFF2). + do 2 eexists; split; eauto. Qed. (** * Refinement of the symbolic execution *) -- cgit From 485a4c0dd450e65745c83e59acdb40b42058e556 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 10 Jun 2021 13:43:14 +0200 Subject: theorem rexec_simu_correct --- scheduling/BTL_SEsimuref.v | 136 ++++++++++++++++++++++++++++++++++++--------- 1 file changed, 111 insertions(+), 25 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index f39da275..852bced0 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -619,7 +619,7 @@ Proof. destruct fi; simpl; eauto. Qed. -Lemma si_ok_tr_sis_down ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis. +Lemma si_ok_tr_down ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis. Admitted. Lemma si_ok_op_down ctx op args dest sis: si_ok ctx (sexec_op op args dest sis) -> si_ok ctx sis. @@ -632,7 +632,26 @@ Lemma si_ok_store_down ctx chunk addr args src sis: si_ok ctx (sexec_store chunk Admitted. (* TODO: a revoir: ce n'est pas une bonne idee de mettre ça dans core => ça fait ramer "eauto" !!! *) -Local Hint Resolve si_ok_tr_sis_down si_ok_op_down si_ok_trap_down si_ok_store_down: core. +Local Hint Resolve si_ok_tr_down si_ok_op_down si_ok_trap_down si_ok_store_down: core. + +Lemma sexec_rec_down ctx ib: + forall k + (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) + sis lsis sfv + (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv)) + (OK: si_ok ctx lsis) + ,si_ok ctx sis. +Proof. + induction ib; simpl; try (autodestruct; simpl). + 1-6: try_simplify_someHyps. + - intros. eapply IHib1. 2-3: eauto. + eapply IHib2; eauto. + - intros k CONT sis lsis sfv. + do 2 autodestruct. + + intros; eapply IHib1; eauto. + + intros; eapply IHib2; eauto. +Qed. + (** RAFFINEMENT EXEC SYMBOLIQUE **) @@ -656,30 +675,15 @@ Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := end . +Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). + + Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ok_tr_ris_imp rexec_op_correct rexec_load_correct rexec_store_correct: core. Local Hint Constructors rst_refines: core. -Lemma sexec_rec_down ctx ib: - forall k - (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) - sis lsis sfv - (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv)) - (OK: si_ok ctx lsis) - ,si_ok ctx sis. -Proof. - induction ib; simpl; try (autodestruct; simpl). - 1-6: try_simplify_someHyps. - - intros. eapply IHib1. 2-3: eauto. - eapply IHib2; eauto. - - intros k CONT sis lsis sfv. - do 2 autodestruct. - + intros; eapply IHib1; eauto. - + intros; eapply IHib2; eauto. -Qed. - -Lemma rexec_rec_correct ctx ib: +Lemma rexec_rec_correct1 ctx ib: forall rk k (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) (CONT: forall ris sis lsis sfv st, ris_refines ctx ris sis -> k sis = st -> get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> rst_refines ctx (rk ris) (k sis)) @@ -710,12 +714,94 @@ Proof. econstructor; try_simplify_someHyps. Qed. -Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). - -Lemma rexec_correct ctx ib sis sfv: +Lemma rexec_correct1 ctx ib sis sfv: get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> (si_ok ctx sis) -> rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). Proof. - unfold sexec; intros; eapply rexec_rec_correct; eauto; simpl; congruence. + unfold sexec; intros; eapply rexec_rec_correct1; eauto; simpl; congruence. +Qed. + + +(** COPIER-COLLER ... Y a-t-il moyen de l'eviter ? **) + +Lemma ris_ok_tr_down ctx fi ris: ris_ok ctx (tr_ris (cf ctx) fi ris) -> ris_ok ctx ris. +Admitted. + +Lemma ris_ok_op_down ctx op args dest ris: ris_ok ctx (rexec_op op args dest ris) -> ris_ok ctx ris. +Admitted. + +Lemma ris_ok_trap_down ctx chunk addr args dest ris: ris_ok ctx (rexec_load TRAP chunk addr args dest ris) -> ris_ok ctx ris. +Admitted. + +Lemma ris_ok_store_down ctx chunk addr args src ris: ris_ok ctx (rexec_store chunk addr args src ris) -> ris_ok ctx ris. +Admitted. + +(* TODO: a revoir: ce n'est pas une bonne idee de mettre ça dans core => ça fait ramer "eauto" !!! *) +Local Hint Resolve ris_ok_tr_down ris_ok_op_down ris_ok_trap_down ris_ok_store_down: core. + +Lemma rexec_rec_down ctx ib: + forall k + (CONT: forall ris lris rfv, get_routcome ctx (k ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris) + ris lris rfv + (ROUT: get_routcome ctx (rexec_rec (cf ctx) ib ris k) = Some (rout lris rfv)) + (OK: ris_ok ctx lris) + ,ris_ok ctx ris. +Proof. + induction ib; simpl; try (autodestruct; simpl). + 1-6: try_simplify_someHyps. + - intros. eapply IHib1. 2-3: eauto. + eapply IHib2; eauto. + - intros k CONT sis lsis sfv. + do 2 autodestruct. + + intros; eapply IHib1; eauto. + + intros; eapply IHib2; eauto. +Qed. + +Lemma rexec_rec_correct2 ctx ib: + forall rk k + (CONTh: forall ris lris rfv, get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris) + (CONT: forall ris sis lris rfv st, ris_refines ctx ris sis -> rk ris = st -> get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> rst_refines ctx (rk ris) (k sis)) + ris sis lris rfv st + (REF: ris_refines ctx ris sis) + (EXEC: rexec_rec (cf ctx) ib ris rk = st) + (SOUT: get_routcome ctx st = Some (rout lris rfv)) + (OK: ris_ok ctx lris) + , rst_refines ctx st (sexec_rec (cf ctx) ib sis k). +Proof. + induction ib; simpl; try (intros; subst; eauto; fail). + - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. + - (* seq *) + intros; subst. + eapply IHib1. 3-6: eauto. + + simpl. eapply rexec_rec_down; eauto. + + intros; subst. eapply IHib2; eauto. + - (* cond *) + intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst. + assert (OK0: ris_ok ctx ris). { + eapply rexec_rec_down with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. + } + generalize OUT; clear OUT; simpl. + autodestruct. + intros COND; generalize COND. + erewrite seval_condition_refpreserv; eauto. + econstructor; try_simplify_someHyps. +Qed. + +Lemma rexec_correct2 ctx ib ris rfv: + get_routcome ctx (rexec (cf ctx) ib) = Some (rout ris rfv) -> + (ris_ok ctx ris) -> + rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib). +Proof. + unfold rexec; intros; eapply rexec_rec_correct2; eauto; simpl; congruence. +Qed. + +Theorem rexec_simu_correct f1 f2 ib1 ib2: + rst_simu (rexec f1 ib1) (rexec f2 ib2) -> + symbolic_simu f1 f2 ib1 ib2. +Proof. + intros SIMU ctx. + eapply rst_simu_correct; eauto. + + intros; eapply rexec_correct1; eauto. + + intros; eapply rexec_correct2; eauto. Qed. -- cgit From 21c979fce33b068ce4b86e67d3d866b203411c6c Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 11 Jun 2021 11:37:41 +0200 Subject: proving the remaining lemma for sexec_rec_okpreserv --- scheduling/BTL_SEsimuref.v | 285 ++++++++++++++++++++++++++------------------- 1 file changed, 162 insertions(+), 123 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 852bced0..da680e63 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -31,7 +31,9 @@ Record ristate := ris_smem: smem; (** For the values in registers: 1) we store a list of sval evaluations - 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval *) + 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval + See [ris_sreg_get] below. + *) ris_input_init: bool; ok_rsval: list sval; ris_sreg:> PTree.t sval @@ -51,18 +53,32 @@ Record ris_ok ctx (ris: ristate) : Prop := { Local Hint Resolve OK_RMEM OK_RREG: core. Local Hint Constructors ris_ok: core. -(* TODO: Is it useful ? -Definition ris_abs (ris: ristate) : sistate := {| - si_pre := fun ctx => ris_ok ctx ris; - si_sreg := ris_sreg_get ris; - si_smem := ris.(ris_smem) -|}. -*) +(** + NOTE that this refinement relation *cannot* be decomposed into a abstraction function of type + ristate -> sistate & an equivalence relation on istate. + + Indeed, any [sis] satisfies [forall ctx r, si_ok ctx sis -> eval_sval ctx (sis r) <> None]. + whereas this is generally not true for [ris] that [forall ctx r, ris_ok ctx ris -> eval_sval ctx (ris r) <> None], + except when, precisely, [ris_refines ris sis]. + + An alternative design enabling to define ris_refines as the composition of an equivalence on sistate + and a abstraction function would consist in constraining sistate with an additional [wf] field: + Record sistate := + { si_pre: iblock_exec_context -> Prop; + si_sreg:> reg -> sval; + si_smem: smem; + si_wf: forall ctx, si_pre ctx -> si_smem <> None /\ forall r, si_sreg r <> None + }. + + Such a constraint should also appear in [ristate]. This is not clear whether this alternative design + would be really simpler. + +*) Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := { OK_EQUIV: si_ok ctx sis <-> ris_ok ctx ris; MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem); - REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r) + REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris r) = eval_sval ctx (sis r) }. Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ: core. Local Hint Constructors ris_refines: core. @@ -70,7 +86,7 @@ Local Hint Constructors ris_refines: core. Record ris_simu ris1 ris2: Prop := { SIMU_FAILS: forall sv, List.In sv ris2.(ok_rsval) -> List.In sv ris1.(ok_rsval); SIMU_MEM: ris1.(ris_smem) = ris2.(ris_smem); - SIMU_REG: forall r, ris_sreg_get ris1 r = ris_sreg_get ris2 r + SIMU_REG: forall r, ris1 r = ris2 r }. Local Hint Resolve SIMU_FAILS SIMU_MEM SIMU_REG: core. Local Hint Constructors ris_simu: core. @@ -186,7 +202,6 @@ Inductive rstate := | Rabort . - Record routcome := rout { _ris: ristate; _rfv: sfval; @@ -257,7 +272,7 @@ Proof. destruct b; simpl; auto. Qed. -Lemma rst_refines_outcome_down ctx rst st: +Lemma rst_refines_outcome_okpreserv ctx rst st: rst_refines ctx rst st -> forall sis sfv, get_soutcome ctx st = Some (sout sis sfv) -> @@ -280,7 +295,7 @@ Proof. intros sis1 sfv1 SOUT OK1. exploit REF1; eauto. clear REF1; intros REF1. - exploit rst_refines_outcome_down; eauto. clear REF1 SOUT. + exploit rst_refines_outcome_okpreserv; eauto. clear REF1 SOUT. intros (ris1 & rfv1 & ROUT1 & REFI1 & REFF1). rewrite OK_EQUIV in OK1; eauto. exploit REFF1; eauto. clear REFF1; intros REFF1. @@ -294,7 +309,114 @@ Proof. do 2 eexists; split; eauto. Qed. -(** * Refinement of the symbolic execution *) + +(** * Properties of the (abstract) operators involved in symbolic execution *) + +Lemma ok_set_mem ctx sm sis: + si_ok ctx (set_smem sm sis) + <-> (si_ok ctx sis /\ eval_smem ctx sm <> None). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. + intuition eauto. +Qed. + +Lemma ok_set_sreg ctx r sv sis: + si_ok ctx (set_sreg r sv sis) + <-> (si_ok ctx sis /\ eval_sval ctx sv <> None). +Proof. + unfold set_sreg; split. + + intros [(SVAL & PRE) SMEM SREG]; simpl in *; split. + - econstructor; eauto. + intros r0; generalize (SREG r0); destruct (Pos.eq_dec r r0); try congruence. + - generalize (SREG r); destruct (Pos.eq_dec r r); try congruence. + + intros ([PRE SMEM SREG] & SVAL). + econstructor; simpl; eauto. + intros r0; destruct (Pos.eq_dec r r0); try congruence. +Qed. + +Lemma si_ok_op_okpreserv ctx op args dest sis: si_ok ctx (sexec_op op args dest sis) -> si_ok ctx sis. +Proof. + unfold sexec_op. rewrite ok_set_sreg. intuition. +Qed. + +Lemma si_ok_load_okpreserv ctx chunk addr args dest sis: si_ok ctx (sexec_load TRAP chunk addr args dest sis) -> si_ok ctx sis. +Proof. + unfold sexec_load. rewrite ok_set_sreg. intuition. +Qed. + +Lemma si_ok_store_okpreserv ctx chunk addr args src sis: si_ok ctx (sexec_store chunk addr args src sis) -> si_ok ctx sis. +Proof. + unfold sexec_store. rewrite ok_set_mem. intuition. +Qed. + +Lemma si_ok_tr_okpreserv ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis. +Proof. + unfold tr_sis. intros OK; inv OK. simpl in *. intuition. +Qed. + +(* These lemmas are very bad for eauto: we put them into a dedicated basis. *) +Local Hint Resolve si_ok_tr_okpreserv si_ok_op_okpreserv si_ok_load_okpreserv si_ok_store_okpreserv: sis_ok. + +Lemma sexec_rec_okpreserv ctx ib: + forall k + (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) + sis lsis sfv + (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv)) + (OK: si_ok ctx lsis) + ,si_ok ctx sis. +Proof. + induction ib; simpl; try (autodestruct; simpl). + 1-6: try_simplify_someHyps; eauto with sis_ok. + - intros. eapply IHib1. 2-3: eauto. + eapply IHib2; eauto. + - intros k CONT sis lsis sfv. + do 2 autodestruct; eauto. +Qed. + +(* an alternative tr_sis *) + +Definition transfer_sis (inputs: list reg) (sis:sistate): sistate := + {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); + si_sreg := transfer_sreg inputs sis; + si_smem := sis.(si_smem) |}. + +Lemma ok_transfer_sis ctx inputs sis: + si_ok ctx (transfer_sis inputs sis) + <-> (si_ok ctx sis). +Proof. + unfold transfer_sis. induction inputs as [|r l]; simpl. + + split; destruct 1; econstructor; simpl in *; intuition eauto. congruence. + + split. + * destruct 1; econstructor; simpl in *; intuition eauto. + * intros X; generalize X. rewrite <- IHl in X; clear IHl. + intros [PRE SMEM SREG]. + econstructor; simpl; eauto. + intros r0; destruct (Pos.eq_dec r r0); try congruence. + intros H; eapply OK_SREG; eauto. +Qed. + +Definition alt_tr_sis := poly_tr (fun f tbl or => transfer_sis (Regset.elements (pre_inputs f tbl or))). + +Lemma tr_sis_alt_def f fi sis: + alt_tr_sis f fi sis = tr_sis f fi sis. +Proof. + unfold tr_sis, str_inputs. destruct fi; simpl; auto. +Qed. + + +(** * Refinement of the symbolic execution (e.g. refinement of [sexec] into [rexec]). + +TODO: Est-ce qu'on garde cette vision purement fonctionnelle de l'implementation ? +=> ça pourrait être pratique quand on va compliquer encore le vérificateur ! + +Du coup, on introduirait une version avec hash-consing qui serait en correspondance +pour une relation d'equivalence sur les ristate ? + +Attention toutefois, il est possible que certaines parties de l'execution soient pénibles à formuler +en programmation fonctionnelle pure (exemple: usage de l'égalité de pointeur comme test d'égalité rapide!) + + +*) Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core. @@ -361,14 +483,6 @@ Definition rset_smem rm (ris:ristate): ristate := ris_sreg:= ris.(ris_sreg) |}. -Lemma ok_set_mem ctx sm sis: - si_ok ctx (set_smem sm sis) - <-> (si_ok ctx sis /\ eval_smem ctx sm <> None). -Proof. - split; destruct 1; econstructor; simpl in *; eauto. - intuition eauto. -Qed. - Lemma ok_rset_mem ctx rm (ris: ristate): (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) -> ris_ok ctx (rset_smem rm ris) @@ -414,20 +528,6 @@ Definition rset_sreg r rsv (ris:ristate): ristate := ris_sreg:= PTree.set r rsv ris.(ris_sreg) (* TODO: A CHANGER *) |}. -Lemma ok_set_sreg ctx r sv sis: - si_ok ctx (set_sreg r sv sis) - <-> (si_ok ctx sis /\ eval_sval ctx sv <> None). -Proof. - unfold set_sreg; split. - + intros [(SVAL & PRE) SMEM SREG]; simpl in *; split. - - econstructor; eauto. - intros r0; generalize (SREG r0); destruct (Pos.eq_dec r r0); try congruence. - - generalize (SREG r); destruct (Pos.eq_dec r r); try congruence. - + intros ([PRE SMEM SREG] & SVAL). - econstructor; simpl; eauto. - intros r0; destruct (Pos.eq_dec r r0); try congruence. -Qed. - Lemma ok_rset_sreg ctx r rsv ris: ris_ok ctx (rset_sreg r rsv ris) <-> (ris_ok ctx ris /\ eval_sval ctx rsv <> None). @@ -528,26 +628,6 @@ Fixpoint transfer_ris (inputs: list reg) (ris:ristate): ristate := | r::l => rseto_sreg r (ris_sreg_get ris r) (transfer_ris l ris) end. -Definition transfer_sis (inputs: list reg) (sis:sistate): sistate := - {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None); - si_sreg := transfer_sreg inputs sis; - si_smem := sis.(si_smem) |}. - -Lemma ok_transfer_sis ctx inputs sis: - si_ok ctx (transfer_sis inputs sis) - <-> (si_ok ctx sis). -Proof. - unfold transfer_sis. induction inputs as [|r l]; simpl. - + split; destruct 1; econstructor; simpl in *; intuition eauto. congruence. - + split. - * destruct 1; econstructor; simpl in *; intuition eauto. - * intros X; generalize X. rewrite <- IHl in X; clear IHl. - intros [PRE SMEM SREG]. - econstructor; simpl; eauto. - intros r0; destruct (Pos.eq_dec r r0); try congruence. - intros H; eapply OK_SREG; eauto. -Qed. - Lemma ok_transfer_ris ctx inputs ris: ris_ok ctx (transfer_ris inputs ris) <-> (ris_ok ctx ris). @@ -583,14 +663,6 @@ Proof. - rewrite ok_rseto_sreg, ok_transfer_ris. auto. Qed. -Definition alt_tr_sis := poly_tr (fun f tbl or => transfer_sis (Regset.elements (pre_inputs f tbl or))). - -Lemma tr_sis_alt_def f fi sis: - alt_tr_sis f fi sis = tr_sis f fi sis. -Proof. - unfold tr_sis, str_inputs. destruct fi; simpl; auto. -Qed. - Definition tr_ris := poly_tr (fun f tbl or => transfer_ris (Regset.elements (pre_inputs f tbl or))). Local Hint Resolve transfer_ris_correct ok_transfer_ris: core. @@ -603,14 +675,11 @@ Proof. destruct fi; simpl; eauto. Qed. -Lemma ok_tr_ris_imp ctx fi ris: - ris_ok ctx (tr_ris (cf ctx) fi ris) - -> (ris_ok ctx ris). +Lemma ris_ok_tr_okpreserv ctx fi ris: ris_ok ctx (tr_ris (cf ctx) fi ris) -> ris_ok ctx ris. Proof. rewrite ok_tr_ris; auto. Qed. - Lemma tr_ris_correct ctx fi ris sis: ris_refines ctx ris sis -> ris_refines ctx (tr_ris (cf ctx) fi ris) (tr_sis (cf ctx) fi sis). @@ -619,40 +688,6 @@ Proof. destruct fi; simpl; eauto. Qed. -Lemma si_ok_tr_down ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis. -Admitted. - -Lemma si_ok_op_down ctx op args dest sis: si_ok ctx (sexec_op op args dest sis) -> si_ok ctx sis. -Admitted. - -Lemma si_ok_trap_down ctx chunk addr args dest sis: si_ok ctx (sexec_load TRAP chunk addr args dest sis) -> si_ok ctx sis. -Admitted. - -Lemma si_ok_store_down ctx chunk addr args src sis: si_ok ctx (sexec_store chunk addr args src sis) -> si_ok ctx sis. -Admitted. - -(* TODO: a revoir: ce n'est pas une bonne idee de mettre ça dans core => ça fait ramer "eauto" !!! *) -Local Hint Resolve si_ok_tr_down si_ok_op_down si_ok_trap_down si_ok_store_down: core. - -Lemma sexec_rec_down ctx ib: - forall k - (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) - sis lsis sfv - (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv)) - (OK: si_ok ctx lsis) - ,si_ok ctx sis. -Proof. - induction ib; simpl; try (autodestruct; simpl). - 1-6: try_simplify_someHyps. - - intros. eapply IHib1. 2-3: eauto. - eapply IHib2; eauto. - - intros k CONT sis lsis sfv. - do 2 autodestruct. - + intros; eapply IHib1; eauto. - + intros; eapply IHib2; eauto. -Qed. - - (** RAFFINEMENT EXEC SYMBOLIQUE **) Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := @@ -678,7 +713,7 @@ Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort). -Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ok_tr_ris_imp +Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ris_ok_tr_okpreserv rexec_op_correct rexec_load_correct rexec_store_correct: core. Local Hint Constructors rst_refines: core. @@ -699,13 +734,13 @@ Proof. - (* seq *) intros; subst. eapply IHib1. 3-6: eauto. - + simpl. eapply sexec_rec_down; eauto. + + simpl. eapply sexec_rec_okpreserv; eauto. + intros; subst. eapply IHib2; eauto. - (* cond *) intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst. assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV. 2: eauto. - eapply sexec_rec_down with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. + eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. } generalize OUT; clear OUT; simpl. autodestruct. @@ -725,22 +760,28 @@ Qed. (** COPIER-COLLER ... Y a-t-il moyen de l'eviter ? **) -Lemma ris_ok_tr_down ctx fi ris: ris_ok ctx (tr_ris (cf ctx) fi ris) -> ris_ok ctx ris. -Admitted. - -Lemma ris_ok_op_down ctx op args dest ris: ris_ok ctx (rexec_op op args dest ris) -> ris_ok ctx ris. -Admitted. +Lemma ris_ok_op_okpreserv ctx op args dest ris: ris_ok ctx (rexec_op op args dest ris) -> ris_ok ctx ris. +Proof. + unfold rexec_op. rewrite ok_rset_sreg. intuition. +Qed. -Lemma ris_ok_trap_down ctx chunk addr args dest ris: ris_ok ctx (rexec_load TRAP chunk addr args dest ris) -> ris_ok ctx ris. -Admitted. +Lemma ris_ok_load_okpreserv ctx chunk addr args dest ris: ris_ok ctx (rexec_load TRAP chunk addr args dest ris) -> ris_ok ctx ris. +Proof. + unfold rexec_load. rewrite ok_rset_sreg. intuition. +Qed. -Lemma ris_ok_store_down ctx chunk addr args src ris: ris_ok ctx (rexec_store chunk addr args src ris) -> ris_ok ctx ris. -Admitted. +Lemma ris_ok_store_okpreserv ctx chunk addr args src ris: ris_ok ctx (rexec_store chunk addr args src ris) -> ris_ok ctx ris. +Proof. + unfold rexec_store. rewrite ok_rset_mem; simpl. + + intuition. + + intros X; rewrite X; simpl. + repeat autodestruct. +Qed. -(* TODO: a revoir: ce n'est pas une bonne idee de mettre ça dans core => ça fait ramer "eauto" !!! *) -Local Hint Resolve ris_ok_tr_down ris_ok_op_down ris_ok_trap_down ris_ok_store_down: core. +(* These lemmas are very bad for eauto: we put them into a dedicated basis. *) +Local Hint Resolve ris_ok_store_okpreserv ris_ok_op_okpreserv ris_ok_load_okpreserv: ris_ok. -Lemma rexec_rec_down ctx ib: +Lemma rexec_rec_okpreserv ctx ib: forall k (CONT: forall ris lris rfv, get_routcome ctx (k ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris) ris lris rfv @@ -749,13 +790,11 @@ Lemma rexec_rec_down ctx ib: ,ris_ok ctx ris. Proof. induction ib; simpl; try (autodestruct; simpl). - 1-6: try_simplify_someHyps. + 1-6: try_simplify_someHyps; eauto with ris_ok. - intros. eapply IHib1. 2-3: eauto. eapply IHib2; eauto. - intros k CONT sis lsis sfv. - do 2 autodestruct. - + intros; eapply IHib1; eauto. - + intros; eapply IHib2; eauto. + do 2 autodestruct; eauto. Qed. Lemma rexec_rec_correct2 ctx ib: @@ -774,12 +813,12 @@ Proof. - (* seq *) intros; subst. eapply IHib1. 3-6: eauto. - + simpl. eapply rexec_rec_down; eauto. + + simpl. eapply rexec_rec_okpreserv; eauto. + intros; subst. eapply IHib2; eauto. - (* cond *) intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst. assert (OK0: ris_ok ctx ris). { - eapply rexec_rec_down with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. + eapply rexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. } generalize OUT; clear OUT; simpl. autodestruct. -- cgit From b03e5a23bbe1370ba0cbb417d81a4505c317da9a Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 11 Jun 2021 15:33:22 +0200 Subject: Roadmap: precision sur le cout de verification des superblocks --- scheduling/BTLroadmap.md | 11 +++++++++-- 1 file changed, 9 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 9dd21be9..23677ba2 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -142,7 +142,6 @@ Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter 1. Verif du "scheduling" (aka analogue de RTLpathScheduler & RTLSchedulerproof). 2. raffinement intermediaire avant le hash-consing ? ça permettrait de décomposer encore davantage la preuve ? -Y a-t-il un moyen simple de faire le tests d'inclusion des fails potentiels à coût linéaire plutôt que quadratique (contrairement à RTLpath) ? ## Port of RTLpath optimizations to BTL @@ -160,6 +159,14 @@ Approche proposée: utiliser un mécanisme de vérification "simple", basée sur Ci-dessous, on liste quelques "techniques" de collaboration oracle/vérificateur pour contrôler l'explosion des chemins. Idée: les conditions comportent une liste d'annotations qui permet le guidage du vérificateur par l'oracle. +**Remarque** déjà sur les superblocs, le test de simulation par +exécution symbolique est quadratique dans le pire cas. Si on a "n" +variables live sur les sorties et "m" sorties, la vérif est en +"n*m". En pratique, le nombre de sorties est limité, notamment à cause +des contraintes pour respecter la structure de superblocs. En RTLpath, +malgré ce coût quadratique théorique dans le pire cas, en pratique, le +vérificateur passe bien à l'échelle sur nos benchmarks. + ### Contrôle des "joins internes" dans le bloc. Si dans le bloc, toute condition a au plus un "predecesseur" (au sens @@ -179,7 +186,7 @@ On considère le bloc BTL ci-dessous (où les `i*` sont des séquences d'instruc if (c2) { i2; goto pc2 } else { i2' } if (c3} { i3; goto pc3 } else { i3'; goto pc3' } -Sur un tel bloc, il n'y a aucun "join interne" (l'exécution symbolique est donc linéaire). +Sur un tel bloc, il n'y a aucun "join interne" (l'exécution symbolique est donc linéaire et le test de simulation quadratique). Mais représenter en RTLpath un tel bloc nécessite au moins 4 blocks (1 superbloc et 3 basic-blocs): i0; c1; i1; c2; i2'; i3'; goto pc3' -- cgit From 9c01536d6bb0696091228cb4a4d52cdcd0c55416 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 11 Jun 2021 16:25:33 +0200 Subject: add hid in BTL_SEtheory: this avoids duplication of types (and should not be harmful) --- scheduling/BTL_SEsimuref.v | 38 +++---- scheduling/BTL_SEtheory.v | 269 ++++++++++++++++++++++++--------------------- 2 files changed, 157 insertions(+), 150 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index da680e63..9c84532b 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -1,13 +1,5 @@ (** Refinement of BTL_SEtheory data-structures in order to introduce (and prove correct) a lower-level specification of the simulation test. - - Ceci est un "bac à sable". - - - On introduit une représentation plus concrète pour les types d'état symbolique [sistate] et [sstate]. - - Etant donné une spécification intuitive "*_simu" pour tester la simulation sur cette représentation des états symboliques, - on essaye déjà de trouver les bonnes notions de raffinement "*_refines" qui permette de prouver les lemmes "*_simu_correct". - - Il faudra ensuite vérifier que l'exécution symbolique préserve ces relations de raffinement ! - *) Require Import Coqlib Maps Floats. @@ -41,7 +33,7 @@ Record ristate := Definition ris_sreg_get (ris: ristate) r: sval := match PTree.get r ris with - | None => if ris_input_init ris then Sinput r else Sundef + | None => if ris_input_init ris then fSinput r else fSundef | Some sv => sv end. Coercion ris_sreg_get: ristate >-> Funclass. @@ -211,7 +203,7 @@ Fixpoint get_routcome ctx (rst:rstate): option routcome := match rst with | Rfinal ris rfv => Some (rout ris rfv) | Rcond cond args ifso ifnot => - SOME b <- seval_condition ctx cond args IN + SOME b <- eval_scondition ctx cond args IN get_routcome ctx (if b then ifso else ifnot) | Rabort => None end. @@ -241,7 +233,7 @@ Lemma rst_simu_lroutcome rst1 rst2: exists ris2 rfv2, get_routcome (bctx2 ctx) rst2 = Some (rout ris2 rfv2) /\ ris_simu ris1 ris2 /\ rfv_simu rfv1 rfv2. Proof. induction 1; simpl; intros f1 f2 ctx lris1 lrfv1 ROUT; try_simplify_someHyps. - erewrite <- seval_condition_preserved. + erewrite <- eval_scondition_preserved. autodestruct. destruct b; simpl; auto. Qed. @@ -252,9 +244,9 @@ Inductive rst_refines ctx: rstate -> sstate -> Prop := (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) : rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv) | Refcond cond rargs args rifso rifnot ifso ifnot - (RCOND: seval_condition ctx cond rargs = seval_condition ctx cond args) - (REFso: seval_condition ctx cond rargs = Some true -> rst_refines ctx rifso ifso) - (REFnot: seval_condition ctx cond rargs = Some false -> rst_refines ctx rifnot ifnot) + (RCOND: eval_scondition ctx cond rargs = eval_scondition ctx cond args) + (REFso: eval_scondition ctx cond rargs = Some true -> rst_refines ctx rifso ifso) + (REFnot: eval_scondition ctx cond rargs = Some false -> rst_refines ctx rifnot ifnot) : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args ifso ifnot) | Refabort : rst_refines ctx Rabort Sabort @@ -464,7 +456,7 @@ Proof. destruct i; simpl; unfold sum_left_map; try autodestruct; eauto. Qed. -Definition ris_init: ristate := {| ris_smem:= Sinit; ris_input_init:=true; ok_rsval := nil; ris_sreg := PTree.empty _ |}. +Definition ris_init: ristate := {| ris_smem:= fSinit; ris_input_init:=true; ok_rsval := nil; ris_sreg := PTree.empty _ |}. Lemma ris_init_correct ctx: ris_refines ctx ris_init sis_init. @@ -507,7 +499,7 @@ Qed. Definition rexec_store chunk addr args src ris: ristate := let args := list_sval_inj (List.map (ris_sreg_get ris) args) in let src := ris_sreg_get ris src in - let rm := Sstore ris.(ris_smem) chunk addr args src in + let rm := fSstore ris.(ris_smem) chunk addr args src in rset_smem rm ris. Lemma rexec_store_correct ctx chunk addr args src ris sis: @@ -554,7 +546,7 @@ Qed. Definition rexec_op op args dst (ris:ristate): ristate := let args := list_sval_inj (List.map ris args) in - rset_sreg dst (Sop op args) ris. + rset_sreg dst (fSop op args) ris. Lemma rexec_op_correct ctx op args dst ris sis: ris_refines ctx ris sis -> @@ -566,7 +558,7 @@ Qed. Definition rexec_load trap chunk addr args dst (ris:ristate): ristate := let args := list_sval_inj (List.map ris args) in - rset_sreg dst (Sload ris.(ris_smem) trap chunk addr args) ris. + rset_sreg dst (fSload ris.(ris_smem) trap chunk addr args) ris. Lemma rexec_load_correct ctx trap chunk addr args dst ris sis: ris_refines ctx ris sis -> @@ -576,12 +568,12 @@ Proof. intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ; eauto. Qed. -Lemma seval_condition_refpreserv ctx cond args ris sis: +Lemma eval_scondition_refpreserv ctx cond args ris sis: ris_refines ctx ris sis -> ris_ok ctx ris -> - seval_condition ctx cond (list_sval_inj (map ris args)) = seval_condition ctx cond (list_sval_inj (map sis args)). + eval_scondition ctx cond (list_sval_inj (map ris args)) = eval_scondition ctx cond (list_sval_inj (map sis args)). Proof. - intros; unfold seval_condition. + intros; unfold eval_scondition. erewrite eval_list_sval_refpreserv; eauto. Qed. @@ -745,7 +737,7 @@ Proof. generalize OUT; clear OUT; simpl. autodestruct. intros COND; generalize COND. - erewrite <- seval_condition_refpreserv; eauto. + erewrite <- eval_scondition_refpreserv; eauto. econstructor; try_simplify_someHyps. Qed. @@ -823,7 +815,7 @@ Proof. generalize OUT; clear OUT; simpl. autodestruct. intros COND; generalize COND. - erewrite seval_condition_refpreserv; eauto. + erewrite eval_scondition_refpreserv; eauto. econstructor; try_simplify_someHyps. Qed. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 5a94b235..2a335790 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -13,54 +13,77 @@ Require Import Coqlib Maps Floats. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL BTL OptionMonad. +Require Export Impure.ImpHCons. +Import HConsing. +(** * Syntax and semantics of symbolic values *) + +(** The semantics of symbolic execution is parametrized by the context of the execution of a block *) Record iblock_exec_context := Bctx { - cge: BTL.genv; - (* cstk: list stackframe; *) (* having the stack here does seem not a good idea *) - cf: function; - csp: val; - crs0: regset; - cm0: mem + cge: BTL.genv; (** usual environment for identifiers *) + cf: function; (** ambient function of the block *) + csp: val; (** stack pointer *) + crs0: regset; (** initial state of registers (at the block entry) *) + cm0: mem (** initial memory state *) }. -(** * Syntax and semantics of symbolic values *) -(* TODO: introduire les hash-code directement ici - avec les "fake" smart constructors qui mettent un unknown_hid ? *) - -(* symbolic value *) +(** symbolic value *) Inductive sval := - | Sundef - | Sinput (r: reg) - | Sop (op:operation) (lsv: list_sval) - | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) -with list_sval := - | Snil - | Scons (sv: sval) (lsv: list_sval) -(* symbolic memory *) + | Sundef (hid: hashcode) + | Sinput (r: reg) (hid: hashcode) + | Sop (op:operation) (lsv: list_sval) (hid: hashcode) + | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (hid: hashcode) +(** list of symbolic values *) +with list_sval := + | Snil (hid: hashcode) + | Scons (sv: sval) (lsv: list_sval) (hid: hashcode) +(** symbolic memory *) with smem := - | Sinit - | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval). + | Sinit (hid: hashcode) + | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval) (hid: hashcode) +. Scheme sval_mut := Induction for sval Sort Prop with list_sval_mut := Induction for list_sval Sort Prop with smem_mut := Induction for smem Sort Prop. +(** "fake" smart-constructors using an [unknown_hid] instead of the one provided by hash-consing. +These smart-constructors are those used in the abstract model of symbolic execution. +They will also appear in the implementation of rewriting rules (in order to avoid hash-consing handling +in proofs of rewriting rules). +*) + +Definition fSundef := Sundef unknown_hid. +Definition fSinput (r: reg) := Sinput r unknown_hid. +Definition fSop (op:operation) (lsv: list_sval) := Sop op lsv unknown_hid. +Definition fSload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) + := Sload sm trap chunk addr lsv unknown_hid. + +Definition fSnil := Snil unknown_hid. +Definition fScons (sv: sval) (lsv: list_sval) := Scons sv lsv unknown_hid. + +Definition fSinit := Sinit unknown_hid. +Definition fSstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval) + := Sstore sm chunk addr lsv srce unknown_hid. + Fixpoint list_sval_inj (l: list sval): list_sval := match l with - | nil => Snil - | v::l => Scons v (list_sval_inj l) + | nil => fSnil + | v::l => fScons v (list_sval_inj l) end. Local Open Scope option_monad_scope. +(** Semantics *) Fixpoint eval_sval ctx (sv: sval): option val := match sv with - | Sundef => Some Vundef - | Sinput r => Some ((crs0 ctx)#r) - | Sop op l => + | Sundef _ => Some Vundef + | Sinput r _ => Some ((crs0 ctx)#r) + | Sop op l _ => SOME args <- eval_list_sval ctx l IN eval_operation (cge ctx) (csp ctx) op args (cm0 ctx) - | Sload sm trap chunk addr lsv => + | Sload sm trap chunk addr lsv _ => match trap with | TRAP => SOME args <- eval_list_sval ctx lsv IN @@ -82,16 +105,16 @@ Fixpoint eval_sval ctx (sv: sval): option val := end with eval_list_sval ctx (lsv: list_sval): option (list val) := match lsv with - | Snil => Some nil - | Scons sv lsv' => + | Snil _ => Some nil + | Scons sv lsv' _ => SOME v <- eval_sval ctx sv IN SOME lv <- eval_list_sval ctx lsv' IN Some (v::lv) end with eval_smem ctx (sm: smem): option mem := match sm with - | Sinit => Some (cm0 ctx) - | Sstore sm chunk addr lsv srce => + | Sinit _ => Some (cm0 ctx) + | Sstore sm chunk addr lsv srce _ => SOME args <- eval_list_sval ctx lsv IN SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN SOME m <- eval_smem ctx sm IN @@ -99,7 +122,12 @@ with eval_smem ctx (sm: smem): option mem := Mem.storev chunk m a sv end. +(** The symbolic memory preserves predicate Mem.valid_pointer with respect to initial memory. + Hence, arithmetic operations and Boolean conditions do not depend on the current memory of the block + (their semantics only depends on the initial memory of the block). + The correctness of this idea is proved on lemmas [sexec_op_correct] and [eval_scondition_eq]. +*) Lemma valid_pointer_preserv ctx sm: forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer (cm0 ctx) b ofs = Mem.valid_pointer m b ofs. Proof. @@ -116,7 +144,7 @@ Proof. intros H; induction l as [|r l]; simpl; repeat autodestruct; auto. Qed. -Definition seval_condition ctx (cond: condition) (lsv: list_sval): option bool := +Definition eval_scondition ctx (cond: condition) (lsv: list_sval): option bool := SOME args <- eval_list_sval ctx lsv IN eval_condition cond args (cm0 ctx). @@ -124,60 +152,60 @@ Definition seval_condition ctx (cond: condition) (lsv: list_sval): option bool : (** * Auxiliary definitions on Builtins *) (* TODO: clean this. Some generic stuffs could be put in [AST.v] *) -Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *) +Section EVAL_BUILTIN_SARG. (* adapted from Events.v *) Variable ctx: iblock_exec_context. Variable m: mem. -Inductive seval_builtin_arg: builtin_arg sval -> val -> Prop := +Inductive eval_builtin_sarg: builtin_arg sval -> val -> Prop := | seval_BA: forall x v, eval_sval ctx x = Some v -> - seval_builtin_arg (BA x) v + eval_builtin_sarg (BA x) v | seval_BA_int: forall n, - seval_builtin_arg (BA_int n) (Vint n) + eval_builtin_sarg (BA_int n) (Vint n) | seval_BA_long: forall n, - seval_builtin_arg (BA_long n) (Vlong n) + eval_builtin_sarg (BA_long n) (Vlong n) | seval_BA_float: forall n, - seval_builtin_arg (BA_float n) (Vfloat n) + eval_builtin_sarg (BA_float n) (Vfloat n) | seval_BA_single: forall n, - seval_builtin_arg (BA_single n) (Vsingle n) + eval_builtin_sarg (BA_single n) (Vsingle n) | seval_BA_loadstack: forall chunk ofs v, Mem.loadv chunk m (Val.offset_ptr (csp ctx) ofs) = Some v -> - seval_builtin_arg (BA_loadstack chunk ofs) v + eval_builtin_sarg (BA_loadstack chunk ofs) v | seval_BA_addrstack: forall ofs, - seval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr (csp ctx) ofs) + eval_builtin_sarg (BA_addrstack ofs) (Val.offset_ptr (csp ctx) ofs) | seval_BA_loadglobal: forall chunk id ofs v, Mem.loadv chunk m (Senv.symbol_address (cge ctx) id ofs) = Some v -> - seval_builtin_arg (BA_loadglobal chunk id ofs) v + eval_builtin_sarg (BA_loadglobal chunk id ofs) v | seval_BA_addrglobal: forall id ofs, - seval_builtin_arg (BA_addrglobal id ofs) (Senv.symbol_address (cge ctx) id ofs) + eval_builtin_sarg (BA_addrglobal id ofs) (Senv.symbol_address (cge ctx) id ofs) | seval_BA_splitlong: forall hi lo vhi vlo, - seval_builtin_arg hi vhi -> seval_builtin_arg lo vlo -> - seval_builtin_arg (BA_splitlong hi lo) (Val.longofwords vhi vlo) + eval_builtin_sarg hi vhi -> eval_builtin_sarg lo vlo -> + eval_builtin_sarg (BA_splitlong hi lo) (Val.longofwords vhi vlo) | seval_BA_addptr: forall a1 a2 v1 v2, - seval_builtin_arg a1 v1 -> seval_builtin_arg a2 v2 -> - seval_builtin_arg (BA_addptr a1 a2) + eval_builtin_sarg a1 v1 -> eval_builtin_sarg a2 v2 -> + eval_builtin_sarg (BA_addptr a1 a2) (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2) . -Definition seval_builtin_args (al: list (builtin_arg sval)) (vl: list val) : Prop := - list_forall2 seval_builtin_arg al vl. +Definition eval_builtin_sargs (al: list (builtin_arg sval)) (vl: list val) : Prop := + list_forall2 eval_builtin_sarg al vl. -Lemma seval_builtin_arg_determ: - forall a v, seval_builtin_arg a v -> forall v', seval_builtin_arg a v' -> v' = v. +Lemma eval_builtin_sarg_determ: + forall a v, eval_builtin_sarg a v -> forall v', eval_builtin_sarg a v' -> v' = v. Proof. induction 1; intros v' EV; inv EV; try congruence. f_equal; eauto. - apply IHseval_builtin_arg1 in H3. apply IHseval_builtin_arg2 in H5. subst; auto. + apply IHeval_builtin_sarg1 in H3. apply IHeval_builtin_sarg2 in H5. subst; auto. Qed. -Lemma eval_builtin_args_determ: - forall al vl, seval_builtin_args al vl -> forall vl', seval_builtin_args al vl' -> vl' = vl. +Lemma eval_builtin_sargs_determ: + forall al vl, eval_builtin_sargs al vl -> forall vl', eval_builtin_sargs al vl' -> vl' = vl. Proof. - induction 1; intros v' EV; inv EV; f_equal; eauto using seval_builtin_arg_determ. + induction 1; intros v' EV; inv EV; f_equal; eauto using eval_builtin_sarg_determ. Qed. -End SEVAL_BUILTIN_ARG. +End EVAL_BUILTIN_SARG. (* NB: generic function that could be put into [AST] file *) Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B := @@ -195,10 +223,10 @@ Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2) end. -Lemma seval_builtin_arg_correct ctx rs m sreg: forall arg varg, +Lemma eval_builtin_sarg_correct ctx rs m sreg: forall arg varg, (forall r, eval_sval ctx (sreg r) = Some rs # r) -> eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg -> - seval_builtin_arg ctx m (builtin_arg_map sreg arg) varg. + eval_builtin_sarg ctx m (builtin_arg_map sreg arg) varg. Proof. induction arg. all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence). @@ -208,20 +236,20 @@ Proof. eapply IHarg1; eauto. eapply IHarg2; eauto. Qed. -Lemma seval_builtin_args_correct ctx rs m sreg args vargs: +Lemma eval_builtin_sargs_correct ctx rs m sreg args vargs: (forall r, eval_sval ctx (sreg r) = Some rs # r) -> eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs -> - seval_builtin_args ctx m (map (builtin_arg_map sreg) args) vargs. + eval_builtin_sargs ctx m (map (builtin_arg_map sreg) args) vargs. Proof. induction 2. - constructor. - simpl. constructor; [| assumption]. - eapply seval_builtin_arg_correct; eauto. + eapply eval_builtin_sarg_correct; eauto. Qed. -Lemma seval_builtin_arg_exact ctx rs m sreg: forall arg varg, +Lemma eval_builtin_sarg_exact ctx rs m sreg: forall arg varg, (forall r, eval_sval ctx (sreg r) = Some rs # r) -> - seval_builtin_arg ctx m (builtin_arg_map sreg arg) varg -> + eval_builtin_sarg ctx m (builtin_arg_map sreg arg) varg -> eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg. Proof. induction arg. @@ -233,16 +261,16 @@ Proof. eapply IHarg1; eauto. eapply IHarg2; eauto. Qed. -Lemma seval_builtin_args_exact ctx rs m sreg: forall args vargs, +Lemma eval_builtin_sargs_exact ctx rs m sreg: forall args vargs, (forall r, eval_sval ctx (sreg r) = Some rs # r) -> - seval_builtin_args ctx m (map (builtin_arg_map sreg) args) vargs -> + eval_builtin_sargs ctx m (map (builtin_arg_map sreg) args) vargs -> eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs. Proof. induction args. - simpl. intros. inv H0. constructor. - intros vargs SEVAL BARG. simpl in BARG. inv BARG. constructor; [| eapply IHargs; eauto]. - eapply seval_builtin_arg_exact; eauto. + eapply eval_builtin_sarg_exact; eauto. Qed. Fixpoint eval_builtin_sval ctx bsv := @@ -285,7 +313,7 @@ Lemma eval_builtin_sval_arg ctx bs: forall ba m v, eval_builtin_sval ctx bs = Some ba -> eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v -> - seval_builtin_arg ctx m bs v. + eval_builtin_sarg ctx m bs v. Proof. induction bs; simpl; try (intros ba m v H; inversion H; subst; clear H; @@ -309,8 +337,8 @@ Proof. econstructor; eauto. Qed. -Lemma seval_builtin_arg_sval ctx m v: forall bs, - seval_builtin_arg ctx m bs v -> +Lemma eval_builtin_sarg_sval ctx m v: forall bs, + eval_builtin_sarg ctx m bs v -> exists ba, eval_builtin_sval ctx bs = Some ba /\ eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v. @@ -321,13 +349,13 @@ Proof. - eexists. constructor. + simpl. rewrite H. reflexivity. + constructor. - - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1). - destruct IHseval_builtin_arg2 as (ba2 & A2 & B2). + - destruct IHeval_builtin_sarg1 as (ba1 & A1 & B1). + destruct IHeval_builtin_sarg2 as (ba2 & A2 & B2). eexists. constructor. + simpl. rewrite A1. rewrite A2. reflexivity. + constructor; assumption. - - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1). - destruct IHseval_builtin_arg2 as (ba2 & A2 & B2). + - destruct IHeval_builtin_sarg1 as (ba1 & A1 & B1). + destruct IHeval_builtin_sarg2 as (ba2 & A2 & B2). eexists. constructor. + simpl. rewrite A1. rewrite A2. reflexivity. + constructor; assumption. @@ -337,9 +365,9 @@ Lemma eval_builtin_sval_args ctx lbs: forall lba m v, eval_list_builtin_sval ctx lbs = Some lba -> list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba v -> - seval_builtin_args ctx m lbs v. + eval_builtin_sargs ctx m lbs v. Proof. - unfold seval_builtin_args; induction lbs; simpl; intros lba m v. + unfold eval_builtin_sargs; induction lbs; simpl; intros lba m v. - intros H; inversion H; subst; clear H. intros H; inversion H. econstructor. - destruct (eval_builtin_sval _ _) eqn:SV; try congruence. @@ -350,8 +378,8 @@ Proof. eapply eval_builtin_sval_arg; eauto. Qed. -Lemma seval_builtin_args_sval ctx m lv: forall lbs, - seval_builtin_args ctx m lbs lv -> +Lemma eval_builtin_sargs_sval ctx m lv: forall lbs, + eval_builtin_sargs ctx m lbs lv -> exists lba, eval_list_builtin_sval ctx lbs = Some lba /\ list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba lv. @@ -361,29 +389,29 @@ Proof. + simpl. reflexivity. + constructor. - destruct IHlist_forall2 as (lba & A & B). - apply seval_builtin_arg_sval in H. destruct H as (ba & A' & B'). + apply eval_builtin_sarg_sval in H. destruct H as (ba & A' & B'). eexists. constructor. + simpl. rewrite A'. rewrite A. reflexivity. + constructor; assumption. Qed. Lemma eval_builtin_sval_correct ctx m: forall bs1 v bs2, - seval_builtin_arg ctx m bs1 v -> + eval_builtin_sarg ctx m bs1 v -> (eval_builtin_sval ctx bs1) = (eval_builtin_sval ctx bs2) -> - seval_builtin_arg ctx m bs2 v. + eval_builtin_sarg ctx m bs2 v. Proof. - intros. exploit seval_builtin_arg_sval; eauto. + intros. exploit eval_builtin_sarg_sval; eauto. intros (ba & X1 & X2). eapply eval_builtin_sval_arg; eauto. congruence. Qed. Lemma eval_list_builtin_sval_correct ctx m vargs: forall lbs1, - seval_builtin_args ctx m lbs1 vargs -> + eval_builtin_sargs ctx m lbs1 vargs -> forall lbs2, (eval_list_builtin_sval ctx lbs1) = (eval_list_builtin_sval ctx lbs2) -> - seval_builtin_args ctx m lbs2 vargs. + eval_builtin_sargs ctx m lbs2 vargs. Proof. - intros. exploit seval_builtin_args_sval; eauto. + intros. exploit eval_builtin_sargs_sval; eauto. intros (ba & X1 & X2). eapply eval_builtin_sval_args; eauto. congruence. @@ -391,6 +419,9 @@ Qed. (** * Symbolic (final) value of a block *) +(** TODO: faut-il hash-conser les valeurs symboliques finales. Pas très utile si pas de join interne. +Mais peut être utile dans le cas contraire. *) + Inductive sfval := | Sgoto (pc: exit) | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:exit) @@ -434,7 +465,7 @@ Inductive sem_sfval ctx stk: sfval -> regset -> mem -> trace -> state -> Prop := sem_sfval ctx stk (Stailcall sig svos lsv) rs m E0 (Callstate stk fd args m') | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs: - seval_builtin_args ctx m sargs vargs -> + eval_builtin_sargs ctx m sargs vargs -> external_call ef (cge ctx) vargs m t vres m' -> sem_sfval ctx stk (Sbuiltin ef sargs res pc) rs m t (State stk (cf ctx) (csp ctx) pc (regmap_setres res vres (tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs)) m') @@ -496,12 +527,12 @@ Proof. rewrite REG; eauto. - erewrite eval_list_sval_inj; simpl; auto. + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto. - eapply seval_builtin_args_correct; eauto. + eapply eval_builtin_sargs_correct; eauto. + (* Bjumptable *) intros. eapply exec_Sjumptable; eauto. congruence. Qed. Local Hint Constructors final_step: core. -Local Hint Resolve seval_builtin_args_exact: core. +Local Hint Resolve eval_builtin_sargs_exact: core. Lemma sexec_final_sfv_exact ctx stk i sis t rs m s: sem_sistate ctx sis rs m -> @@ -530,7 +561,7 @@ Qed. (** * symbolic execution of basic instructions *) -Definition sis_init : sistate := {| si_pre:= fun _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. +Definition sis_init : sistate := {| si_pre:= fun _ => True; si_sreg:= fun r => fSinput r; si_smem:= fSinit |}. Lemma sis_init_correct ctx: sem_sistate ctx sis_init (crs0 ctx) (cm0 ctx). @@ -575,7 +606,7 @@ Qed. Definition sexec_op op args dst sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in - set_sreg dst (Sop op args) sis. + set_sreg dst (fSop op args) sis. Lemma sexec_op_correct ctx op args dst sis rs m v (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = Some v) @@ -593,7 +624,7 @@ Qed. Definition sexec_load trap chunk addr args dst sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in - set_sreg dst (Sload sis.(si_smem) trap chunk addr args) sis. + set_sreg dst (fSload sis.(si_smem) trap chunk addr args) sis. Lemma sexec_load_TRAP_correct ctx chunk addr args dst sis rs m a v (EVAL: eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a) @@ -612,7 +643,7 @@ Qed. Definition sexec_store chunk addr args src sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in let src := sis.(si_sreg) src in - let sm := Sstore sis.(si_smem) chunk addr args src in + let sm := fSstore sis.(si_smem) chunk addr args src in set_smem sm sis. Lemma sexec_store_correct ctx chunk addr args src sis rs m m' a @@ -628,11 +659,11 @@ Proof. rewrite REG; auto. Qed. -Lemma seval_condition_eq ctx cond args sis rs m +Lemma eval_scondition_eq ctx cond args sis rs m (SIS : sem_sistate ctx sis rs m) - :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) = eval_condition cond rs ## args m. + :eval_scondition ctx cond (list_sval_inj (map (si_sreg sis) args)) = eval_condition cond rs ## args m. Proof. - destruct SIS as (PRE&MEM®); unfold seval_condition; simpl. + destruct SIS as (PRE&MEM®); unfold eval_scondition; simpl. erewrite eval_list_sval_inj; simpl; auto. eapply cond_valid_pointer_eq; eauto. Qed. @@ -640,7 +671,7 @@ Qed. (** * Compute sistate associated to final values *) Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval := match inputs with - | nil => fun r => Sundef + | nil => fun r => fSundef | r1::l => fun r => if Pos.eq_dec r1 r then sreg r1 else transfer_sreg l sreg r end. @@ -723,7 +754,7 @@ Fixpoint get_soutcome ctx (st:sstate): option soutcome := match st with | Sfinal sis sfv => Some (sout sis sfv) | Scond cond args ifso ifnot => - SOME b <- seval_condition ctx cond args IN + SOME b <- eval_scondition ctx cond args IN get_soutcome ctx (if b then ifso else ifnot) | Sabort => None end. @@ -735,7 +766,7 @@ Inductive sem_sstate ctx stk t s: sstate -> Prop := (SFV: sem_sfval ctx stk sfv rs m t s) : sem_sstate ctx stk t s (Sfinal sis sfv) | sem_Scond b cond args ifso ifnot - (SEVAL: seval_condition ctx cond args = Some b) + (SEVAL: eval_scondition ctx cond args = Some b) (SELECT: sem_sstate ctx stk t s (if b then ifso else ifnot)) : sem_sstate ctx stk t s (Scond cond args ifso ifnot) (* NB: Sabort: fails to produce a transition *) @@ -768,14 +799,10 @@ Proof. Qed. -(** * Idée de l'execution symbolique en Continuation Passing Style - -[k] ci-dessous est la continuation (c-a-d. la suite de la construction de l'arbre qu'on va appliquer dans chaque branche) +(** * Model of Symbolic Execution with Continuation Passing Style -Rem: si manipuler une telle continuation s'avère compliqué dans les preuves, -il faudra faire autrement dans le modèle -- par exemple avec une structure de donnée -pour représenter l'ensemble des chemins. -(même si on peut conserver le CPS dans l'implem finale, pour l'efficacité). +Parameter [k] is the continuation, i.e. the [sstate] construction that will be applied in each execution branch. +Its input parameter is the symbolic internal state of the branch. *) @@ -821,7 +848,7 @@ Proof. (* condition *) all: intros; eapply sem_Scond; eauto; [ - erewrite seval_condition_eq; eauto | + erewrite eval_scondition_eq; eauto | replace (if b then sexec_rec (cf ctx) ifso sis k else sexec_rec (cf ctx) ifnot sis k) with (sexec_rec (cf ctx) (if b then ifso else ifnot) sis k); try autodestruct; eauto ]. Qed. @@ -1044,7 +1071,7 @@ Proof. exploit IHib2; eauto. - (* Bcond *) inv SEXEC. - erewrite seval_condition_eq in SEVAL; eauto. + erewrite eval_scondition_eq in SEVAL; eauto. rewrite SEVAL. destruct b. + exploit IHib1; eauto. @@ -1139,7 +1166,6 @@ Inductive sfv_simu {f1 f2} (ctx: simu_proof_context f1 f2): sfval -> sfval -> Pr Definition sistate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (sis1 sis2:sistate): Prop := forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sem_sistate (bctx2 ctx) sis2 rs m. - Record si_ok ctx (sis: sistate): Prop := { OK_PRE: (sis.(si_pre) ctx); OK_SMEM: eval_smem ctx sis.(si_smem) <> None; @@ -1154,7 +1180,6 @@ Proof. intuition congruence. Qed. - Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop := forall sis1 sfv1, get_soutcome (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> si_ok (bctx1 ctx) sis1 -> exists sis2 sfv2, get_soutcome (bctx2 ctx) st2 = Some (sout sis2 sfv2) @@ -1255,23 +1280,13 @@ Proof. reflexivity. Qed. -Lemma seval_condition_preserved cond lsv: - seval_condition (bctx1 ctx) cond lsv = seval_condition (bctx2 ctx) cond lsv. +Lemma eval_scondition_preserved cond lsv: + eval_scondition (bctx1 ctx) cond lsv = eval_scondition (bctx2 ctx) cond lsv. Proof. - unfold seval_condition. + unfold eval_scondition. rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. Qed. -(* TODO: useless ? -Lemma get_soutcome_preserved sis: - get_soutcome (bctx1 ctx) sis = get_soutcome (bctx2 ctx) sis. -Proof. - induction sis; simpl; eauto. - erewrite seval_condition_preserved. - repeat (autodestruct; auto). -Qed. -*) - (* additional preservation properties under this additional hypothesis *) Hypothesis senv_preserved_BTL: Senv.equiv (sge1 ctx) (sge2 ctx). @@ -1288,9 +1303,9 @@ Proof. reflexivity. Qed. -Lemma seval_builtin_arg_preserved m: forall bs varg, - seval_builtin_arg (bctx1 ctx) m bs varg -> - seval_builtin_arg (bctx2 ctx) m bs varg. +Lemma eval_builtin_sarg_preserved m: forall bs varg, + eval_builtin_sarg (bctx1 ctx) m bs varg -> + eval_builtin_sarg (bctx2 ctx) m bs varg. Proof. induction 1; simpl. all: try (constructor; auto). @@ -1299,12 +1314,12 @@ Proof. - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal. Qed. -Lemma seval_builtin_args_preserved m lbs vargs: - seval_builtin_args (bctx1 ctx) m lbs vargs -> - seval_builtin_args (bctx2 ctx) m lbs vargs. +Lemma eval_builtin_sargs_preserved m lbs vargs: + eval_builtin_sargs (bctx1 ctx) m lbs vargs -> + eval_builtin_sargs (bctx2 ctx) m lbs vargs. Proof. induction 1; constructor; eauto. - eapply seval_builtin_arg_preserved; auto. + eapply eval_builtin_sarg_preserved; auto. Qed. End SymbValPreserved. -- cgit From 6199b498cc309d371ba773966a34ff6a9ee47c2e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 11 Jun 2021 17:17:02 +0200 Subject: stub match_states --- scheduling/BTL_Schedulerproof.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index c8c2a76f..3811b170 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -22,10 +22,12 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. +Inductive match_states (s1 s2: BTL.state): Prop :=. (* STUB *) + Theorem transf_program_correct: forward_simulation (fsem prog) (fsem tprog). Proof. - eapply forward_simulation_step with equiv_state. (* lock-step with respect to equiv_states *) + eapply forward_simulation_step with match_states. (* lock-step with respect to match_states *) Admitted. End PRESERVATION. -- cgit From 8252a8b7678ca4b82191ce0159b93b976f8c58d9 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 14 Jun 2021 13:35:51 +0200 Subject: begin scheduler BTL proof --- scheduling/BTL_Scheduler.v | 1 + scheduling/BTL_Schedulerproof.v | 146 ++++++++++++++++++++++++++++++++++++++-- 2 files changed, 142 insertions(+), 5 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index ec83b3c1..8240c861 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -37,6 +37,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; + (*exists_entrypoint: exists ibf, (fn_code f1)!(fn_entrypoint f1) = Some ibf;*) (* 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); diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 3811b170..f186581f 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -1,7 +1,7 @@ Require Import AST Linking Values Maps Globalenvs Smallstep Registers. -Require Import Coqlib Maps Events Errors Op. +Require Import Coqlib Maps Events Errors Op OptionMonad. Require Import RTL BTL BTL_SEtheory. -Require Import BTL_Scheduler. +Require Import BTLmatchRTL BTL_Scheduler. Definition match_prog (p tp: program) := match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp. @@ -22,12 +22,148 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. -Inductive match_states (s1 s2: BTL.state): Prop :=. (* STUB *) +Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s. +Proof. + rewrite <- (Genv.find_symbol_match TRANSL). reflexivity. +Qed. + +Lemma senv_preserved: + Senv.equiv pge tpge. +Proof. + eapply (Genv.senv_match TRANSL). +Qed. + +(* TODO gourdinl move this to BTL_Scheduler.v? *) +Inductive match_fundef: fundef -> fundef -> Prop := + | match_Internal f f': match_function f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: stackframe -> stackframe -> Prop := + | match_stackframe_intro + res f sp pc rs f' + (TRANSF: match_function f f') + : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs). + +Inductive match_states: state -> state -> Prop := + | match_states_intro + st f sp pc rs m st' f' + (*(ENTRY: (fn_code f) ! pc = Some ibf)*) + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function f f') + : match_states (State st f sp pc rs m) (State st' f' sp pc rs m) + | match_states_call + st st' f f' args m + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_fundef f f') + : match_states (Callstate st f args m) (Callstate st' f' args m) + | match_states_return + st st' v m + (STACKS: list_forall2 match_stackframes st st') + : match_states (Returnstate st v m) (Returnstate st' v m) + . + +Lemma transf_function_correct f f': + transf_function f = OK f' -> match_function f f'. +Proof. + unfold transf_function. intros H. monadInv H. + econstructor; eauto. eapply check_symbolic_simu_correct; eauto. +Qed. + +Lemma transf_fundef_correct f f': + transf_fundef f = OK f' -> match_fundef f f'. +Proof. +Admitted. + +Lemma function_ptr_preserved: + forall v f, + Genv.find_funct_ptr pge v = Some f -> + exists tf, + Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf. +Proof. + intros. + exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. +Qed. + +Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = funsig f. +Proof. + intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto. + symmetry; erewrite preserv_fnsig; eauto. +Qed. + +Theorem transf_initial_states: + forall s1, initial_state prog s1 -> + exists s2, initial_state tprog s2 /\ match_states s1 s2. +Proof. + intros. inv H. + exploit function_ptr_preserved; eauto. intros (tf & FIND & TRANSF). + eexists. split. + - econstructor; eauto. + + intros; apply (Genv.init_mem_match TRANSL); eauto. + + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. + symmetry. eapply match_program_main. eauto. + + erewrite function_sig_translated; eauto. + - constructor; eauto. + + constructor. + + apply transf_fundef_correct; auto. +Qed. + +Lemma transf_final_states s1 s2 r: + match_states s1 s2 -> final_state s1 r -> final_state s2 r. +Proof. + intros. inv H0. inv H. inv STACKS. constructor. +Qed. + +Lemma iblock_step_simulation stk1 stk2 f f' sp rs m ibf t s1 pc + (STEP : iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s1) + (ENTRY : (fn_code f) ! pc = Some ibf) + (STACKS : list_forall2 match_stackframes stk1 stk2) + (TRANSF : match_function f f') + :exists s2 : state, + (fn_code f') ! pc = Some ibf /\ + iblock_step tr_inputs tpge stk2 f' sp rs m ibf.(entry) t s2 /\ + match_states s1 s2. +Admitted. + +Local Hint Constructors step: core. + +Theorem step_simulation s1 s1' t s2 + (STEP: step tr_inputs pge s1 t s1') + (MS: match_states s1 s2) + :exists s2', + step tr_inputs tpge s2 t s2' + /\ match_states s1' s2'. +Proof. + destruct STEP as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; inv MS. + - (* iblock *) + simplify_someHyps. intros ENTRY. + exploit iblock_step_simulation; eauto. + intros (s2 & ENTRY2 & STEP2 & MS2). + eexists; split; eauto. + - (* function internal *) + inversion TRANSF as [xf tf MF |]; subst. + eexists; split. + + econstructor. erewrite <- preserv_fnstacksize; eauto. + + erewrite preserv_fnparams; eauto. + erewrite preserv_entrypoint; eauto. + econstructor; eauto. + - (* function external *) + inv TRANSF. eexists; split; econstructor; eauto. + eapply external_call_symbols_preserved; eauto. apply senv_preserved. + - (* return *) + inv STACKS. destruct b1 as [res' f' sp' pc' rs']. + eexists; split. + + econstructor. + + inv H1. econstructor; eauto. +Qed. Theorem transf_program_correct: forward_simulation (fsem prog) (fsem tprog). Proof. - eapply forward_simulation_step with match_states. (* lock-step with respect to match_states *) -Admitted. + eapply forward_simulation_step with match_states; simpl; eauto. (* lock-step with respect to match_states *) + - eapply senv_preserved. + - eapply transf_initial_states. + - eapply transf_final_states. + - intros; eapply step_simulation; eauto. +Qed. End PRESERVATION. -- cgit From 26d9fbb36b2e9c8f1262250035cd804bf87f7228 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 14 Jun 2021 19:02:39 +0200 Subject: Preparation for scheduling proof, main lemmas ok --- scheduling/BTL_Livecheck.v | 1 - scheduling/BTL_Scheduler.v | 1 - scheduling/BTL_Schedulerproof.v | 53 +++++++++++++++++++++++++++++++---------- scheduling/RTLtoBTLproof.v | 2 +- 4 files changed, 41 insertions(+), 16 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 25a1c545..9f96e74e 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -583,7 +583,6 @@ Lemma cfgsem2fsem_ibistep_simu_Some sp f stk1 stk2 ib: forall s t rs1 m rs1' m' alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2) (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2)) (LIVE: liveness_ok_function f) - (*(PC : (fn_code f) ! pc = Some ibf)*) (STACKS: list_forall2 eqlive_stackframes stk1 stk2), exists rs2' s', iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' (Some fin)) diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 8240c861..ec83b3c1 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -37,7 +37,6 @@ 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; - (*exists_entrypoint: exists ibf, (fn_code f1)!(fn_entrypoint f1) = Some ibf;*) (* 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); diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index f186581f..aaa2bf80 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -46,11 +46,11 @@ Inductive match_stackframes: stackframe -> stackframe -> Prop := Inductive match_states: state -> state -> Prop := | match_states_intro - st f sp pc rs m st' f' - (*(ENTRY: (fn_code f) ! pc = Some ibf)*) + st f sp pc rs rs' m st' f' + (EQREGS: forall r, rs # r = rs' # r) (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function f f') - : match_states (State st f sp pc rs m) (State st' f' sp pc rs m) + : match_states (State st f sp pc rs m) (State st' f' sp pc rs' m) | match_states_call st st' f f' args m (STACKS: list_forall2 match_stackframes st st') @@ -113,17 +113,44 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -Lemma iblock_step_simulation stk1 stk2 f f' sp rs m ibf t s1 pc - (STEP : iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s1) - (ENTRY : (fn_code f) ! pc = Some ibf) +(* TODO gourdinl + generalize with a None version of the lemma *) +Lemma iblock_istep_simulation stk1 stk2 f f' sp pc fin rs1 m rs1' m' rs2 ib1 t s + (ISTEP: iblock_istep pge sp rs1 m ib1 rs1' m' (Some fin)) + (FSTEP : final_step tr_inputs pge stk1 f sp rs1' m' fin t s) + (EQREGS: forall r : positive, rs1 # r = rs2 # r) (STACKS : list_forall2 match_stackframes stk1 stk2) (TRANSF : match_function f f') - :exists s2 : state, - (fn_code f') ! pc = Some ibf /\ - iblock_step tr_inputs tpge stk2 f' sp rs m ibf.(entry) t s2 /\ - match_states s1 s2. + :exists ib2 rs2' s', + (fn_code f') ! pc = Some ib2 + /\ iblock_istep tpge sp rs2 m ib2.(entry) rs2' m' (Some fin) + /\ final_step tr_inputs tpge stk2 f' sp rs2' m' fin t s' + /\ match_states s s'. +Proof. + induction ib1; inv ISTEP; eauto. + - (* BF *) + admit. + - (* Bseq continue *) + eauto. Admitted. +Lemma iblock_step_simulation stk1 stk2 f f' sp rs rs' m ibf t s1 pc + (STEP: iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s1) + (PC: (fn_code f) ! pc = Some ibf) + (EQREGS: forall r : positive, rs # r = rs' # r) + (STACKS: list_forall2 match_stackframes stk1 stk2) + (TRANSF: match_function f f') + :exists ibf' s2, + (fn_code f') ! pc = Some ibf' /\ + iblock_step tr_inputs tpge stk2 f' sp rs' m ibf'.(entry) t s2 /\ + match_states s1 s2. +Proof. + destruct STEP as (rs1 & m1 & fin & ISTEP & FSTEP). + exploit iblock_istep_simulation; eauto. + intros (ibf2 & rs2' & s' & PC' & ISTEP' & FSTEP' & MS'). + repeat eexists; eauto. +Qed. + Local Hint Constructors step: core. Theorem step_simulation s1 s1' t s2 @@ -133,11 +160,11 @@ Theorem step_simulation s1 s1' t s2 step tr_inputs tpge s2 t s2' /\ match_states s1' s2'. Proof. - destruct STEP as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; inv MS. + destruct STEP as [stack ibf f sp n rs m t s PC STEP | | | ]; inv MS. - (* iblock *) - simplify_someHyps. intros ENTRY. + simplify_someHyps. intros PC. exploit iblock_step_simulation; eauto. - intros (s2 & ENTRY2 & STEP2 & MS2). + intros (ibf' & s2 & PC2 & STEP2 & MS2). eexists; split; eauto. - (* function internal *) inversion TRANSF as [xf tf MF |]; subst. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 18ff8d5f..513ed40a 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -375,7 +375,7 @@ Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = RTL.f Proof. intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto. erewrite preserv_fnsig; eauto. -Admitted. +Qed. Lemma transf_initial_states s1: RTL.initial_state prog s1 -> -- cgit From 870d0cf9d06f453e2ba3cbdaf0184bc1f657c04b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 15 Jun 2021 17:11:25 +0200 Subject: some advance in sched proof --- scheduling/BTL_Schedulerproof.v | 116 +++++++++++++++++++++++++++++++--------- 1 file changed, 92 insertions(+), 24 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index aaa2bf80..c7ecb0dc 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -40,9 +40,10 @@ Inductive match_fundef: fundef -> fundef -> Prop := Inductive match_stackframes: stackframe -> stackframe -> Prop := | match_stackframe_intro - res f sp pc rs f' + res f sp pc rs rs' f' + (EQREGS: forall r, rs # r = rs' # r) (TRANSF: match_function f f') - : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs). + : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs'). Inductive match_states: state -> state -> Prop := | match_states_intro @@ -62,6 +63,24 @@ Inductive match_states: state -> state -> Prop := : match_states (Returnstate st v m) (Returnstate st' v m) . +Lemma match_stack_equiv stk1 stk2: + list_forall2 match_stackframes stk1 stk2 -> + forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> + list_forall2 match_stackframes stk1 stk3. +Proof. + (*Local Hint Resolve match_stackframes_equiv: core.*) + induction 1; intros stk3 EQ; inv EQ; constructor; eauto. + inv H3; inv H; econstructor; eauto. + intros; rewrite <- EQUIV; auto. +Qed. + +Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3. +Proof. + Local Hint Resolve match_stack_equiv: core. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; rewrite <- EQUIV; auto. +Qed. + Lemma transf_function_correct f f': transf_function f = OK f' -> match_function f f'. Proof. @@ -113,43 +132,89 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -(* TODO gourdinl - generalize with a None version of the lemma *) -Lemma iblock_istep_simulation stk1 stk2 f f' sp pc fin rs1 m rs1' m' rs2 ib1 t s +Theorem symbolic_simu_ssem stk1 stk2 f1 f2 ib1 ib2: + forall (ctx: simu_proof_context f1 f2) t s, + sem_sstate (bctx1 ctx) stk1 t s (sexec f1 ib1) -> + exists s', + sem_sstate (bctx2 ctx) stk2 t s' (sexec f2 ib2) + /\ match_states s s'. +Admitted. + +Theorem symbolic_simu_correct stk1 stk2 f1 f2 ib1 ib2: + symbolic_simu f1 f2 ib1 ib2 -> + forall (ctx: simu_proof_context f1 f2) t s, iblock_step tr_inputs (sge1 ctx) stk1 f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s -> + exists s', + iblock_step tr_inputs (sge2 ctx) stk2 f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s' + /\ match_states s s'. +Proof. + unfold symbolic_simu, sstate_simu. + intros SIMU ctx t s1 STEP1. + exploit (sexec_correct (bctx1 ctx)); simpl; eauto. + intros SST1. exploit sem_sstate_run; eauto. + intros (sis1 & sfv1 & rs & m & OUTC1 & SIS1 & SV1). + exploit sem_si_ok; eauto. intros SOK. + intros; exploit SIMU; eauto. + intros (sis2 & sfv2 & OUTC2 & SSIMU & SEM). + exploit symbolic_simu_ssem; eauto. + intros (s2 & SST2 & MS). + exploit (sexec_exact (bctx2 ctx)); simpl; eauto. + intros (s3 & STEP & EQUIV). + exists s3. intuition eauto. + eapply match_states_equiv; eauto. +Qed. + +(* +Lemma iblock_istep_simulation_None stk1 stk2 f f' sp pc rs1 m rs1' m' rs2 ib1 ibf2 s + (ISTEP: iblock_istep pge sp rs1 m ib1 rs1' m' None) + (PC: (fn_code f') ! pc = Some ibf2) + (EQREGS: forall r : positive, rs1 # r = rs2 # r) + (STACKS: list_forall2 match_stackframes stk1 stk2) + (TRANSF: match_function f f') + :exists rs2' s', + iblock_istep_run tpge sp ibf2.(entry) rs2 m = Some (out rs2' m' None) + /\ match_states s s'. +Proof. + induction ib1; inv ISTEP; eauto. +Admitted. + +Lemma iblock_istep_simulation_Some stk1 stk2 f f' sp ib1: forall fin ibf2 rs1 m rs1' m' rs2 pc t s (ISTEP: iblock_istep pge sp rs1 m ib1 rs1' m' (Some fin)) - (FSTEP : final_step tr_inputs pge stk1 f sp rs1' m' fin t s) + (FSTEP: final_step tr_inputs pge stk1 f sp rs1' m' fin t s) + (PC: (fn_code f') ! pc = Some ibf2) (EQREGS: forall r : positive, rs1 # r = rs2 # r) - (STACKS : list_forall2 match_stackframes stk1 stk2) - (TRANSF : match_function f f') - :exists ib2 rs2' s', - (fn_code f') ! pc = Some ib2 - /\ iblock_istep tpge sp rs2 m ib2.(entry) rs2' m' (Some fin) + (STACKS: list_forall2 match_stackframes stk1 stk2) + (TRANSF: match_function f f'), + exists rs2' s', + iblock_istep_run tpge sp ibf2.(entry) rs2 m = Some (out rs2' m' (Some fin)) /\ final_step tr_inputs tpge stk2 f' sp rs2' m' fin t s' /\ match_states s s'. Proof. - induction ib1; inv ISTEP; eauto. + induction ib1; simpl; try_simplify_someHyps; inv ISTEP; eauto; intros. - (* BF *) admit. - (* Bseq continue *) - eauto. + exploit iblock_istep_simulation_None; eauto. + intros (rs2' & s' & ISTEP & MS). + admit. + - Admitted. - -Lemma iblock_step_simulation stk1 stk2 f f' sp rs rs' m ibf t s1 pc - (STEP: iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s1) + *) +Lemma iblock_step_simulation stk1 stk2 f f' sp rs rs' m ibf t s pc + (STEP: iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s) (PC: (fn_code f) ! pc = Some ibf) (EQREGS: forall r : positive, rs # r = rs' # r) (STACKS: list_forall2 match_stackframes stk1 stk2) (TRANSF: match_function f f') - :exists ibf' s2, - (fn_code f') ! pc = Some ibf' /\ - iblock_step tr_inputs tpge stk2 f' sp rs' m ibf'.(entry) t s2 /\ - match_states s1 s2. + :exists ibf' s', + (fn_code f') ! pc = Some ibf' + /\ iblock_step tr_inputs tpge stk2 f' sp rs' m ibf'.(entry) t s' + /\ match_states s s'. Proof. destruct STEP as (rs1 & m1 & fin & ISTEP & FSTEP). - exploit iblock_istep_simulation; eauto. - intros (ibf2 & rs2' & s' & PC' & ISTEP' & FSTEP' & MS'). - repeat eexists; eauto. -Qed. + inversion TRANSF. apply symbolic_simu_ok in PC as SYMB. + destruct SYMB as (ibf2 & PC' & SYMBS). clear symbolic_simu_ok. + unfold symbolic_simu, sstate_simu in SYMBS. +Admitted. Local Hint Constructors step: core. @@ -181,6 +246,9 @@ Proof. eexists; split. + econstructor. + inv H1. econstructor; eauto. + intros; destruct (Pos.eq_dec res' r); subst. + * rewrite !Regmap.gss; reflexivity. + * rewrite !Regmap.gso; auto. Qed. Theorem transf_program_correct: -- cgit From ef95bd7f4afe159bcedc3ec5732579bfc6ba08c6 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 17 Jun 2021 17:36:42 +0200 Subject: some advance, new section to simplify context from symbolic exec --- scheduling/BTL_Scheduler.v | 12 ++- scheduling/BTL_Schedulerproof.v | 203 ++++++++++++++++++++++++++++------------ scheduling/RTLtoBTLproof.v | 2 +- 3 files changed, 151 insertions(+), 66 deletions(-) (limited to 'scheduling') 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 -> diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index c7ecb0dc..85b185c5 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -33,6 +33,18 @@ Proof. eapply (Genv.senv_match TRANSL). Qed. +Lemma functions_preserved: + forall (v: val) (f: fundef), + Genv.find_funct pge v = Some f -> + exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tpge v = Some tf /\ linkorder cunit prog. +Proof. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & A & B & C). + repeat eexists; intuition eauto. + + unfold incl; auto. + + eapply linkorder_refl. +Qed. + (* TODO gourdinl move this to BTL_Scheduler.v? *) Inductive match_fundef: fundef -> fundef -> Prop := | match_Internal f f': match_function f f' -> match_fundef (Internal f) (Internal f') @@ -85,7 +97,9 @@ Lemma transf_function_correct f f': transf_function f = OK f' -> match_function f f'. Proof. unfold transf_function. intros H. monadInv H. - econstructor; eauto. eapply check_symbolic_simu_correct; eauto. + econstructor; eauto. + eapply check_symbolic_simu_input_equiv; eauto. + eapply check_symbolic_simu_correct; eauto. Qed. Lemma transf_fundef_correct f f': @@ -132,30 +146,122 @@ Proof. intros. inv H0. inv H. inv STACKS. constructor. Qed. -Theorem symbolic_simu_ssem stk1 stk2 f1 f2 ib1 ib2: - forall (ctx: simu_proof_context f1 f2) t s, - sem_sstate (bctx1 ctx) stk1 t s (sexec f1 ib1) -> +Section SYMBOLIC_CTX. + +Variables f1 f2: function. +Variable sp0: val. +Variable rs0: regset. +Variable m0: Memory.Mem.mem. + +Lemma symbols_preserved_rev s: Genv.find_symbol pge s = Genv.find_symbol tpge s. +Proof. + symmetry; apply symbols_preserved. +Qed. + +Let ctx := Sctx f1 f2 pge tpge symbols_preserved_rev sp0 rs0 m0. + +Lemma str_regs_equiv sfv1 sfv2 rs: + sfv_simu ctx sfv1 sfv2 -> + match_function f1 f2 -> + str_regs (cf (bctx1 ctx)) sfv1 rs = str_regs (cf (bctx2 ctx)) sfv2 rs. +Proof. + intros SFV MF; inv MF. + inv SFV; simpl; + erewrite equiv_input_regs_tr_inputs; auto. +Qed. + +Lemma sfind_function_preserved ros1 ros2 fd: + svident_simu ctx ros1 ros2 -> + sfind_function (bctx1 ctx) ros1 = Some fd -> + exists fd', + sfind_function (bctx2 ctx) ros2 = Some fd' + /\ transf_fundef fd = OK fd'. +Proof. + unfold sfind_function. intros [sv1 sv2 SIMU|]; simpl in *. + + rewrite !SIMU. clear SIMU. + destruct (eval_sval _ _); try congruence. + intros; exploit functions_preserved; eauto. + intros (fd' & cunit & (X1 & X2 & X3)). eexists; eauto. + + subst. rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. + intros; exploit function_ptr_preserved; eauto. +Qed. + +Theorem symbolic_simu_ssem stk1 stk2 st st' s t: + sstate_simu ctx st st' -> + sem_sstate (bctx1 ctx) stk1 t s st -> + list_forall2 match_stackframes stk1 stk2 -> + match_function f1 f2 -> exists s', - sem_sstate (bctx2 ctx) stk2 t s' (sexec f2 ib2) + sem_sstate (bctx2 ctx) stk2 t s' st' /\ match_states s s'. +Proof. + intros SIMU SST1 STACKS MF. inversion MF. + (*destruct SST1.*) + (*- [> sem_Sfinal <]*) + (*admit.*) + (*- [> sem_Scond <]*) + + exploit sem_sstate_run; eauto. + intros (sis1 & sfv1 & rs' & m' & SOUT1 & SIS1 & SF). + exploit sem_si_ok; eauto. intros SOK. + exploit SIMU; simpl; eauto. (* rewrite SEVAL; eauto. *) + intros (sis2 & sfv2 & SOUT2 & SSIMU & SFVSIM). + remember SIS1 as EQSFV. clear HeqEQSFV. + eapply SFVSIM in EQSFV. + + + exploit SSIMU; eauto. + intros SIS2. + + + + destruct SIS1 as (PRE1 & SMEM1 & SREGS1). + try_simplify_someHyps; intros. + (*exploit sem_sfval_equiv; eauto.*) + (*intros (s' & SF' & EQUIV).*) + set (EQSFV2:=sfv2). + inversion EQSFV; subst. + - + inversion SF; subst. + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; simpl; eauto. + + econstructor. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + erewrite equiv_input_regs_tr_inputs; eauto. + + - + inversion SF; subst. + exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND). + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; eauto. + + econstructor; eauto. + * apply function_sig_translated; auto. + * erewrite <- ARGS; eauto. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + * erewrite equiv_input_regs_tr_inputs; eauto. + repeat (econstructor; eauto). + * apply transf_fundef_correct; auto. + - + + +admit. Admitted. -Theorem symbolic_simu_correct stk1 stk2 f1 f2 ib1 ib2: - symbolic_simu f1 f2 ib1 ib2 -> - forall (ctx: simu_proof_context f1 f2) t s, iblock_step tr_inputs (sge1 ctx) stk1 f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s -> +Theorem symbolic_simu_correct stk1 stk2 ibf1 ibf2 s t: + sstate_simu ctx (sexec f1 ibf1.(entry)) (sexec f2 ibf2.(entry)) -> + iblock_step tr_inputs (sge1 ctx) stk1 f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ibf1.(entry) t s -> + list_forall2 match_stackframes stk1 stk2 -> + match_function f1 f2 -> exists s', - iblock_step tr_inputs (sge2 ctx) stk2 f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s' + iblock_step tr_inputs (sge2 ctx) stk2 f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ibf2.(entry) t s' /\ match_states s s'. Proof. - unfold symbolic_simu, sstate_simu. - intros SIMU ctx t s1 STEP1. + unfold sstate_simu. + intros SIMU STEP1 STACKS MF. exploit (sexec_correct (bctx1 ctx)); simpl; eauto. - intros SST1. exploit sem_sstate_run; eauto. - intros (sis1 & sfv1 & rs & m & OUTC1 & SIS1 & SV1). - exploit sem_si_ok; eauto. intros SOK. - intros; exploit SIMU; eauto. - intros (sis2 & sfv2 & OUTC2 & SSIMU & SEM). - exploit symbolic_simu_ssem; eauto. + intros SST1. exploit symbolic_simu_ssem; eauto. intros (s2 & SST2 & MS). exploit (sexec_exact (bctx2 ctx)); simpl; eauto. intros (s3 & STEP & EQUIV). @@ -163,57 +269,34 @@ Proof. eapply match_states_equiv; eauto. Qed. -(* -Lemma iblock_istep_simulation_None stk1 stk2 f f' sp pc rs1 m rs1' m' rs2 ib1 ibf2 s - (ISTEP: iblock_istep pge sp rs1 m ib1 rs1' m' None) - (PC: (fn_code f') ! pc = Some ibf2) - (EQREGS: forall r : positive, rs1 # r = rs2 # r) - (STACKS: list_forall2 match_stackframes stk1 stk2) - (TRANSF: match_function f f') - :exists rs2' s', - iblock_istep_run tpge sp ibf2.(entry) rs2 m = Some (out rs2' m' None) - /\ match_states s s'. +Lemma iblock_step_eqregs tr ge stk f sp ib: + forall rs m rs' t s, + (forall r, rs # r = rs' # r) -> + iblock_step tr ge stk f sp rs m ib t s <-> + iblock_step tr ge stk f sp rs' m ib t s. Proof. - induction ib1; inv ISTEP; eauto. + induction ib; try_simplify_someHyps. Admitted. -Lemma iblock_istep_simulation_Some stk1 stk2 f f' sp ib1: forall fin ibf2 rs1 m rs1' m' rs2 pc t s - (ISTEP: iblock_istep pge sp rs1 m ib1 rs1' m' (Some fin)) - (FSTEP: final_step tr_inputs pge stk1 f sp rs1' m' fin t s) - (PC: (fn_code f') ! pc = Some ibf2) - (EQREGS: forall r : positive, rs1 # r = rs2 # r) - (STACKS: list_forall2 match_stackframes stk1 stk2) - (TRANSF: match_function f f'), - exists rs2' s', - iblock_istep_run tpge sp ibf2.(entry) rs2 m = Some (out rs2' m' (Some fin)) - /\ final_step tr_inputs tpge stk2 f' sp rs2' m' fin t s' - /\ match_states s s'. -Proof. - induction ib1; simpl; try_simplify_someHyps; inv ISTEP; eauto; intros. - - (* BF *) - admit. - - (* Bseq continue *) - exploit iblock_istep_simulation_None; eauto. - intros (rs2' & s' & ISTEP & MS). - admit. - - -Admitted. - *) -Lemma iblock_step_simulation stk1 stk2 f f' sp rs rs' m ibf t s pc - (STEP: iblock_step tr_inputs pge stk1 f sp rs m ibf.(entry) t s) - (PC: (fn_code f) ! pc = Some ibf) +End SYMBOLIC_CTX. + +Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s pc + (STEP: iblock_step tr_inputs pge stk1 f1 sp rs m ibf.(entry) t s) + (PC: (fn_code f1) ! pc = Some ibf) (EQREGS: forall r : positive, rs # r = rs' # r) (STACKS: list_forall2 match_stackframes stk1 stk2) - (TRANSF: match_function f f') + (TRANSF: match_function f1 f2) :exists ibf' s', - (fn_code f') ! pc = Some ibf' - /\ iblock_step tr_inputs tpge stk2 f' sp rs' m ibf'.(entry) t s' + (fn_code f2) ! pc = Some ibf' + /\ iblock_step tr_inputs tpge stk2 f2 sp rs' m ibf'.(entry) t s' /\ match_states s s'. Proof. - destruct STEP as (rs1 & m1 & fin & ISTEP & FSTEP). - inversion TRANSF. apply symbolic_simu_ok in PC as SYMB. - destruct SYMB as (ibf2 & PC' & SYMBS). clear symbolic_simu_ok. - unfold symbolic_simu, sstate_simu in SYMBS. + (*destruct STEP as (rs1 & m1 & fin & ISTEP & FSTEP).*) + exploit symbolic_simu_ok; eauto. intros (ib2 & PC' & SYMBS). + exploit symbolic_simu_correct; repeat (eauto; simpl). + intros (s' & STEP2 & MS). + exists ib2; exists s'; repeat split; auto. + erewrite iblock_step_eqregs; eauto. Admitted. Local Hint Constructors step: core. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 513ed40a..f5c2f3db 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -743,7 +743,7 @@ Theorem transf_program_correct: Proof. eapply compose_forward_simulations. - eapply transf_program_correct_cfg. - - eapply cfgsem2fsem. + - eapply cfgsem2fsem. eauto. Admitted. End BTL_SIMULATES_RTL. -- cgit From 86e108c7b95456e986605cb1861cb5128f348ec0 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 18 Jun 2021 11:08:52 +0200 Subject: End of main scheduling proof --- scheduling/BTL_Scheduler.v | 1 + scheduling/BTL_Schedulerproof.v | 139 ++++++++++++++++++++++++---------------- scheduling/RTLtoBTLproof.v | 9 ++- 3 files changed, 90 insertions(+), 59 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 10f780e2..c50eb364 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -52,6 +52,7 @@ Definition check_symbolic_simu (f tf: BTL.function): res unit := OK tt. (* TODO: Lemma check_symbolic_simu_input_equiv x f1 f2: check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. +Proof. Admitted. Lemma check_symbolic_simu_correct x f1 f2: diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 85b185c5..6963ee8b 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -105,6 +105,11 @@ Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. + intros TRANSF; destruct f; simpl. monadInv TRANSF. + + exploit transf_function_correct; eauto. + (*intros (dupmap & MATCH_F).*) + (*eapply match_Internal; eauto.*) + (*+ eapply match_External.*) Admitted. Lemma function_ptr_preserved: @@ -196,58 +201,79 @@ Theorem symbolic_simu_ssem stk1 stk2 st st' s t: /\ match_states s s'. Proof. intros SIMU SST1 STACKS MF. inversion MF. - (*destruct SST1.*) - (*- [> sem_Sfinal <]*) - (*admit.*) - (*- [> sem_Scond <]*) - exploit sem_sstate_run; eauto. intros (sis1 & sfv1 & rs' & m' & SOUT1 & SIS1 & SF). exploit sem_si_ok; eauto. intros SOK. - exploit SIMU; simpl; eauto. (* rewrite SEVAL; eauto. *) + exploit SIMU; simpl; eauto. intros (sis2 & sfv2 & SOUT2 & SSIMU & SFVSIM). remember SIS1 as EQSFV. clear HeqEQSFV. eapply SFVSIM in EQSFV. - - - exploit SSIMU; eauto. - intros SIS2. - - - + exploit SSIMU; eauto. intros SIS2. destruct SIS1 as (PRE1 & SMEM1 & SREGS1). try_simplify_someHyps; intros. - (*exploit sem_sfval_equiv; eauto.*) - (*intros (s' & SF' & EQUIV).*) set (EQSFV2:=sfv2). - inversion EQSFV; subst. - - - inversion SF; subst. - exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. - + erewrite <- str_regs_equiv; simpl; eauto. - + econstructor. - + intros SST2. eexists; split; eauto. - econstructor; simpl; eauto. - erewrite equiv_input_regs_tr_inputs; eauto. - - - - inversion SF; subst. - exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND). - exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. - + erewrite <- str_regs_equiv; eauto. - + econstructor; eauto. - * apply function_sig_translated; auto. - * erewrite <- ARGS; eauto. - + intros SST2. eexists; split; eauto. - econstructor; simpl; eauto. - * erewrite equiv_input_regs_tr_inputs; eauto. - repeat (econstructor; eauto). - * apply transf_fundef_correct; auto. - - - - -admit. -Admitted. + inversion EQSFV; subst; inversion SF; subst. + - (* goto *) + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; simpl; eauto. + + econstructor. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + erewrite equiv_input_regs_tr_inputs; eauto. + - (* call *) + exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND). + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; eauto. + + econstructor; eauto. + * apply function_sig_translated; auto. + * erewrite <- ARGS; eauto. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + * erewrite equiv_input_regs_tr_inputs; eauto. + repeat (econstructor; eauto). + * apply transf_fundef_correct; auto. + - (* tailcall *) + exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND). + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + econstructor; eauto. + * apply function_sig_translated; auto. + * simpl; erewrite <- preserv_fnstacksize; eauto. + * erewrite <- ARGS; eauto. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + apply transf_fundef_correct; auto. + - (* builtin *) + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; simpl; eauto. + + econstructor. + * eapply eval_builtin_sargs_preserved. + eapply senv_preserved. + eapply eval_list_builtin_sval_correct; eauto. + inv BARGS. + erewrite <- eval_list_builtin_sval_preserved in H0; auto. + * simpl in *; eapply external_call_symbols_preserved; eauto. + eapply senv_preserved. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + erewrite equiv_input_regs_tr_inputs; eauto. + - (* jumptable *) + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + erewrite <- str_regs_equiv; simpl; eauto. + + econstructor; eauto. + rewrite <- VAL; auto. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + erewrite equiv_input_regs_tr_inputs; eauto. + - (* return *) + exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto. + + econstructor; eauto. + * simpl; erewrite <- preserv_fnstacksize; eauto. + * inv OPT; eauto. + rewrite <- SIMU0; auto. + + intros SST2. eexists; split; eauto. + econstructor; simpl; eauto. + Unshelve. all: auto. +Qed. Theorem symbolic_simu_correct stk1 stk2 ibf1 ibf2 s t: sstate_simu ctx (sexec f1 ibf1.(entry)) (sexec f2 ibf2.(entry)) -> @@ -269,17 +295,21 @@ Proof. eapply match_states_equiv; eauto. Qed. -Lemma iblock_step_eqregs tr ge stk f sp ib: - forall rs m rs' t s, - (forall r, rs # r = rs' # r) -> - iblock_step tr ge stk f sp rs m ib t s <-> - iblock_step tr ge stk f sp rs' m ib t s. +End SYMBOLIC_CTX. + +Lemma iblock_step_eqregs sp ib: + forall rs rs' + (EQREGS: forall r, rs # r = rs' # r) + fin m rs2 m2 (ISTEP: iblock_istep tpge sp rs m ib rs2 m2 (Some fin)), + iblock_istep tpge sp rs' m ib rs2 m2 (Some fin). Proof. - induction ib; try_simplify_someHyps. + induction ib; intros; inv ISTEP. + - + admit. + - econstructor; eauto. + - econstructor. eapply IHib1; eauto. Admitted. -End SYMBOLIC_CTX. - Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s pc (STEP: iblock_step tr_inputs pge stk1 f1 sp rs m ibf.(entry) t s) (PC: (fn_code f1) ! pc = Some ibf) @@ -291,12 +321,13 @@ Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s pc /\ iblock_step tr_inputs tpge stk2 f2 sp rs' m ibf'.(entry) t s' /\ match_states s s'. Proof. - (*destruct STEP as (rs1 & m1 & fin & ISTEP & FSTEP).*) exploit symbolic_simu_ok; eauto. intros (ib2 & PC' & SYMBS). exploit symbolic_simu_correct; repeat (eauto; simpl). intros (s' & STEP2 & MS). - exists ib2; exists s'; repeat split; auto. - erewrite iblock_step_eqregs; eauto. + exists ib2; exists s'. repeat split; auto. + destruct STEP2 as (rs2 & m2 & fin & ISTEP2 & FSTEP2). + unfold iblock_step. exists rs2; exists m2; exists fin. + split; auto. Admitted. Local Hint Constructors step: core. diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index f5c2f3db..4af5668c 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -252,11 +252,10 @@ Lemma transf_function_correct f f': transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. Proof. unfold transf_function; unfold bind. repeat autodestruct. - intros H _ _ X. inversion X; subst; clear X. -(* - eexists; eapply verify_function_correct; simpl; eauto. -*) -Admitted. + intros LIVE VER _ _ X. inv X. eexists; econstructor. + - eapply verify_function_correct; simpl; eauto. + - unfold liveness_ok_function; destruct u0; auto. +Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. -- cgit From 0321d5f0aa2478ab830990a3637c37566bd9b634 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 18 Jun 2021 11:46:10 +0200 Subject: finish RTLtoBTL proof --- scheduling/RTLtoBTLproof.v | 20 ++++++++++++++++++-- 1 file changed, 18 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 4af5668c..8a352434 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -737,12 +737,28 @@ Proof. - eapply senv_preserved. Qed. +Theorem all_fundef_liveness_ok b f: + Genv.find_funct_ptr tge b = Some f -> liveness_ok_fundef f. +Proof. + unfold match_prog, match_program in TRANSL. + unfold Genv.find_funct_ptr, tge; simpl; intro X. + destruct (Genv.find_def_match_2 TRANSL b) as [|f0 y H]; try congruence. + destruct y as [tf0|]; try congruence. + inversion X as [H1]. subst. clear X. + remember (@Gfun fundef unit f) as f2. + destruct H as [ctx' f1 f2 H0|]; try congruence. + inversion Heqf2 as [H2]. subst; clear Heqf2. + exploit transf_fundef_correct; eauto. + destruct f; econstructor. + inv H1; eapply liveness_ok; eauto. +Qed. + Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (BTL.fsem tprog). Proof. eapply compose_forward_simulations. - eapply transf_program_correct_cfg. - - eapply cfgsem2fsem. eauto. -Admitted. + - eapply cfgsem2fsem. apply all_fundef_liveness_ok. +Qed. End BTL_SIMULATES_RTL. -- cgit From 58312db34d43803a2cb63e20667a5059988dc0d1 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 18 Jun 2021 15:16:21 +0200 Subject: lemma on eqregs? --- scheduling/BTL_Scheduler.v | 2 +- scheduling/BTL_Schedulerproof.v | 19 ++++++------------- 2 files changed, 7 insertions(+), 14 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index c50eb364..12840f62 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -41,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; 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); }. diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 6963ee8b..1a4c5757 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -80,7 +80,6 @@ Lemma match_stack_equiv stk1 stk2: forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> list_forall2 match_stackframes stk1 stk3. Proof. - (*Local Hint Resolve match_stackframes_equiv: core.*) induction 1; intros stk3 EQ; inv EQ; constructor; eauto. inv H3; inv H; econstructor; eauto. intros; rewrite <- EQUIV; auto. @@ -297,17 +296,13 @@ Qed. End SYMBOLIC_CTX. -Lemma iblock_step_eqregs sp ib: - forall rs rs' +Lemma iblock_step_eqregs stk sp f ib: + forall rs rs' m s t (EQREGS: forall r, rs # r = rs' # r) - fin m rs2 m2 (ISTEP: iblock_istep tpge sp rs m ib rs2 m2 (Some fin)), - iblock_istep tpge sp rs' m ib rs2 m2 (Some fin). + (STEP: iblock_step tr_inputs tpge stk f sp rs m ib s t), + iblock_step tr_inputs tpge stk f sp rs' m ib s t. Proof. - induction ib; intros; inv ISTEP. - - - admit. - - econstructor; eauto. - - econstructor. eapply IHib1; eauto. + Admitted. Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s pc @@ -325,9 +320,7 @@ Proof. exploit symbolic_simu_correct; repeat (eauto; simpl). intros (s' & STEP2 & MS). exists ib2; exists s'. repeat split; auto. - destruct STEP2 as (rs2 & m2 & fin & ISTEP2 & FSTEP2). - unfold iblock_step. exists rs2; exists m2; exists fin. - split; auto. + eapply iblock_step_eqregs; eauto. Admitted. Local Hint Constructors step: core. -- cgit From 274f32429e43efa5b4031d2dfe1c32f30a7fff3f Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sat, 19 Jun 2021 23:51:34 +0200 Subject: finish main proof --- scheduling/BTL_Schedulerproof.v | 146 +++++++++++++++++++++++++++++++++++----- 1 file changed, 128 insertions(+), 18 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 1a4c5757..13ea61a6 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -296,32 +296,145 @@ Qed. End SYMBOLIC_CTX. -Lemma iblock_step_eqregs stk sp f ib: - forall rs rs' m s t +Lemma tr_inputs_eqregs_list f tbl: forall rs rs' r' ores, + (forall r : positive, rs # r = rs' # r) -> + (tr_inputs f tbl ores rs) # r' = + (tr_inputs f tbl ores rs') # r'. +Proof. + induction tbl as [|pc tbl IHtbl]; + intros; destruct ores; simpl; + rewrite !tr_inputs_get; autodestruct. +Qed. + +Lemma find_function_eqregs rs rs' ros fd: + (forall r : positive, rs # r = rs' # r) -> + find_function tpge ros rs = Some fd -> + find_function tpge ros rs' = Some fd. +Proof. + intros EQREGS; destruct ros; simpl. + rewrite EQREGS; auto. + autodestruct. +Qed. + +Lemma args_eqregs (args: list reg): forall (rs rs': regset), + (forall r : positive, rs # r = rs' # r) -> + rs ## args = rs' ## args. +Proof. + induction args; simpl; trivial. + intros rs rs' EQREGS; rewrite EQREGS. + apply f_equal; eauto. +Qed. + +Lemma f_arg_eqregs (rs rs': regset): + (forall r : positive, rs # r = rs' # r) -> + (fun r : positive => rs # r) = (fun r : positive => rs' # r). +Proof. + intros EQREGS; apply functional_extensionality; eauto. +Qed. + +Lemma eqregs_update (rs rs': regset) v dest: + (forall r, rs # r = rs' # r) -> + (forall r, (rs # dest <- v) # r = (rs' # dest <- v) # r). +Proof. + intros EQREGS r; destruct (Pos.eq_dec dest r); subst. + * rewrite !Regmap.gss; reflexivity. + * rewrite !Regmap.gso; auto. +Qed. + +Local Hint Resolve equiv_stack_refl tr_inputs_eqregs_list find_function_eqregs eqregs_update: core. + +Lemma iblock_istep_eqregs_None sp ib: forall rs m rs' rs1 m1 (EQREGS: forall r, rs # r = rs' # r) - (STEP: iblock_step tr_inputs tpge stk f sp rs m ib s t), - iblock_step tr_inputs tpge stk f sp rs' m ib s t. + (ISTEP: iblock_istep tpge sp rs m ib rs1 m1 None), + exists rs1', + iblock_istep tpge sp rs' m ib rs1' m1 None + /\ (forall r, rs1 # r = rs1' # r). Proof. + induction ib; intros; inv ISTEP. + - repeat econstructor; eauto. + - repeat econstructor; eauto. + erewrite args_eqregs; eauto. + - repeat econstructor; eauto. + erewrite args_eqregs; eauto. + - repeat econstructor; eauto. + erewrite args_eqregs; eauto. + rewrite <- EQREGS; eauto. + - exploit IHib1; eauto. + intros (rs1' & ISTEP1 & EQREGS1). + exploit IHib2; eauto. + intros (rs2' & ISTEP2 & EQREGS2). + eexists; split. econstructor; eauto. + eauto. + - destruct b. + 1: exploit IHib1; eauto. + 2: exploit IHib2; eauto. + all: + intros (rs1' & ISTEP & EQREGS'); + eexists; split; try eapply exec_cond; + try erewrite args_eqregs; eauto. +Qed. -Admitted. +Lemma iblock_step_eqregs stk sp f ib: forall rs rs' m t s + (EQREGS: forall r, rs # r = rs' # r) + (STEP: iblock_step tr_inputs tpge stk f sp rs m ib t s), + exists s', + iblock_step tr_inputs tpge stk f sp rs' m ib t s' + /\ equiv_state s s'. +Proof. + induction ib; intros; + destruct STEP as (rs2 & m2 & fin & ISTEP & FSTEP); inv ISTEP. + - inv FSTEP; + eexists; split; repeat econstructor; eauto. + + destruct (or); simpl; try rewrite EQREGS; constructor; eauto. + + erewrite args_eqregs; eauto. repeat econstructor; eauto. + + erewrite args_eqregs; eauto. econstructor; eauto. + + erewrite f_arg_eqregs; eauto. + + destruct res; simpl; eauto. + + rewrite <- EQREGS; auto. + - exploit IHib1; eauto. + + econstructor. do 2 eexists; split; eauto. + + intros (s' & STEP & EQUIV). clear FSTEP. + destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP). + repeat (econstructor; eauto). + - exploit iblock_istep_eqregs_None; eauto. + intros (rs1' & ISTEP1 & EQREGS1). + exploit IHib2; eauto. + + econstructor. do 2 eexists; split; eauto. + + intros (s' & STEP & EQUIV). clear FSTEP. + destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP). + eexists; split; eauto. + econstructor; do 2 eexists; split; eauto. + eapply exec_seq_continue; eauto. + - destruct b. + 1: exploit IHib1; eauto. + 3: exploit IHib2; eauto. + 1,3: econstructor; do 2 eexists; split; eauto. + all: + intros (s' & STEP & EQUIV); clear FSTEP; + destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP); + eexists; split; eauto; econstructor; + do 2 eexists; split; eauto; + eapply exec_cond; try erewrite args_eqregs; eauto. +Qed. -Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s pc - (STEP: iblock_step tr_inputs pge stk1 f1 sp rs m ibf.(entry) t s) +Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s1 pc + (STEP: iblock_step tr_inputs pge stk1 f1 sp rs m ibf.(entry) t s1) (PC: (fn_code f1) ! pc = Some ibf) (EQREGS: forall r : positive, rs # r = rs' # r) (STACKS: list_forall2 match_stackframes stk1 stk2) (TRANSF: match_function f1 f2) - :exists ibf' s', - (fn_code f2) ! pc = Some ibf' - /\ iblock_step tr_inputs tpge stk2 f2 sp rs' m ibf'.(entry) t s' - /\ match_states s s'. + :exists ibf' s2, + (fn_code f2) ! pc = Some ibf' + /\ iblock_step tr_inputs tpge stk2 f2 sp rs' m ibf'.(entry) t s2 + /\ match_states s1 s2. Proof. exploit symbolic_simu_ok; eauto. intros (ib2 & PC' & SYMBS). exploit symbolic_simu_correct; repeat (eauto; simpl). - intros (s' & STEP2 & MS). - exists ib2; exists s'. repeat split; auto. - eapply iblock_step_eqregs; eauto. -Admitted. + intros (s2 & STEP2 & MS). + exploit iblock_step_eqregs; eauto. intros (s3 & STEP3 & EQUIV). + clear STEP2. exists ib2; exists s3. repeat split; auto. + eapply match_states_equiv; eauto. +Qed. Local Hint Constructors step: core. @@ -353,9 +466,6 @@ Proof. eexists; split. + econstructor. + inv H1. econstructor; eauto. - intros; destruct (Pos.eq_dec res' r); subst. - * rewrite !Regmap.gss; reflexivity. - * rewrite !Regmap.gso; auto. Qed. Theorem transf_program_correct: -- cgit From 99b1514f31869bc11c4eb89a7a9b46e5fef81881 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sun, 20 Jun 2021 00:13:50 +0200 Subject: transf fct correct in BTL scheduler proof --- scheduling/BTL_Schedulerproof.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 13ea61a6..95363194 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -104,12 +104,12 @@ Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. - intros TRANSF; destruct f; simpl. monadInv TRANSF. + intros TRANSF; destruct f; simpl; inv TRANSF. + exploit transf_function_correct; eauto. - (*intros (dupmap & MATCH_F).*) - (*eapply match_Internal; eauto.*) - (*+ eapply match_External.*) -Admitted. + econstructor. intros MATCH_F. + eapply match_Internal; eauto. + + eapply match_External. +Qed. Lemma function_ptr_preserved: forall v f, -- cgit From 6781fad2c0cf9403e761692ce248252f13b4aefe Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 24 Jun 2021 09:30:13 +0200 Subject: update BTLroadmap --- scheduling/BTLroadmap.md | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 23677ba2..3d799235 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -123,9 +123,7 @@ simultanée. **STATUS** -1. See [BTL.fsem] -2. fsem -> cfgsem: fait (ou presque). -3. cfgsem -> fsem: a faire (via verif de liveness). +DONE. ## "Basic" symbolic execution / symbolic simulation @@ -138,10 +136,13 @@ Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter - model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms wrt BTL.fsem - high-level specification [sstate_simu] of "symbolic simulation" over iblocks (wrt BTL.fsem). +- BTL_SEsimuref.v: DONE (raffinement de la notion d'état symbolique) + +- BTL_Schedulerproof: DONE (preuve que les simulations symboliques locales de blocks/blocks garantissent la simulation globale du programme pour BTL.fsem). + **TODO** -1. Verif du "scheduling" (aka analogue de RTLpathScheduler & RTLSchedulerproof). -2. raffinement intermediaire avant le hash-consing ? ça permettrait de décomposer encore davantage la preuve ? +implem du hash-consing. ## Port of RTLpath optimizations to BTL -- cgit From 6617e2aaac097301a950667a0c402d3ed8cd57be Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 25 Jun 2021 07:45:41 +0200 Subject: BTLroadmap: combinaison liveness et invariants ? --- scheduling/BTLroadmap.md | 56 ++++++++++++++++++++++++++++++++++++++++-------- 1 file changed, 47 insertions(+), 9 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md index 3d799235..1514c5d4 100644 --- a/scheduling/BTLroadmap.md +++ b/scheduling/BTLroadmap.md @@ -350,18 +350,39 @@ Au final, on obtient deux BDD identiques sur cet exemple (sans explosion combina **Rem** les mécanismes de propagation/réécritures décrits ci-dessus peuvent être aussi très utile pour la simulation symbolique modulo invariants (cf. ci-dessous) ! -## Invariants at block entry +## Peut-on coder l'égalité modulo liveness comme un cas simple d'invariant symbolique (et faire la vérif de liveneness par execution symbolique) ? -Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. Case-study: support of strength-reduction. +On peut peut-être simplifier le cadre actuel pour rendre l'intégration des invariants symboliques plus facile !! Une piste à creuser... + +Si la sémantique `BTL.fsem` a permis de rendre l'architecture un peu plus modulaire (mais sans vraiment simplifier sur le fond). +Elle ne simplifie pas vraiment l'intégration de la liveness avec les invariants. +Une approche qui pourrait réellement simplifier: voir le raisonnement modulo liveness comme un cas particulier d'invariant symbolique. + +**Idée** pour tout chemin d'exécution passant de `pc` et `pc'`, on ramène la simulation modulo liveness de `ib1` par `ib2` au test: + + [ib1;str_inputs(pc')] lessdef [str_inputs(pc);ib2] + +**Avantages possibles** -**PRELIMINARY IDEAS** +- plus de verif de liveness spécifique. +- plus de sémantique `BTL.fsem` +- les `str_inputs` ne seraient plus dans l'exécution symbolique, mais dans le test de simulation: ça évite les difficultés sur les jumptables. -- En pratique, il faudrait affiner la sémantique symbolique des "jumptable" pour que ça marche bien: avec un "état symbolique interne" par sortie (au lieu d'un état symbolique interne global dans la sémantique actuelle). Ça semble nécessaire pour gérer les invariants de renommage par exemple: en effet, il y a potentiellement un renommage différent pour chaque sortie. -Une solution possible: calquer Bjumptable sur Bcond (c-a-d autoriser les Bjumptable en milieu de blocs). On pourrait étendre la prédiction de branchement par profiling aux jumptables (par exemple, avoir des superblocks avec des jumptables au milieu). Un autre intérêt: simplifier (un peu) le support pour des formes SSA partielles (cf. ci-dessous). Ceci dit, ça compliquerait les choses à plein d'endroits (type Coq [iblock] mutuellement inductif avec [list_iblock] pour avoir les bons principes d'inductions, etc) pour des gains minimes en pratiques ? +Par contre, il faut réintroduire une preuve `lessdef` modulo liveness dans `BTL_Schedulerproof`. -- Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). +### Plan d'implem. incrémentale (dans une branche) + +1. Repartir d'une version de BTL sans info de liveness pour simplifier au maximum!! (ie on supprime la semantique fsem et on repart sur cfgsem). +2. Introduire la preuve `lessdef` modulo liveness dans `BTL_Schedulerproof` en *supposant* qu'on a le test de simulation symbolique qui nous arrange. + L'idée est de dégager des conditions suffisantes sur le test de simulation symbolique. +3. Réaliser le test de simulation symbolique. + +## Symbolic Invariants at block entry + +Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. Case-study: support of strength-reduction. -**EXAMPLE OF STRENGTH REDUCTION** + +### EXAMPLE OF STRENGTH REDUCTION On veut passer du code C1: @@ -382,7 +403,7 @@ au code C2: loop: // inputs: int *t, int s, int *ti, int *tn if (ti >= tn) goto exit; s += *ti; - ti += 4; + ti += 1; // i.e. sizeof(int) goto loop; exit; // inputs: int *t, int s @@ -405,11 +426,28 @@ Ci-dessus `ib1[TRANSFER]` dit qu'on a appliqué `TRANSFER(pc')` sur chacune des **REM** pour que ce soit correct, il faut sans doute vérifier une condition du style `ok(ib1) => ok(ib1[TRANSFER])`... +### Comment gérer le cas des registres résultat de call et builtin ? + +Rem: c'est la gestion de ces cas particuliers qui conduit à la notion de `pre_output_regs` (ou `BTL.pre_input`) dans la verif de la simulation modulo liveness. +On va retrouver ça dans la verif de la simulation modulo invariants symboliques. +Mais pas évident à éviter. + +Concrètement, on ne pourra pas mettre d'invariant qui porte sur le résultat d'un call ou d'un builtin (à part un pur renommage de registre). + +### GENERALISATION (What comes next ?) + +Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins). + +En gros, on pourrait passer d'invariants purement symboliques à des invariants combinant valeurs symboliques et valeurs abstraites. + ## Support of SSA-optimizations Minimum feature: extends BTL with "register renamings" at exits. This should enable to represent SSA-forms in BTL IR, more or less like in MLIR. -Maximum feature: add a basic instruction that performs parallel renaming of registers. If we also support Bjumptable in the middle of a block (see above), this simple feature would suffice to represent a very general notion of "partial SSA forms": since they could appear in the middle of a block, or just before an exit (such a parallel assignment would be forbidden in BTL<->RTL matching). +Maximum feature: add also a basic instruction that performs parallel renaming of registers. +This simple feature would help to represent a very general notion of "partial SSA forms": since they could appear in the middle of a block. + +In both case, such register renamings would be forbidden in BTL<->RTL matching (SSA-optimizations could only happen within BTL passes). ## Alias analysis in the symbolic simulation -- cgit From e907de8f2c30e5e004941b1ff6079eed503d83c3 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 30 Jun 2021 12:42:12 +0200 Subject: Starting symbolic execution implementation --- scheduling/BTL_SEimpl.v | 314 ++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 314 insertions(+) create mode 100644 scheduling/BTL_SEimpl.v (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v new file mode 100644 index 00000000..5c974efd --- /dev/null +++ b/scheduling/BTL_SEimpl.v @@ -0,0 +1,314 @@ +Require Import AST Maps. +Require Import Op Registers. +Require Import RTL BTL. +Require Import BTL_SEsimuref BTL_SEtheory OptionMonad. + +Require Import Impure.ImpHCons. +Import Notations. +Import HConsing. + +Local Open Scope option_monad_scope. +Local Open Scope impure. + +Import ListNotations. +Local Open Scope list_scope. + +(** Notations to make lemmas more readable *) +Notation "'sval_refines' ctx sv1 sv2" := (eval_sval ctx sv1 = eval_sval ctx sv2) + (only parsing, at level 0, ctx at next level, sv1 at next level, sv2 at next level): hse. + +Local Open Scope hse. + +Notation "'list_sval_refines' ctx lsv1 lsv2" := (eval_list_sval ctx lsv1 = eval_list_sval ctx lsv2) + (only parsing, at level 0, ctx at next level, lsv1 at next level, lsv2 at next level): hse. + +Notation "'smem_refines' ctx sm1 sm2" := (eval_smem ctx sm1 = eval_smem ctx sm2) + (only parsing, at level 0, ctx at next level, sm1 at next level, sm2 at next level): hse. + +(** Debug printer *) +Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := RET tt. (* TO REMOVE DEBUG INFO *) +(*Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := DO s <~ k x;; println ("DEBUG simu_check:" +; s). (* TO INSERT DEBUG INFO *)*) + +Definition DEBUG (s: pstring): ?? unit := XDEBUG tt (fun _ => RET s). + +(** * Implementation of Data-structure use in Hash-consing *) + +Definition sval_get_hid (sv: sval): hashcode := + match sv with + | Sundef hid => hid + | Sinput _ hid => hid + | Sop _ _ hid => hid + | Sload _ _ _ _ _ hid => hid + end. + +Definition list_sval_get_hid (lsv: list_sval): hashcode := + match lsv with + | Snil hid => hid + | Scons _ _ hid => hid + end. + +Definition smem_get_hid (sm: smem): hashcode := + match sm with + | Sinit hid => hid + | Sstore _ _ _ _ _ hid => hid + end. + +Definition sval_set_hid (sv: sval) (hid: hashcode): sval := + match sv with + | Sundef _ => Sundef hid + | Sinput r _ => Sinput r hid + | Sop o lsv _ => Sop o lsv hid + | Sload sm trap chunk addr lsv _ => Sload sm trap chunk addr lsv hid + end. + +Definition list_sval_set_hid (lsv: list_sval) (hid: hashcode): list_sval := + match lsv with + | Snil _ => Snil hid + | Scons sv lsv _ => Scons sv lsv hid + end. + +Definition smem_set_hid (sm: smem) (hid: hashcode): smem := + match sm with + | Sinit _ => Sinit hid + | Sstore sm chunk addr lsv srce _ => Sstore sm chunk addr lsv srce hid + end. + +(** * Implementation of symbolic execution *) +Section CanonBuilding. + +Variable hC_sval: hashinfo sval -> ?? sval. + +Hypothesis hC_sval_correct: forall s, + WHEN hC_sval s ~> s' THEN forall ctx, + sval_refines ctx (hdata s) s'. + +Variable hC_list_sval: hashinfo list_sval -> ?? list_sval. +Hypothesis hC_list_sval_correct: forall lh, + WHEN hC_list_sval lh ~> lh' THEN forall ctx, + list_sval_refines ctx (hdata lh) lh'. + +Variable hC_smem: hashinfo smem -> ?? smem. +Hypothesis hC_smem_correct: forall hm, + WHEN hC_smem hm ~> hm' THEN forall ctx, + smem_refines ctx (hdata hm) hm'. + +(* First, we wrap constructors for hashed values !*) + +Definition reg_hcode := 1. +Definition op_hcode := 2. +Definition load_hcode := 3. +Definition undef_code := 4. + +Definition hSinput_hcodes (r: reg) := + DO hc <~ hash reg_hcode;; + DO hv <~ hash r;; + RET [hc;hv]. +Extraction Inline hSinput_hcodes. + +Definition hSinput (r:reg): ?? sval := + DO hv <~ hSinput_hcodes r;; + hC_sval {| hdata:=Sinput r unknown_hid; hcodes :=hv; |}. + +Lemma hSinput_correct r: + WHEN hSinput r ~> hv THEN forall ctx, + sval_refines ctx hv (Sinput r unknown_hid). +Proof. + wlp_simplify. +Qed. +Global Opaque hSinput. +Local Hint Resolve hSinput_correct: wlp. + +Definition hSop_hcodes (op:operation) (lsv: list_sval) := + DO hc <~ hash op_hcode;; + DO hv <~ hash op;; + RET [hc;hv;list_sval_get_hid lsv]. +Extraction Inline hSop_hcodes. + +Definition hSop (op:operation) (lsv: list_sval): ?? sval := + DO hv <~ hSop_hcodes op lsv;; + hC_sval {| hdata:=Sop op lsv unknown_hid; hcodes :=hv |}. + +Lemma hSop_fSop_correct op lsv: + WHEN hSop op lsv ~> hv THEN forall ctx, + sval_refines ctx hv (fSop op lsv). +Proof. + wlp_simplify. +Qed. +Global Opaque hSop. +Local Hint Resolve hSop_fSop_correct: wlp_raw. + +Lemma hSop_correct op lsv1: + WHEN hSop op lsv1 ~> hv THEN forall ctx lsv2 + (LR: list_sval_refines ctx lsv1 lsv2), + sval_refines ctx hv (Sop op lsv2 unknown_hid). +Proof. + wlp_xsimplify ltac:(intuition eauto with wlp wlp_raw). + rewrite <- LR. erewrite H; eauto. +Qed. +Local Hint Resolve hSop_correct: wlp. + +Definition hSload_hcodes (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval):= + DO hc <~ hash load_hcode;; + DO hv1 <~ hash trap;; + DO hv2 <~ hash chunk;; + DO hv3 <~ hash addr;; + RET [hc; smem_get_hid sm; hv1; hv2; hv3; list_sval_get_hid lsv]. +Extraction Inline hSload_hcodes. + +Definition hSload (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval): ?? sval := + DO hv <~ hSload_hcodes sm trap chunk addr lsv;; + hC_sval {| hdata := Sload sm trap chunk addr lsv unknown_hid; hcodes := hv |}. + +Lemma hSload_correct sm1 trap chunk addr lsv1: + WHEN hSload sm1 trap chunk addr lsv1 ~> hv THEN forall ctx lsv2 sm2 hid + (LR: list_sval_refines ctx lsv1 lsv2) + (MR: smem_refines ctx sm1 sm2), + sval_refines ctx hv (Sload sm2 trap chunk addr lsv2 hid). +Proof. + wlp_simplify. + rewrite <- LR, <- MR. + auto. +Qed. +Global Opaque hSload. +Local Hint Resolve hSload_correct: wlp. + +Definition hSundef_hcodes := + DO hc <~ hash undef_code;; + RET [hc]. +Extraction Inline hSundef_hcodes. + +Definition hSundef : ?? sval := + DO hv <~ hSundef_hcodes;; + hC_sval {| hdata:=Sundef unknown_hid; hcodes :=hv; |}. + +Lemma hSundef_correct: + WHEN hSundef ~> hv THEN forall ctx, + sval_refines ctx hv (Sundef unknown_hid). +Proof. + wlp_simplify. +Qed. +Global Opaque hSundef. +Local Hint Resolve hSundef_correct: wlp. + +Definition hSnil (_: unit): ?? list_sval := + hC_list_sval {| hdata := Snil unknown_hid; hcodes := nil |}. + +Lemma hSnil_correct: + WHEN hSnil() ~> hv THEN forall ctx, + list_sval_refines ctx hv (Snil unknown_hid). +Proof. + wlp_simplify. +Qed. +Global Opaque hSnil. +Local Hint Resolve hSnil_correct: wlp. + +Definition hScons (sv: sval) (lsv: list_sval): ?? list_sval := + hC_list_sval {| hdata := Scons sv lsv unknown_hid; hcodes := [sval_get_hid sv; list_sval_get_hid lsv] |}. + +Lemma hScons_correct sv1 lsv1: + WHEN hScons sv1 lsv1 ~> lsv1' THEN forall ctx sv2 lsv2 + (VR: sval_refines ctx sv1 sv2) + (LR: list_sval_refines ctx lsv1 lsv2), + list_sval_refines ctx lsv1' (Scons sv2 lsv2 unknown_hid). +Proof. + wlp_simplify. + rewrite <- VR, <- LR. + auto. +Qed. +Global Opaque hScons. +Local Hint Resolve hScons_correct: wlp. + +Definition hSinit (_: unit): ?? smem := + hC_smem {| hdata := Sinit unknown_hid; hcodes := nil |}. + +Lemma hSinit_correct: + WHEN hSinit() ~> hm THEN forall ctx, + smem_refines ctx hm (Sinit unknown_hid). +Proof. + wlp_simplify. +Qed. +Global Opaque hSinit. +Local Hint Resolve hSinit_correct: wlp. + +Definition hSstore_hcodes (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval):= + DO hv1 <~ hash chunk;; + DO hv2 <~ hash addr;; + RET [smem_get_hid sm; hv1; hv2; list_sval_get_hid lsv; sval_get_hid srce]. +Extraction Inline hSstore_hcodes. + +Definition hSstore (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval): ?? smem := + DO hv <~ hSstore_hcodes sm chunk addr lsv srce;; + hC_smem {| hdata := Sstore sm chunk addr lsv srce unknown_hid; hcodes := hv |}. + +Lemma hSstore_correct sm1 chunk addr lsv1 sv1: + WHEN hSstore sm1 chunk addr lsv1 sv1 ~> sm1' THEN forall ctx lsv2 sm2 sv2 + (LR: list_sval_refines ctx lsv1 lsv2) + (MR: smem_refines ctx sm1 sm2) + (VR: sval_refines ctx sv1 sv2), + smem_refines ctx sm1' (Sstore sm2 chunk addr lsv2 sv2 unknown_hid). +Proof. + wlp_simplify. + rewrite <- LR, <- MR, <- VR. + auto. +Qed. +Global Opaque hSstore. +Local Hint Resolve hSstore_correct: wlp. + +Definition hris_sreg_eval ctx hris r := eval_sval ctx (ris_sreg_get hris r). + +Definition hris_sreg_get (hris: ristate) r: ?? sval := + match PTree.get r hris with + | None => if ris_input_init hris then hSinput r else hSundef + | Some sv => RET sv + end. +Coercion hris_sreg_get: ristate >-> Funclass. + +Lemma hris_sreg_get_correct hris r: + WHEN hris_sreg_get hris r ~> sv THEN forall ctx (f: reg -> sval) + (RR: forall r, hris_sreg_eval ctx hris r = eval_sval ctx (f r)), + sval_refines ctx sv (f r). +Proof. + unfold hris_sreg_eval, ris_sreg_get. wlp_simplify; rewrite <- RR; rewrite H; auto; + rewrite H0, H1; simpl; auto. +Qed. +Global Opaque hris_sreg_get. +Local Hint Resolve hris_sreg_get_correct: wlp. + +Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := + match o with + | Some x => RET x + | None => FAILWITH msg + end. + +Definition hris_init: ?? ristate + := DO hm <~ hSinit ();; + RET {| ris_smem := hm; ris_input_init := true; ok_rsval := nil; ris_sreg := PTree.empty _ |}. + +Lemma ris_init_correct: + WHEN hris_init ~> hris THEN + forall ctx, ris_refines ctx hris sis_init. +Proof. + unfold hris_init, sis_init; wlp_simplify. + econstructor; simpl in *; eauto. + + split; destruct 1; econstructor; simpl in *; + try rewrite H; try congruence; trivial. + + destruct 1; simpl in *. unfold ris_sreg_get; simpl. + intros; rewrite PTree.gempty; eauto. +Qed. + +(* +Definition hrexec f ib := hrexec_rec f ib hris_init (fun _ => Rabort). + +Definition hsexec (f: function) (pc:node): ?? ristate := + DO path <~ some_or_fail ((fn_code f)!pc) "hsexec.internal_error.1";; + DO hinit <~ init_ristate;; + DO hst <~ hsiexec_path path.(psize) f hinit;; + DO i <~ some_or_fail ((fn_code f)!(hst.(hsi_pc))) "hsexec.internal_error.2";; + DO ohst <~ hsiexec_inst i hst;; + match ohst with + | Some hst' => RET {| hinternal := hst'; hfinal := HSnone |} + | None => DO hsvf <~ hsexec_final i hst.(hsi_local);; + RET {| hinternal := hst; hfinal := hsvf |} + end.*) + +End CanonBuilding. -- cgit From f03d8bc434ea992bd390b949c1e0ce7dd99d2ddc Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 30 Jun 2021 14:57:48 +0200 Subject: some advance --- scheduling/BTL_SEimpl.v | 168 ++++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 154 insertions(+), 14 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 5c974efd..bf47b6da 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -1,4 +1,4 @@ -Require Import AST Maps. +Require Import Coqlib AST Maps. Require Import Op Registers. Require Import RTL BTL. Require Import BTL_SEsimuref BTL_SEtheory OptionMonad. @@ -13,6 +13,10 @@ Local Open Scope impure. Import ListNotations. Local Open Scope list_scope. +(** Tactics *) + +Ltac simplify_SOME x := repeat inversion_SOME x; try_simplify_someHyps. + (** Notations to make lemmas more readable *) Notation "'sval_refines' ctx sv1 sv2" := (eval_sval ctx sv1 = eval_sval ctx sv2) (only parsing, at level 0, ctx at next level, sv1 at next level, sv2 at next level): hse. @@ -160,10 +164,10 @@ Definition hSload (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: hC_sval {| hdata := Sload sm trap chunk addr lsv unknown_hid; hcodes := hv |}. Lemma hSload_correct sm1 trap chunk addr lsv1: - WHEN hSload sm1 trap chunk addr lsv1 ~> hv THEN forall ctx lsv2 sm2 hid + WHEN hSload sm1 trap chunk addr lsv1 ~> hv THEN forall ctx lsv2 sm2 (LR: list_sval_refines ctx lsv1 lsv2) (MR: smem_refines ctx sm1 sm2), - sval_refines ctx hv (Sload sm2 trap chunk addr lsv2 hid). + sval_refines ctx hv (Sload sm2 trap chunk addr lsv2 unknown_hid). Proof. wlp_simplify. rewrite <- LR, <- MR. @@ -254,25 +258,161 @@ Qed. Global Opaque hSstore. Local Hint Resolve hSstore_correct: wlp. -Definition hris_sreg_eval ctx hris r := eval_sval ctx (ris_sreg_get hris r). +Definition hrs_sreg_eval ctx hrs r := eval_sval ctx (ris_sreg_get hrs r). -Definition hris_sreg_get (hris: ristate) r: ?? sval := - match PTree.get r hris with - | None => if ris_input_init hris then hSinput r else hSundef +Definition hrs_sreg_get (hrs: ristate) r: ?? sval := + match PTree.get r hrs with + | None => if ris_input_init hrs then hSinput r else hSundef | Some sv => RET sv end. -Coercion hris_sreg_get: ristate >-> Funclass. +Coercion hrs_sreg_get: ristate >-> Funclass. -Lemma hris_sreg_get_correct hris r: - WHEN hris_sreg_get hris r ~> sv THEN forall ctx (f: reg -> sval) - (RR: forall r, hris_sreg_eval ctx hris r = eval_sval ctx (f r)), +Lemma hrs_sreg_get_correct hrs r: + WHEN hrs_sreg_get hrs r ~> sv THEN forall ctx (f: reg -> sval) + (RR: forall r, hrs_sreg_eval ctx hrs r = eval_sval ctx (f r)), sval_refines ctx sv (f r). Proof. - unfold hris_sreg_eval, ris_sreg_get. wlp_simplify; rewrite <- RR; rewrite H; auto; + unfold hrs_sreg_eval, ris_sreg_get. wlp_simplify; rewrite <- RR; rewrite H; auto; rewrite H0, H1; simpl; auto. Qed. -Global Opaque hris_sreg_get. -Local Hint Resolve hris_sreg_get_correct: wlp. +Global Opaque hrs_sreg_get. +Local Hint Resolve hrs_sreg_get_correct: wlp. + +Fixpoint hlist_args (hrs: ristate) (l: list reg): ?? list_sval := + match l with + | nil => hSnil() + | r::l => + DO v <~ hrs_sreg_get hrs r;; + DO lsv <~ hlist_args hrs l;; + hScons v lsv + end. + +Lemma hlist_args_correct hrs l: + WHEN hlist_args hrs l ~> lsv THEN forall ctx (f: reg -> sval) + (RR: forall r, hrs_sreg_eval ctx hrs r = eval_sval ctx (f r)), + list_sval_refines ctx lsv (list_sval_inj (List.map f l)). +Proof. + induction l; wlp_simplify. +Qed. +Global Opaque hlist_args. +Local Hint Resolve hlist_args_correct: wlp. + +(** Convert a "fake" hash-consed term into a "real" hash-consed term *) + +Fixpoint fsval_proj sv: ?? sval := + match sv with + | Sundef hc => + DO b <~ phys_eq hc unknown_hid;; + if b then hSundef else RET sv + | Sinput r hc => + DO b <~ phys_eq hc unknown_hid;; + if b then hSinput r (* was not yet really hash-consed *) + else RET sv (* already hash-consed *) + | Sop op hl hc => + DO b <~ phys_eq hc unknown_hid;; + if b then (* was not yet really hash-consed *) + DO hl' <~ fsval_list_proj hl;; + hSop op hl' + else RET sv (* already hash-consed *) + | Sload hm t chk addr hl _ => RET sv (* FIXME TODO gourdinl ? *) + end +with fsval_list_proj sl: ?? list_sval := + match sl with + | Snil hc => + DO b <~ phys_eq hc unknown_hid;; + if b then hSnil() (* was not yet really hash-consed *) + else RET sl (* already hash-consed *) + | Scons hv hl hc => + DO b <~ phys_eq hc unknown_hid;; + if b then (* was not yet really hash-consed *) + DO hv' <~ fsval_proj hv;; + DO hl' <~ fsval_list_proj hl;; + hScons hv' hl' + else RET sl (* already hash-consed *) + end. + +Lemma fsval_proj_correct sv: + WHEN fsval_proj sv ~> sv' THEN forall ctx, + eval_sval ctx sv = eval_sval ctx sv'. +Proof. + induction sv using sval_mut + with (P0 := fun lsv => + WHEN fsval_list_proj lsv ~> lsv' THEN forall ctx, + eval_list_sval ctx lsv = eval_list_sval ctx lsv') + (P1 := fun sm => True); try (wlp_simplify; tauto). + - wlp_xsimplify ltac:(intuition eauto with wlp_raw wlp). + rewrite H, H0; auto. + - wlp_simplify; erewrite H0, H1; eauto. +Qed. +Global Opaque fsval_proj. +Local Hint Resolve fsval_proj_correct: wlp. + +Lemma fsval_list_proj_correct lsv: + WHEN fsval_list_proj lsv ~> lsv' THEN forall ctx, + eval_list_sval ctx lsv = eval_list_sval ctx lsv'. +Proof. + induction lsv; wlp_simplify. + erewrite H0, H1; eauto. +Qed. +Global Opaque fsval_list_proj. +Local Hint Resolve fsval_list_proj_correct: wlp. + +(** ** Assignment of memory *) + +Definition hrexec_store chunk addr args src hrs: ?? ristate := + DO hargs <~ hlist_args hrs args;; + DO hsrc <~ hrs_sreg_get hrs src;; + DO hm <~ hSstore hrs.(ris_smem) chunk addr hargs hsrc;; + RET (rset_smem hm hrs). + +Lemma hrexec_store_correct chunk addr args src hrs: + WHEN hrexec_store chunk addr args src hrs ~> hrs' THEN forall ctx sis + (REF: ris_refines ctx hrs sis), + ris_refines ctx hrs' (sexec_store chunk addr args src sis). +Proof. + wlp_simplify. + eapply rset_mem_correct; simpl; eauto. + - intros X; erewrite H1; eauto. + rewrite X. simplify_SOME z. + - intros X; inversion REF. + erewrite H1; eauto. +Qed. + +(** ** Assignment of registers *) + +(** locally new symbolic values during symbolic execution *) +Inductive root_sval: Type := +| Rop (op: operation) +| Rload (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) +. + +Definition root_apply (rsv: root_sval) (lr: list reg) (st: sistate): sval := + let lsv := list_sval_inj (List.map (si_sreg st) lr) in + let sm := si_smem st in + match rsv with + | Rop op => fSop op lsv + | Rload trap chunk addr => fSload sm trap chunk addr lsv + end. +Coercion root_apply: root_sval >-> Funclass. + +Definition root_happly (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := + DO lsv <~ hlist_args hrs lr;; + match rsv with + | Rop op => hSop op lsv + | Rload trap chunk addr => hSload hrs.(ris_smem) trap chunk addr lsv + end. + +Lemma root_happly_correct (rsv: root_sval) lr hrs: + WHEN root_happly rsv lr hrs ~> sv THEN forall ctx st + (REF: ris_refines ctx hrs st) + (OK: ris_ok ctx hrs), + sval_refines ctx sv (rsv lr st). +Proof. + unfold root_apply, root_happly; destruct rsv; wlp_simplify; inv REF; + erewrite H0, H; eauto. +Qed. +Global Opaque root_happly. +Hint Resolve root_happly_correct: wlp. Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := match o with -- cgit From 7bc2a5e6d2694f144e49fe60af6ddfd0e8a0aba0 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 30 Jun 2021 15:59:25 +0200 Subject: coercions and simplify --- scheduling/BTL_SEimpl.v | 92 +++++++++++++++++++++++++++++++++++++++++++++- scheduling/BTL_SEsimuref.v | 2 +- 2 files changed, 91 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index bf47b6da..f01f1cc0 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -362,7 +362,7 @@ Local Hint Resolve fsval_list_proj_correct: wlp. Definition hrexec_store chunk addr args src hrs: ?? ristate := DO hargs <~ hlist_args hrs args;; DO hsrc <~ hrs_sreg_get hrs src;; - DO hm <~ hSstore hrs.(ris_smem) chunk addr hargs hsrc;; + DO hm <~ hSstore hrs chunk addr hargs hsrc;; RET (rset_smem hm hrs). Lemma hrexec_store_correct chunk addr args src hrs: @@ -399,7 +399,7 @@ Definition root_happly (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval : DO lsv <~ hlist_args hrs lr;; match rsv with | Rop op => hSop op lsv - | Rload trap chunk addr => hSload hrs.(ris_smem) trap chunk addr lsv + | Rload trap chunk addr => hSload hrs trap chunk addr lsv end. Lemma root_happly_correct (rsv: root_sval) lr hrs: @@ -414,6 +414,94 @@ Qed. Global Opaque root_happly. Hint Resolve root_happly_correct: wlp. +Local Open Scope lazy_bool_scope. + +(* NB: return [false] if the rsv cannot fail *) +Definition may_trap (rsv: root_sval) (lr: list reg): bool := + match rsv with + | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lr) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) + | Rload TRAP _ _ => true + | _ => false + end. + +Lemma lazy_orb_negb_false (b1 b2:bool): + (b1 ||| negb b2) = false <-> (b1 = false /\ b2 = true). +Proof. + unfold negb. repeat autodestruct; simpl; intuition (try congruence). +Qed. + +Lemma eval_list_sval_length ctx (f: reg -> sval) (l:list reg): + forall l', eval_list_sval ctx (list_sval_inj (List.map f l)) = Some l' -> + Datatypes.length l = Datatypes.length l'. +Proof. + induction l. + - simpl. intros. inv H. reflexivity. + - simpl. intros. destruct (eval_sval _ _); [|discriminate]. + destruct (eval_list_sval _ _) eqn:SLS; [|discriminate]. inv H. simpl. + erewrite IHl; eauto. +Qed. + +Lemma may_trap_correct ctx (rsv: root_sval) (lr: list reg) st: + may_trap rsv lr = false -> + eval_list_sval ctx (list_sval_inj (List.map (si_sreg st) lr)) <> None -> + eval_smem ctx (si_smem st) <> None -> + eval_sval ctx (rsv lr st) <> None. +Proof. + destruct rsv; simpl; try congruence. + - rewrite lazy_orb_negb_false. intros (TRAP1 & LEN) OK1 OK2. + autodestruct; try congruence. intros. + eapply is_trapping_op_sound; eauto. + erewrite <- eval_list_sval_length; eauto. + apply Nat.eqb_eq in LEN. + assumption. + - intros X OK1 OK2. + repeat autodestruct; try congruence. +Qed. + +(** simplify a symbolic value before assignment to a register *) +Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := + match rsv with + | Rop op => + match is_move_operation op lr with + | Some arg => hrs_sreg_get hrs arg (* optimization of Omove *) + | None => + DO lsv <~ hlist_args hrs lr;; + hSop op lsv + (* TODO gourdinl + match target_op_simplify op lr hst with + | Some fhv => fsval_proj fhv + | None => + hSop op lhsv + end*) + end + | Rload _ chunk addr => + DO lhsv <~ hlist_args hrs lr;; + hSload hrs NOTRAP chunk addr lhsv + end. + +Lemma simplify_correct rsv lr hrs: + WHEN simplify rsv lr hrs ~> hv THEN forall ctx st + (REF: ris_refines ctx hrs st) + (OK0: ris_ok ctx hrs) + (OK1: eval_sval ctx (rsv lr st) <> None), + sval_refines ctx hv (rsv lr st). +Proof. + destruct rsv; simpl; auto. + - (* Rop *) + destruct (is_move_operation _ _) eqn: Hmove. + { wlp_simplify; exploit is_move_operation_correct; eauto. + intros (Hop & Hlsv); subst; simpl in *. inv REF. + simplify_SOME z; erewrite H; eauto. } + wlp_simplify; inv REF. erewrite H0; eauto. + - (* Rload *) + destruct trap; wlp_simplify; inv REF. + + erewrite H0, H, MEM_EQ; eauto. + repeat simplify_SOME z. + * destruct (Memory.Mem.loadv _ _ _); try congruence. + * rewrite H1 in OK1; congruence. + + erewrite H0; eauto. +Qed. + Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := match o with | Some x => RET x diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 9c84532b..66abe6e1 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -20,7 +20,7 @@ Record ristate := { (** [ris_smem] represents the current smem symbolic evaluations. (we also recover the history of smem in ris_smem) *) - ris_smem: smem; + ris_smem:> smem; (** For the values in registers: 1) we store a list of sval evaluations 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval -- cgit From 4b370424090ab9e87fdf48ea5b04482728f99b7f Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 30 Jun 2021 17:37:03 +0200 Subject: cbranch expanse --- scheduling/BTL_SEimpl.v | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index f01f1cc0..df8cc47f 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -502,6 +502,31 @@ Proof. + erewrite H0; eauto. Qed. +Definition cbranch_expanse (prev: ristate) (cond: condition) (args: list reg): ?? (condition * list_sval) := + (* TODO gourdinl + match target_cbranch_expanse prev cond args with + | Some (cond', vargs) => + DO vargs' <~ fsval_list_proj vargs;; + RET (cond', vargs') + | None =>*) + DO vargs <~ hlist_args prev args ;; + RET (cond, vargs). + (*end.*) + +Lemma cbranch_expanse_correct hrs c l: + WHEN cbranch_expanse hrs c l ~> r THEN forall ctx st + (REF : ris_refines ctx hrs st) + (OK: ris_ok ctx hrs), + eval_scondition ctx (fst r) (snd r) = + eval_scondition ctx c (list_sval_inj (map (si_sreg st) l)). +Proof. + unfold cbranch_expanse. + wlp_simplify; inv REF. + unfold eval_scondition; erewrite <- H; eauto. +Qed. +Local Hint Resolve cbranch_expanse_correct: wlp. +Global Opaque cbranch_expanse. + Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := match o with | Some x => RET x -- cgit From ed7397533af0f850688256f07a2c5a22d6c58615 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 1 Jul 2021 16:25:05 +0200 Subject: red_PTree lemmas --- scheduling/BTL_SEimpl.v | 81 ++++++++++++++++++++++++++++++++++++++++------ scheduling/BTL_SEsimuref.v | 6 ++++ 2 files changed, 77 insertions(+), 10 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index df8cc47f..b007a07e 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -258,21 +258,18 @@ Qed. Global Opaque hSstore. Local Hint Resolve hSstore_correct: wlp. -Definition hrs_sreg_eval ctx hrs r := eval_sval ctx (ris_sreg_get hrs r). - Definition hrs_sreg_get (hrs: ristate) r: ?? sval := match PTree.get r hrs with | None => if ris_input_init hrs then hSinput r else hSundef | Some sv => RET sv end. -Coercion hrs_sreg_get: ristate >-> Funclass. Lemma hrs_sreg_get_correct hrs r: WHEN hrs_sreg_get hrs r ~> sv THEN forall ctx (f: reg -> sval) - (RR: forall r, hrs_sreg_eval ctx hrs r = eval_sval ctx (f r)), + (RR: forall r, sval_refines ctx (hrs r) (f r)), sval_refines ctx sv (f r). Proof. - unfold hrs_sreg_eval, ris_sreg_get. wlp_simplify; rewrite <- RR; rewrite H; auto; + unfold ris_sreg_get; wlp_simplify; rewrite <- RR; rewrite H; auto; rewrite H0, H1; simpl; auto. Qed. Global Opaque hrs_sreg_get. @@ -289,7 +286,7 @@ Fixpoint hlist_args (hrs: ristate) (l: list reg): ?? list_sval := Lemma hlist_args_correct hrs l: WHEN hlist_args hrs l ~> lsv THEN forall ctx (f: reg -> sval) - (RR: forall r, hrs_sreg_eval ctx hrs r = eval_sval ctx (f r)), + (RR: forall r, sval_refines ctx (hrs r) (f r)), list_sval_refines ctx lsv (list_sval_inj (List.map f l)). Proof. induction l; wlp_simplify. @@ -333,7 +330,7 @@ with fsval_list_proj sl: ?? list_sval := Lemma fsval_proj_correct sv: WHEN fsval_proj sv ~> sv' THEN forall ctx, - eval_sval ctx sv = eval_sval ctx sv'. + sval_refines ctx sv sv'. Proof. induction sv using sval_mut with (P0 := fun lsv => @@ -349,7 +346,7 @@ Local Hint Resolve fsval_proj_correct: wlp. Lemma fsval_list_proj_correct lsv: WHEN fsval_list_proj lsv ~> lsv' THEN forall ctx, - eval_list_sval ctx lsv = eval_list_sval ctx lsv'. + list_sval_refines ctx lsv lsv'. Proof. induction lsv; wlp_simplify. erewrite H0, H1; eauto. @@ -495,13 +492,53 @@ Proof. wlp_simplify; inv REF. erewrite H0; eauto. - (* Rload *) destruct trap; wlp_simplify; inv REF. - + erewrite H0, H, MEM_EQ; eauto. + + erewrite H0; eauto. + erewrite H; [|eapply REG_EQ; eauto]. + erewrite MEM_EQ; eauto. repeat simplify_SOME z. * destruct (Memory.Mem.loadv _ _ _); try congruence. * rewrite H1 in OK1; congruence. + erewrite H0; eauto. Qed. +Definition red_PTree_set (r: reg) (sv: sval) (hrs: ristate): ristate := + let sreg := + match (ris_input_init hrs), sv with + | true, Sinput r' _ => + if Pos.eq_dec r r' + then PTree.remove r' hrs + else PTree.set r sv hrs + | false, Sundef _ => + PTree.remove r hrs + | _, _ => PTree.set r sv hrs + end in + ris_sreg_set hrs sreg. + +Lemma red_PTree_set_correct (r r0:reg) sv hrs ctx: + sval_refines ctx ((red_PTree_set r sv hrs) r0) ((ris_sreg_set hrs (PTree.set r sv hrs)) r0). +Proof. + unfold red_PTree_set, ris_sreg_set, ris_sreg_get; simpl. + destruct (ris_input_init hrs) eqn:Hinit, sv; simpl; auto. + 1: destruct (Pos.eq_dec r r1); auto; subst; + rename r1 into r. + all: destruct (Pos.eq_dec r r0); auto; + [ subst; rewrite PTree.grs, PTree.gss; simpl; auto | + rewrite PTree.gro, PTree.gso; simpl; auto]. +Qed. + +Lemma red_PTree_set_refines (r r0:reg) ctx hrs st sv1 sv2: + ris_refines ctx hrs st -> + sval_refines ctx sv1 sv2 -> + ris_ok ctx hrs -> + sval_refines ctx ((red_PTree_set r sv1 hrs) r0) (if Pos.eq_dec r r0 then sv2 else si_sreg st r0). +Proof. + intros REF SREF OK; rewrite red_PTree_set_correct. + unfold ris_sreg_set, ris_sreg_get. + destruct (Pos.eq_dec r r0). + - subst; simpl. rewrite PTree.gss; simpl; auto. + - inv REF; simpl. rewrite PTree.gso; simpl; eauto. +Qed. + Definition cbranch_expanse (prev: ristate) (cond: condition) (args: list reg): ?? (condition * list_sval) := (* TODO gourdinl match target_cbranch_expanse prev cond args with @@ -549,7 +586,31 @@ Proof. intros; rewrite PTree.gempty; eauto. Qed. -(* +Fixpoint hrexec_rec f ib ris (k: ristate -> rstate): ?? rstate := + match ib with + (*| BF fin _ => Rfinal (tr_ris f fin ris) (sexec_final_sfv fin ris)*) + (* basic instructions *) + | Bnop _ => RET (k ris) + | Bop op args res _ => + DO next <~ ris_sreg_set hrs res + k (rexec_op op args res ris) + | Bload TRAP chunk addr args dst _ => k (rexec_load TRAP chunk addr args dst ris) + | Bload NOTRAP chunk addr args dst _ => Rabort + | Bstore chunk addr args src _ => k (rexec_store chunk addr args src ris) + (* composed instructions *) + | Bseq ib1 ib2 => + rexec_rec f ib1 ris (fun ris2 => rexec_rec f ib2 ris2 k) + | Bcond cond args ifso ifnot _ => + let args := list_sval_inj (List.map ris args) in + let ifso := rexec_rec f ifso ris k in + let ifnot := rexec_rec f ifnot ris k in + Rcond cond args ifso ifnot + + (* TODO to remove *) + | _ => RET (k ris) + end + . + Definition hrexec f ib := hrexec_rec f ib hris_init (fun _ => Rabort). Definition hsexec (f: function) (pc:node): ?? ristate := diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 66abe6e1..60c49e04 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -38,6 +38,12 @@ Definition ris_sreg_get (ris: ristate) r: sval := end. Coercion ris_sreg_get: ristate >-> Funclass. +Definition ris_sreg_set (ris: ristate) (sreg: PTree.t sval): ristate := + {| ris_smem := ris_smem ris; + ris_input_init := ris_input_init ris; + ok_rsval := ok_rsval ris; + ris_sreg := sreg |}. + Record ris_ok ctx (ris: ristate) : Prop := { OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None -- cgit From 804339839051759c5c77e63232d0aec25bb3846d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 5 Jul 2021 16:18:47 +0200 Subject: lemma on HC_set_reg --- scheduling/BTL_SEimpl.v | 206 +++++++++++++++++++++++++++++++++------------ scheduling/BTL_SEsimuref.v | 11 +++ 2 files changed, 165 insertions(+), 52 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index b007a07e..6e86bebb 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -400,10 +400,10 @@ Definition root_happly (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval : end. Lemma root_happly_correct (rsv: root_sval) lr hrs: - WHEN root_happly rsv lr hrs ~> sv THEN forall ctx st - (REF: ris_refines ctx hrs st) + WHEN root_happly rsv lr hrs ~> sv THEN forall ctx sis + (REF: ris_refines ctx hrs sis) (OK: ris_ok ctx hrs), - sval_refines ctx sv (rsv lr st). + sval_refines ctx sv (rsv lr sis). Proof. unfold root_apply, root_happly; destruct rsv; wlp_simplify; inv REF; erewrite H0, H; eauto. @@ -472,16 +472,16 @@ Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := end*) end | Rload _ chunk addr => - DO lhsv <~ hlist_args hrs lr;; - hSload hrs NOTRAP chunk addr lhsv + DO lsv <~ hlist_args hrs lr;; + hSload hrs NOTRAP chunk addr lsv end. Lemma simplify_correct rsv lr hrs: - WHEN simplify rsv lr hrs ~> hv THEN forall ctx st - (REF: ris_refines ctx hrs st) + WHEN simplify rsv lr hrs ~> hv THEN forall ctx sis + (REF: ris_refines ctx hrs sis) (OK0: ris_ok ctx hrs) - (OK1: eval_sval ctx (rsv lr st) <> None), - sval_refines ctx hv (rsv lr st). + (OK1: eval_sval ctx (rsv lr sis) <> None), + sval_refines ctx hv (rsv lr sis). Proof. destruct rsv; simpl; auto. - (* Rop *) @@ -500,22 +500,22 @@ Proof. * rewrite H1 in OK1; congruence. + erewrite H0; eauto. Qed. +Global Opaque simplify. +Local Hint Resolve simplify_correct: wlp. + +Definition red_PTree_set (r: reg) (sv: sval) (hrs: ristate): PTree.t sval := + match (ris_input_init hrs), sv with + | true, Sinput r' _ => + if Pos.eq_dec r r' + then PTree.remove r' hrs + else PTree.set r sv hrs + | false, Sundef _ => + PTree.remove r hrs + | _, _ => PTree.set r sv hrs + end. -Definition red_PTree_set (r: reg) (sv: sval) (hrs: ristate): ristate := - let sreg := - match (ris_input_init hrs), sv with - | true, Sinput r' _ => - if Pos.eq_dec r r' - then PTree.remove r' hrs - else PTree.set r sv hrs - | false, Sundef _ => - PTree.remove r hrs - | _, _ => PTree.set r sv hrs - end in - ris_sreg_set hrs sreg. - -Lemma red_PTree_set_correct (r r0:reg) sv hrs ctx: - sval_refines ctx ((red_PTree_set r sv hrs) r0) ((ris_sreg_set hrs (PTree.set r sv hrs)) r0). +Lemma red_PTree_set_correct (r r0:reg) sv (hrs: ristate) ctx: + sval_refines ctx ((ris_sreg_set hrs (red_PTree_set r sv hrs)) r0) ((ris_sreg_set hrs (PTree.set r sv hrs)) r0). Proof. unfold red_PTree_set, ris_sreg_set, ris_sreg_get; simpl. destruct (ris_input_init hrs) eqn:Hinit, sv; simpl; auto. @@ -526,11 +526,11 @@ Proof. rewrite PTree.gro, PTree.gso; simpl; auto]. Qed. -Lemma red_PTree_set_refines (r r0:reg) ctx hrs st sv1 sv2: - ris_refines ctx hrs st -> +Lemma red_PTree_set_refines (r r0:reg) ctx hrs sis sv1 sv2: + ris_refines ctx hrs sis -> sval_refines ctx sv1 sv2 -> ris_ok ctx hrs -> - sval_refines ctx ((red_PTree_set r sv1 hrs) r0) (if Pos.eq_dec r r0 then sv2 else si_sreg st r0). + sval_refines ctx (ris_sreg_set hrs (red_PTree_set r sv1 hrs) r0) (if Pos.eq_dec r r0 then sv2 else si_sreg sis r0). Proof. intros REF SREF OK; rewrite red_PTree_set_correct. unfold ris_sreg_set, ris_sreg_get. @@ -551,11 +551,11 @@ Definition cbranch_expanse (prev: ristate) (cond: condition) (args: list reg): ? (*end.*) Lemma cbranch_expanse_correct hrs c l: - WHEN cbranch_expanse hrs c l ~> r THEN forall ctx st - (REF : ris_refines ctx hrs st) + WHEN cbranch_expanse hrs c l ~> r THEN forall ctx sis + (REF : ris_refines ctx hrs sis) (OK: ris_ok ctx hrs), eval_scondition ctx (fst r) (snd r) = - eval_scondition ctx c (list_sval_inj (map (si_sreg st) l)). + eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)). Proof. unfold cbranch_expanse. wlp_simplify; inv REF. @@ -586,43 +586,145 @@ Proof. intros; rewrite PTree.gempty; eauto. Qed. -Fixpoint hrexec_rec f ib ris (k: ristate -> rstate): ?? rstate := +Definition hrset_sreg r lr rsv (hrs: ristate): ?? ristate := + DO ok_lsv <~ + (if may_trap rsv lr + then DO hv <~ root_happly rsv lr hrs;; + XDEBUG hv (fun hv => DO hv_name <~ string_of_hashcode (sval_get_hid hv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr hv_name))%string);; + RET (hv::(ok_rsval hrs)) + else RET (ok_rsval hrs));; + DO simp <~ simplify rsv lr hrs;; + RET {| ris_smem := hrs.(ris_smem); + ris_input_init := hrs.(ris_input_init); + ok_rsval := ok_lsv; + ris_sreg:= red_PTree_set r simp hrs |}. + +Lemma ok_hrset_sreg (rsv:root_sval) ctx (st: sistate) r lr: + si_ok ctx (set_sreg r (rsv lr st) st) + <-> (si_ok ctx st /\ eval_sval ctx (rsv lr st) <> None). +Proof. + unfold set_sreg; simpl; split. + - intros. destruct H as [[OK_SV OK_PRE] OK_SMEM OK_SREG]; simpl in *. + repeat (split; try tauto). + + intros r0; generalize (OK_SREG r0); clear OK_SREG; destruct (Pos.eq_dec r r0); try congruence. + + generalize (OK_SREG r); clear OK_SREG; destruct (Pos.eq_dec r r); try congruence. + - intros (OK & SEVAL). inv OK. + repeat (split; try tauto; eauto). + intros r0; destruct (Pos.eq_dec r r0) eqn:Heq; simpl; + rewrite Heq; eauto. +Qed. + +(* TODO gourdinl move this in BTL_SEtheory? *) +Lemma eval_list_sval_inj_not_none ctx st: forall l, + (forall r, List.In r l -> eval_sval ctx (si_sreg st r) = None -> False) -> + eval_list_sval ctx (list_sval_inj (map (si_sreg st) l)) = None -> False. +Proof. + induction l. + - intuition discriminate. + - intros ALLR. simpl. + inversion_SOME v. + + intro SVAL. inversion_SOME lv; [discriminate|]. + assert (forall r : reg, In r l -> eval_sval ctx (si_sreg st r) = None -> False). + { intros r INR. eapply ALLR. right. assumption. } + intro SVALLIST. intro. eapply IHl; eauto. + + intros. exploit (ALLR a); simpl; eauto. +Qed. + +Lemma hrset_sreg_correct r lr rsv hrs: + WHEN hrset_sreg r lr rsv hrs ~> hrs' THEN forall ctx sis + (REF: ris_refines ctx hrs sis), + ris_refines ctx hrs' (set_sreg r (rsv lr sis) sis). +Proof. + wlp_simplify; inversion REF. + - (* may_trap -> true *) + assert (X: si_ok ctx (set_sreg r (rsv lr sis) sis) <-> + ris_ok ctx {| ris_smem := hrs; + ris_input_init := ris_input_init hrs; + ok_rsval := exta :: ok_rsval hrs; + ris_sreg := red_PTree_set r exta0 hrs |}). + { + rewrite ok_hrset_sreg, OK_EQUIV. + split. + + intros (ROK & SEVAL); inv ROK. + assert (ROK: ris_ok ctx hrs) by (econstructor; eauto). + econstructor; eauto; simpl. + intuition (subst; eauto). + erewrite H0 in *; eauto. + + intros (OK_RMEM & OK_RREG); simpl in *. + assert (ROK: ris_ok ctx hrs) by (econstructor; eauto). + erewrite <- H0 in *; eauto. } + split; auto; rewrite <- X, ok_hrset_sreg. + + intuition eauto. + + intros (SOK & SEVAL) r0. + rewrite ris_sreg_set_access. + erewrite red_PTree_set_refines; intuition eauto. + - (* may_trap -> false *) + assert (X: si_ok ctx (set_sreg r (rsv lr sis) sis) <-> + ris_ok ctx {| ris_smem := hrs; + ris_input_init := ris_input_init hrs; + ok_rsval := ok_rsval hrs; + ris_sreg := red_PTree_set r exta hrs |}). + { + rewrite ok_hrset_sreg, OK_EQUIV. + split. + + intros (ROK & SEVAL); inv ROK. + econstructor; eauto. + + intros (OK_RMEM & OK_RREG). + assert (ROK: ris_ok ctx hrs) by (econstructor; eauto). + split; auto. + intros SNONE; exploit may_trap_correct; eauto. + * intros LNONE; eapply eval_list_sval_inj_not_none; eauto. + assert (SOK: si_ok ctx sis) by intuition. + inv SOK. intuition eauto. + * rewrite <- MEM_EQ; auto. } + split; auto; rewrite <- X, ok_hrset_sreg. + + intuition eauto. + + intros (SOK & SEVAL) r0. + rewrite ris_sreg_set_access. + erewrite red_PTree_set_refines; intuition eauto. +Qed. + +Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := match ib with - (*| BF fin _ => Rfinal (tr_ris f fin ris) (sexec_final_sfv fin ris)*) + | BF fin _ => RET (Rfinal (tr_ris f fin hrs) (sexec_final_sfv fin hrs)) (* basic instructions *) - | Bnop _ => RET (k ris) - | Bop op args res _ => - DO next <~ ris_sreg_set hrs res - k (rexec_op op args res ris) - | Bload TRAP chunk addr args dst _ => k (rexec_load TRAP chunk addr args dst ris) - | Bload NOTRAP chunk addr args dst _ => Rabort - | Bstore chunk addr args src _ => k (rexec_store chunk addr args src ris) + | Bnop _ => k hrs + | Bop op args dst _ => + DO next <~ hrset_sreg dst args (Rop op) hrs;; + k next + | Bload TRAP chunk addr args dst _ => + DO next <~ hrset_sreg dst args (Rload TRAP chunk addr) hrs;; + k next + | Bload NOTRAP chunk addr args dst _ => RET Rabort + | Bstore chunk addr args src _ => + DO next <~ hrexec_store chunk addr args src hrs;; + k next (* composed instructions *) | Bseq ib1 ib2 => - rexec_rec f ib1 ris (fun ris2 => rexec_rec f ib2 ris2 k) + hrexec_rec f ib1 hrs (fun hrs2 => hrexec_rec f ib2 hrs2 k) | Bcond cond args ifso ifnot _ => - let args := list_sval_inj (List.map ris args) in - let ifso := rexec_rec f ifso ris k in - let ifnot := rexec_rec f ifnot ris k in - Rcond cond args ifso ifnot - - (* TODO to remove *) - | _ => RET (k ris) + DO res <~ cbranch_expanse hrs cond args;; + let (cond, vargs) := res in + DO ifso <~ hrexec_rec f ifso hrs k;; + DO ifnot <~ hrexec_rec f ifnot hrs k;; + RET (Rcond cond vargs ifso ifnot) end . -Definition hrexec f ib := hrexec_rec f ib hris_init (fun _ => Rabort). +Definition hrexec f ib := + DO init <~ hris_init;; + hrexec_rec f ib init (fun _ => RET Rabort). -Definition hsexec (f: function) (pc:node): ?? ristate := +Definition hsexec (f: function) (pc:node): ?? rstate := DO path <~ some_or_fail ((fn_code f)!pc) "hsexec.internal_error.1";; - DO hinit <~ init_ristate;; - DO hst <~ hsiexec_path path.(psize) f hinit;; + (*DO hinit <~ init_ristate;;*) + DO hst <~ hrexec f f.(fn_entrypoint);; DO i <~ some_or_fail ((fn_code f)!(hst.(hsi_pc))) "hsexec.internal_error.2";; DO ohst <~ hsiexec_inst i hst;; match ohst with | Some hst' => RET {| hinternal := hst'; hfinal := HSnone |} | None => DO hsvf <~ hsexec_final i hst.(hsi_local);; RET {| hinternal := hst; hfinal := hsvf |} - end.*) + end. End CanonBuilding. diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 60c49e04..c812c607 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -44,6 +44,17 @@ Definition ris_sreg_set (ris: ristate) (sreg: PTree.t sval): ristate := ok_rsval := ok_rsval ris; ris_sreg := sreg |}. +Lemma ris_sreg_set_access (ris: ristate) (sreg: PTree.t sval) r smem rsval: + {| ris_smem := smem; + ris_input_init := ris_input_init ris; + ok_rsval := rsval; + ris_sreg := sreg |} r = + ris_sreg_set ris sreg r. +Proof. + unfold ris_sreg_set, ris_sreg_get; simpl. + reflexivity. +Qed. + Record ris_ok ctx (ris: ristate) : Prop := { OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None; OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None -- cgit From 0746cadf3bcdcbd3977c52e59352e5e7316b31d2 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 6 Jul 2021 11:45:52 +0200 Subject: Some advance, need to finish canonbuilding proofs --- scheduling/BTL_SEimpl.v | 283 +++++++++++++++++++++++++++++++++++++++++++++--- 1 file changed, 270 insertions(+), 13 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 6e86bebb..8f7881a5 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -1,6 +1,6 @@ Require Import Coqlib AST Maps. Require Import Op Registers. -Require Import RTL BTL. +Require Import RTL BTL Errors. Require Import BTL_SEsimuref BTL_SEtheory OptionMonad. Require Import Impure.ImpHCons. @@ -77,6 +77,95 @@ Definition smem_set_hid (sm: smem) (hid: hashcode): smem := | Sstore sm chunk addr lsv srce _ => Sstore sm chunk addr lsv srce hid end. +(** Now, we build the hash-Cons value from a "hash_eq". + + Informal specification: + [hash_eq] must be consistent with the "hashed" constructors defined above. + + We expect that hashinfo values in the code of these "hashed" constructors verify: + (hash_eq (hdata x) (hdata y) ~> true) <-> (hcodes x)=(hcodes y) +*) + +Definition sval_hash_eq (sv1 sv2: sval): ?? bool := + match sv1, sv2 with + | Sinput r1 _, Sinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *) + | Sop op1 lsv1 _, Sop op2 lsv2 _ => + DO b1 <~ phys_eq lsv1 lsv2;; + if b1 + then struct_eq op1 op2 (* NB: really need a struct_eq here ? *) + else RET false + | Sload sm1 trap1 chk1 addr1 lsv1 _, Sload sm2 trap2 chk2 addr2 lsv2 _ => + DO b1 <~ phys_eq lsv1 lsv2;; + DO b2 <~ phys_eq sm1 sm2;; + DO b3 <~ struct_eq trap1 trap2;; + DO b4 <~ struct_eq chk1 chk2;; + if b1 && b2 && b3 && b4 + then struct_eq addr1 addr2 + else RET false + | _,_ => RET false + end. + +Lemma and_true_split a b: a && b = true <-> a = true /\ b = true. +Proof. + destruct a; simpl; intuition. +Qed. + +Lemma sval_hash_eq_correct x y: + WHEN sval_hash_eq x y ~> b THEN + b = true -> sval_set_hid x unknown_hid = sval_set_hid y unknown_hid. +Proof. + destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence. +Qed. +Global Opaque sval_hash_eq. +Local Hint Resolve sval_hash_eq_correct: wlp. + +Definition list_sval_hash_eq (lsv1 lsv2: list_sval): ?? bool := + match lsv1, lsv2 with + | Snil _, Snil _ => RET true + | Scons sv1 lsv1' _, Scons sv2 lsv2' _ => + DO b <~ phys_eq lsv1' lsv2';; + if b + then phys_eq sv1 sv2 + else RET false + | _,_ => RET false + end. + +Lemma list_sval_hash_eq_correct x y: + WHEN list_sval_hash_eq x y ~> b THEN + b = true -> list_sval_set_hid x unknown_hid = list_sval_set_hid y unknown_hid. +Proof. + destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence. +Qed. +Global Opaque list_sval_hash_eq. +Local Hint Resolve list_sval_hash_eq_correct: wlp. + +Definition smem_hash_eq (sm1 sm2: smem): ?? bool := + match sm1, sm2 with + | Sinit _, Sinit _ => RET true + | Sstore sm1 chk1 addr1 lsv1 sv1 _, Sstore sm2 chk2 addr2 lsv2 sv2 _ => + DO b1 <~ phys_eq lsv1 lsv2;; + DO b2 <~ phys_eq sm1 sm2;; + DO b3 <~ phys_eq sv1 sv2;; + DO b4 <~ struct_eq chk1 chk2;; + if b1 && b2 && b3 && b4 + then struct_eq addr1 addr2 + else RET false + | _,_ => RET false + end. + +Lemma smem_hash_eq_correct x y: + WHEN smem_hash_eq x y ~> b THEN + b = true -> smem_set_hid x unknown_hid = smem_set_hid y unknown_hid. +Proof. + destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence. +Qed. +Global Opaque smem_hash_eq. +Local Hint Resolve smem_hash_eq_correct: wlp. + +Definition hSVAL: hashP sval := {| hash_eq := sval_hash_eq; get_hid:=sval_get_hid; set_hid:=sval_set_hid |}. +Definition hLSVAL: hashP list_sval := {| hash_eq := list_sval_hash_eq; get_hid:= list_sval_get_hid; set_hid:= list_sval_set_hid |}. +Definition hSMEM: hashP smem := {| hash_eq := smem_hash_eq; get_hid:= smem_get_hid; set_hid:= smem_set_hid |}. + (** * Implementation of symbolic execution *) Section CanonBuilding. @@ -465,7 +554,7 @@ Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := DO lsv <~ hlist_args hrs lr;; hSop op lsv (* TODO gourdinl - match target_op_simplify op lr hst with + match target_op_simplify op lr hrs with | Some fhv => fsval_proj fhv | None => hSop op lhsv @@ -715,16 +804,184 @@ Definition hrexec f ib := DO init <~ hris_init;; hrexec_rec f ib init (fun _ => RET Rabort). -Definition hsexec (f: function) (pc:node): ?? rstate := - DO path <~ some_or_fail ((fn_code f)!pc) "hsexec.internal_error.1";; - (*DO hinit <~ init_ristate;;*) - DO hst <~ hrexec f f.(fn_entrypoint);; - DO i <~ some_or_fail ((fn_code f)!(hst.(hsi_pc))) "hsexec.internal_error.2";; - DO ohst <~ hsiexec_inst i hst;; - match ohst with - | Some hst' => RET {| hinternal := hst'; hfinal := HSnone |} - | None => DO hsvf <~ hsexec_final i hst.(hsi_local);; - RET {| hinternal := hst; hfinal := hsvf |} - end. +Lemma hrexec_rec_correct1 ctx ib: + forall rk k + (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) + (CONT: forall hrs sis lsis sfv st, + ris_refines ctx hrs sis -> + k sis = st -> + get_soutcome ctx (k sis) = Some (sout lsis sfv) -> + si_ok ctx lsis -> + WHEN rk hrs ~> ris' THEN + rst_refines ctx ris' (k sis)) + hrs sis lsis sfv st + (REF: ris_refines ctx hrs sis) + (EXEC: sexec_rec (cf ctx) ib sis k = st) + (SOUT: get_soutcome ctx st = Some (sout lsis sfv)) + (OK: si_ok ctx lsis), + WHEN hrexec_rec (cf ctx) ib hrs rk ~> hrs THEN + rst_refines ctx hrs st. +Proof. + (* + induction ib; wlp_simplify. + - admit. + - eapply CONT; eauto. + - try_simplify_someHyps. intros OUT. + eapply CONT; eauto. + intuition eauto. + econstructor; intuition eauto. + inv REF. inv H2. + intuition eauto. + + + - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. + - (* seq *) + intros; subst. + eapply IHib1. 3-6: eauto. + + simpl. eapply sexec_rec_okpreserv; eauto. + + intros; subst. eapply IHib2; eauto. + - (* cond *) + intros rk k CONTh CONT hrs sis lsis sfv st REF EXEC OUT OK. subst. + assert (rOK: ris_ok ctx hrs). { + erewrite <- OK_EQUIV. 2: eauto. + eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. + } + generalize OUT; clear OUT; simpl. + autodestruct. + intros COND; generalize COND. + erewrite <- eval_scondition_refpreserv; eauto. + econstructor; try_simplify_someHyps. + Qed.*) +Admitted. + +Definition hsexec (f: function) ib: ?? rstate := + hrexec f ib. End CanonBuilding. + +(** * Implementing the simulation test with concrete hash-consed symbolic execution *) + +Definition phys_check {A} (x y:A) (msg: pstring): ?? unit := + DO b <~ phys_eq x y;; + assert_b b msg;; + RET tt. + +Definition struct_check {A} (x y: A) (msg: pstring): ?? unit := + DO b <~ struct_eq x y;; + assert_b b msg;; + RET tt. + +Lemma struct_check_correct {A} (a b: A) msg: + WHEN struct_check a b msg ~> _ THEN + a = b. +Proof. wlp_simplify. Qed. +Global Opaque struct_check. +Hint Resolve struct_check_correct: wlp. + +Definition sfval_simu_check (sfv1 sfv2: sfval) := RET tt. + +Fixpoint rst_simu_check (rst1 rst2: rstate) := + match rst1, rst2 with + | Rfinal ris1 sfv1, Rfinal ris2 sfv2 => + sfval_simu_check sfv1 sfv2 + | Rcond cond1 lsv1 rsL1 rsR1, Rcond cond2 lsv2 rsL2 rsR2 => + struct_check cond1 cond2 "hsstate_simu_check: conditions do not match";; + phys_check lsv1 lsv2 "hsstate_simu_check: args do not match";; + rst_simu_check rsL1 rsL2;; + rst_simu_check rsR1 rsR2 + | _, _ => FAILWITH "hsstate_simu_check: simu_check failed" + end. + +Lemma rst_simu_check_correct rst1 rst2: + WHEN rst_simu_check rst1 rst2 ~> _ THEN + rst_simu rst1 rst2. +Proof. + induction rst1, rst2; + wlp_simplify. +Admitted. +Hint Resolve rst_simu_check_correct: wlp. +Global Opaque rst_simu_check. + +Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? (option (node * node)) := + (* creating the hash-consing tables *) + DO hC_sval <~ hCons hSVAL;; + DO hC_list_hsval <~ hCons hLSVAL;; + DO hC_hsmem <~ hCons hSMEM;; + let hsexec := hsexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in + (* performing the hash-consed executions *) + DO hst1 <~ hsexec f1 ib1;; + DO hst2 <~ hsexec f2 ib2;; + (* comparing the executions *) + rst_simu_check hst1 hst2. + +Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): + WHEN simu_check_single f1 f2 ib1 ib2 ~> _ THEN + symbolic_simu f1 f2 ib1 ib2. +Proof. +Admitted. +Global Opaque simu_check_single. +Global Hint Resolve simu_check_single_correct: wlp. + +Fixpoint simu_check_rec (f1 f2: function) (ibf1 ibf2: iblock_info): ?? unit := + DO res <~ simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; + let (pc1', pc2') := res in + simu_check_rec pc1' pc2'. + +Lemma simu_check_rec_correct dm f tf lm: + WHEN simu_check_rec dm f tf lm ~> _ THEN + forall pc1 pc2, In (pc2, pc1) lm -> sexec_simu dm f tf pc1 pc2. +Proof. + induction lm; wlp_simplify. + match goal with + | X: (_,_) = (_,_) |- _ => inversion X; subst + end. + subst; eauto. +Qed. +Global Opaque simu_check_rec. +Global Hint Resolve simu_check_rec_correct: wlp. + +Definition imp_simu_check (f tf: function): ?? unit := + simu_check_rec f tf ;; + DEBUG("simu_check OK!"). + +Local Hint Resolve PTree.elements_correct: core. +Lemma imp_simu_check_correct f tf: + WHEN imp_simu_check f tf ~> _ THEN + forall sexec_simu f tf. +Proof. + wlp_simplify. +Qed. +Global Opaque imp_simu_check. +Global Hint Resolve imp_simu_check_correct: wlp. + +Program Definition aux_simu_check (f tf: function): ?? bool := + DO r <~ + (TRY + imp_simu_check f tf;; + RET true + CATCH_FAIL s, _ => + println ("simu_check_failure:" +; s);; + RET false + ENSURE (sexec_simu f tf));; + RET (`r). +Obligation 1. + split; wlp_simplify. discriminate. +Qed. + +Lemma aux_simu_check_correct f tf: + WHEN aux_simu_check f tf ~> b THEN + sexec_simu f tf. +Proof. + unfold aux_simu_check; wlp_simplify. + destruct exta; simpl; auto. +Qed. + +(* Coerce aux_simu_check into a pure function (this is a little unsafe like all oracles in CompCert). *) + +Import UnsafeImpure. + +Definition simu_check (f tf: function) : res unit := + match unsafe_coerce (aux_simu_check f tf) with + | Some true => OK tt + | _ => Error (msg "simu_check has failed") + end. -- cgit From 8bcada5be5fc93c5526d6381b5d034558e48de03 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 7 Jul 2021 09:58:49 +0200 Subject: progress in simulation test --- scheduling/BTL_SEimpl.v | 307 +++++++++++++++++++++++++++++++++++++++--------- 1 file changed, 250 insertions(+), 57 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 8f7881a5..94d7375d 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -166,6 +166,17 @@ Definition hSVAL: hashP sval := {| hash_eq := sval_hash_eq; get_hid:=sval_get_hi Definition hLSVAL: hashP list_sval := {| hash_eq := list_sval_hash_eq; get_hid:= list_sval_get_hid; set_hid:= list_sval_set_hid |}. Definition hSMEM: hashP smem := {| hash_eq := smem_hash_eq; get_hid:= smem_get_hid; set_hid:= smem_set_hid |}. +Program Definition mk_hash_params: Dict.hash_params sval := + {| + Dict.test_eq := phys_eq; + Dict.hashing := fun (sv: sval) => RET (sval_get_hid sv); + Dict.log := fun sv => + DO sv_name <~ string_of_hashcode (sval_get_hid sv);; + println ("unexpected undef behavior of hashcode:" +; (CamlStr sv_name)) |}. +Obligation 1. + wlp_simplify. +Qed. + (** * Implementation of symbolic execution *) Section CanonBuilding. @@ -194,17 +205,17 @@ Definition undef_code := 4. Definition hSinput_hcodes (r: reg) := DO hc <~ hash reg_hcode;; - DO hv <~ hash r;; - RET [hc;hv]. + DO sv <~ hash r;; + RET [hc;sv]. Extraction Inline hSinput_hcodes. Definition hSinput (r:reg): ?? sval := - DO hv <~ hSinput_hcodes r;; - hC_sval {| hdata:=Sinput r unknown_hid; hcodes :=hv; |}. + DO sv <~ hSinput_hcodes r;; + hC_sval {| hdata:=Sinput r unknown_hid; hcodes :=sv; |}. Lemma hSinput_correct r: - WHEN hSinput r ~> hv THEN forall ctx, - sval_refines ctx hv (Sinput r unknown_hid). + WHEN hSinput r ~> sv THEN forall ctx, + sval_refines ctx sv (Sinput r unknown_hid). Proof. wlp_simplify. Qed. @@ -213,17 +224,17 @@ Local Hint Resolve hSinput_correct: wlp. Definition hSop_hcodes (op:operation) (lsv: list_sval) := DO hc <~ hash op_hcode;; - DO hv <~ hash op;; - RET [hc;hv;list_sval_get_hid lsv]. + DO sv <~ hash op;; + RET [hc;sv;list_sval_get_hid lsv]. Extraction Inline hSop_hcodes. Definition hSop (op:operation) (lsv: list_sval): ?? sval := - DO hv <~ hSop_hcodes op lsv;; - hC_sval {| hdata:=Sop op lsv unknown_hid; hcodes :=hv |}. + DO sv <~ hSop_hcodes op lsv;; + hC_sval {| hdata:=Sop op lsv unknown_hid; hcodes :=sv |}. Lemma hSop_fSop_correct op lsv: - WHEN hSop op lsv ~> hv THEN forall ctx, - sval_refines ctx hv (fSop op lsv). + WHEN hSop op lsv ~> sv THEN forall ctx, + sval_refines ctx sv (fSop op lsv). Proof. wlp_simplify. Qed. @@ -231,9 +242,9 @@ Global Opaque hSop. Local Hint Resolve hSop_fSop_correct: wlp_raw. Lemma hSop_correct op lsv1: - WHEN hSop op lsv1 ~> hv THEN forall ctx lsv2 + WHEN hSop op lsv1 ~> sv THEN forall ctx lsv2 (LR: list_sval_refines ctx lsv1 lsv2), - sval_refines ctx hv (Sop op lsv2 unknown_hid). + sval_refines ctx sv (Sop op lsv2 unknown_hid). Proof. wlp_xsimplify ltac:(intuition eauto with wlp wlp_raw). rewrite <- LR. erewrite H; eauto. @@ -242,21 +253,21 @@ Local Hint Resolve hSop_correct: wlp. Definition hSload_hcodes (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval):= DO hc <~ hash load_hcode;; - DO hv1 <~ hash trap;; - DO hv2 <~ hash chunk;; - DO hv3 <~ hash addr;; - RET [hc; smem_get_hid sm; hv1; hv2; hv3; list_sval_get_hid lsv]. + DO sv1 <~ hash trap;; + DO sv2 <~ hash chunk;; + DO sv3 <~ hash addr;; + RET [hc; smem_get_hid sm; sv1; sv2; sv3; list_sval_get_hid lsv]. Extraction Inline hSload_hcodes. Definition hSload (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval): ?? sval := - DO hv <~ hSload_hcodes sm trap chunk addr lsv;; - hC_sval {| hdata := Sload sm trap chunk addr lsv unknown_hid; hcodes := hv |}. + DO sv <~ hSload_hcodes sm trap chunk addr lsv;; + hC_sval {| hdata := Sload sm trap chunk addr lsv unknown_hid; hcodes := sv |}. Lemma hSload_correct sm1 trap chunk addr lsv1: - WHEN hSload sm1 trap chunk addr lsv1 ~> hv THEN forall ctx lsv2 sm2 + WHEN hSload sm1 trap chunk addr lsv1 ~> sv THEN forall ctx lsv2 sm2 (LR: list_sval_refines ctx lsv1 lsv2) (MR: smem_refines ctx sm1 sm2), - sval_refines ctx hv (Sload sm2 trap chunk addr lsv2 unknown_hid). + sval_refines ctx sv (Sload sm2 trap chunk addr lsv2 unknown_hid). Proof. wlp_simplify. rewrite <- LR, <- MR. @@ -271,12 +282,12 @@ Definition hSundef_hcodes := Extraction Inline hSundef_hcodes. Definition hSundef : ?? sval := - DO hv <~ hSundef_hcodes;; - hC_sval {| hdata:=Sundef unknown_hid; hcodes :=hv; |}. + DO sv <~ hSundef_hcodes;; + hC_sval {| hdata:=Sundef unknown_hid; hcodes :=sv; |}. Lemma hSundef_correct: - WHEN hSundef ~> hv THEN forall ctx, - sval_refines ctx hv (Sundef unknown_hid). + WHEN hSundef ~> sv THEN forall ctx, + sval_refines ctx sv (Sundef unknown_hid). Proof. wlp_simplify. Qed. @@ -287,8 +298,8 @@ Definition hSnil (_: unit): ?? list_sval := hC_list_sval {| hdata := Snil unknown_hid; hcodes := nil |}. Lemma hSnil_correct: - WHEN hSnil() ~> hv THEN forall ctx, - list_sval_refines ctx hv (Snil unknown_hid). + WHEN hSnil() ~> sv THEN forall ctx, + list_sval_refines ctx sv (Snil unknown_hid). Proof. wlp_simplify. Qed. @@ -324,14 +335,14 @@ Global Opaque hSinit. Local Hint Resolve hSinit_correct: wlp. Definition hSstore_hcodes (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval):= - DO hv1 <~ hash chunk;; - DO hv2 <~ hash addr;; - RET [smem_get_hid sm; hv1; hv2; list_sval_get_hid lsv; sval_get_hid srce]. + DO sv1 <~ hash chunk;; + DO sv2 <~ hash addr;; + RET [smem_get_hid sm; sv1; sv2; list_sval_get_hid lsv; sval_get_hid srce]. Extraction Inline hSstore_hcodes. Definition hSstore (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval): ?? smem := - DO hv <~ hSstore_hcodes sm chunk addr lsv srce;; - hC_smem {| hdata := Sstore sm chunk addr lsv srce unknown_hid; hcodes := hv |}. + DO sv <~ hSstore_hcodes sm chunk addr lsv srce;; + hC_smem {| hdata := Sstore sm chunk addr lsv srce unknown_hid; hcodes := sv |}. Lemma hSstore_correct sm1 chunk addr lsv1 sv1: WHEN hSstore sm1 chunk addr lsv1 sv1 ~> sm1' THEN forall ctx lsv2 sm2 sv2 @@ -408,12 +419,12 @@ with fsval_list_proj sl: ?? list_sval := DO b <~ phys_eq hc unknown_hid;; if b then hSnil() (* was not yet really hash-consed *) else RET sl (* already hash-consed *) - | Scons hv hl hc => + | Scons sv hl hc => DO b <~ phys_eq hc unknown_hid;; if b then (* was not yet really hash-consed *) - DO hv' <~ fsval_proj hv;; + DO sv' <~ fsval_proj sv;; DO hl' <~ fsval_list_proj hl;; - hScons hv' hl' + hScons sv' hl' else RET sl (* already hash-consed *) end. @@ -555,7 +566,7 @@ Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := hSop op lsv (* TODO gourdinl match target_op_simplify op lr hrs with - | Some fhv => fsval_proj fhv + | Some fsv => fsval_proj fsv | None => hSop op lhsv end*) @@ -566,11 +577,11 @@ Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := end. Lemma simplify_correct rsv lr hrs: - WHEN simplify rsv lr hrs ~> hv THEN forall ctx sis + WHEN simplify rsv lr hrs ~> sv THEN forall ctx sis (REF: ris_refines ctx hrs sis) (OK0: ris_ok ctx hrs) (OK1: eval_sval ctx (rsv lr sis) <> None), - sval_refines ctx hv (rsv lr sis). + sval_refines ctx sv (rsv lr sis). Proof. destruct rsv; simpl; auto. - (* Rop *) @@ -678,9 +689,9 @@ Qed. Definition hrset_sreg r lr rsv (hrs: ristate): ?? ristate := DO ok_lsv <~ (if may_trap rsv lr - then DO hv <~ root_happly rsv lr hrs;; - XDEBUG hv (fun hv => DO hv_name <~ string_of_hashcode (sval_get_hid hv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr hv_name))%string);; - RET (hv::(ok_rsval hrs)) + then DO sv <~ root_happly rsv lr hrs;; + XDEBUG sv (fun sv => DO sv_name <~ string_of_hashcode (sval_get_hid sv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr sv_name))%string);; + RET (sv::(ok_rsval hrs)) else RET (ok_rsval hrs));; DO simp <~ simplify rsv lr hrs;; RET {| ris_smem := hrs.(ris_smem); @@ -878,11 +889,186 @@ Proof. wlp_simplify. Qed. Global Opaque struct_check. Hint Resolve struct_check_correct: wlp. -Definition sfval_simu_check (sfv1 sfv2: sfval) := RET tt. +Definition option_eq_check {A} (o1 o2: option A): ?? unit := + match o1, o2 with + | Some x1, Some x2 => phys_check x1 x2 "option_eq_check: data physically differ" + | None, None => RET tt + | _, _ => FAILWITH "option_eq_check: structure differs" + end. -Fixpoint rst_simu_check (rst1 rst2: rstate) := +Lemma option_eq_check_correct A (o1 o2: option A): WHEN option_eq_check o1 o2 ~> _ THEN o1=o2. +Proof. + wlp_simplify. +Qed. +Global Opaque option_eq_check. +Hint Resolve option_eq_check_correct:wlp. + +Import PTree. + +Fixpoint PTree_eq_check {A} (d1 d2: PTree.t A): ?? unit := + match d1, d2 with + | Leaf, Leaf => RET tt + | Node l1 o1 r1, Node l2 o2 r2 => + option_eq_check o1 o2;; + PTree_eq_check l1 l2;; + PTree_eq_check r1 r2 + | _, _ => FAILWITH "PTree_eq_check: some key is absent" + end. + +Lemma PTree_eq_check_correct A d1: forall (d2: t A), + WHEN PTree_eq_check d1 d2 ~> _ THEN forall x, PTree.get x d1 = PTree.get x d2. +Proof. + induction d1 as [|l1 Hl1 o1 r1 Hr1]; destruct d2 as [|l2 o2 r2]; simpl; + wlp_simplify. destruct x; simpl; auto. +Qed. +Global Opaque PTree_eq_check. +Local Hint Resolve PTree_eq_check_correct: wlp. + +Definition hrs_simu_check (hrs1 hrs2: ristate) : ?? unit := + DEBUG("? hrs_simu_check");; + phys_check (ris_smem hrs1) (ris_smem hrs2) "hrs_simu_check: ris_smem sets aren't equiv";; + phys_check (ris_input_init hrs1) (ris_input_init hrs2) "hrs_simu_check: ris_input_init bools aren't equiv";; + PTree_eq_check (ris_sreg hrs1) (ris_sreg hrs2);; + Sets.assert_list_incl mk_hash_params (ok_rsval hrs2) (ok_rsval hrs1);; + DEBUG("=> hrs_simu_check: OK"). + +Lemma setoid_in {A: Type} (a: A): forall l, + SetoidList.InA (fun x y => x = y) a l -> + In a l. +Proof. + induction l; intros; inv H. + - constructor. reflexivity. + - right. auto. +Qed. + +Lemma regset_elements_in r rs: + Regset.In r rs -> + In r (Regset.elements rs). +Proof. + intros. exploit Regset.elements_1; eauto. intro SIN. + apply setoid_in. assumption. +Qed. +Local Hint Resolve regset_elements_in: core. + +Lemma hrs_simu_check_correct hrs1 hrs2: + WHEN hrs_simu_check hrs1 hrs2 ~> _ THEN + ris_simu hrs1 hrs2. +Proof. + wlp_simplify; constructor; auto. + unfold ris_sreg_get; intros r; rewrite H, H1; reflexivity. +Qed. +Hint Resolve hrs_simu_check_correct: wlp. +Global Opaque hrs_simu_check. + +Definition svos_simu_check (svos1 svos2: sval + ident) := + match svos1, svos2 with + | inl sv1, inl sv2 => phys_check sv1 sv2 "svos_simu_check: sval mismatch" + | inr id1, inr id2 => phys_check id1 id2 "svos_simu_check: symbol mismatch" + | _, _ => FAILWITH "svos_simu_check: type mismatch" + end. + +Lemma svos_simu_check_correct svos1 svos2: + WHEN svos_simu_check svos1 svos2 ~> _ THEN + svos1 = svos2. +Proof. + destruct svos1; destruct svos2; wlp_simplify. +Qed. +Global Opaque svos_simu_check. +Hint Resolve svos_simu_check_correct: wlp. + +Fixpoint builtin_arg_simu_check (bs bs': builtin_arg sval) := + match bs with + | BA sv => + match bs' with + | BA sv' => phys_check sv sv' "builtin_arg_simu_check: sval mismatch" + | _ => FAILWITH "builtin_arg_simu_check: BA mismatch" + end + | BA_splitlong lo hi => + match bs' with + | BA_splitlong lo' hi' => + builtin_arg_simu_check lo lo';; + builtin_arg_simu_check hi hi' + | _ => FAILWITH "builtin_arg_simu_check: BA_splitlong mismatch" + end + | BA_addptr b1 b2 => + match bs' with + | BA_addptr b1' b2' => + builtin_arg_simu_check b1 b1';; + builtin_arg_simu_check b2 b2' + | _ => FAILWITH "builtin_arg_simu_check: BA_addptr mismatch" + end + | bs => struct_check bs bs' "builtin_arg_simu_check: basic mismatch" + end. + +Lemma builtin_arg_simu_check_correct: forall bs1 bs2, + WHEN builtin_arg_simu_check bs1 bs2 ~> _ THEN + bs1 = bs2. +Proof. + induction bs1. + all: try (wlp_simplify; subst; reflexivity). + all: destruct bs2; wlp_simplify; congruence. +Qed. +Global Opaque builtin_arg_simu_check. +Hint Resolve builtin_arg_simu_check_correct: wlp. + +Fixpoint list_builtin_arg_simu_check lbs1 lbs2 := + match lbs1, lbs2 with + | nil, nil => RET tt + | bs1::lbs1, bs2::lbs2 => + builtin_arg_simu_check bs1 bs2;; + list_builtin_arg_simu_check lbs1 lbs2 + | _, _ => FAILWITH "list_builtin_arg_simu_check: length mismatch" + end. + +Lemma list_builtin_arg_simu_check_correct: forall lbs1 lbs2, + WHEN list_builtin_arg_simu_check lbs1 lbs2 ~> _ THEN + lbs1 = lbs2. +Proof. + induction lbs1; destruct lbs2; wlp_simplify. congruence. +Qed. +Global Opaque list_builtin_arg_simu_check. +Hint Resolve list_builtin_arg_simu_check_correct: wlp. + +Definition sfval_simu_check (sfv1 sfv2: sfval): ?? unit := + match sfv1, sfv2 with + | Sgoto e1, Sgoto e2 => + phys_check e1 e2 "sfval_simu_check: Sgoto successors do not match" + | Scall sig1 svos1 lsv1 res1 e1, Scall sig2 svos2 lsv2 res2 e2 => + phys_check e1 e2 "sfval_simu_check: Scall successors do not match";; + phys_check sig1 sig2 "sfval_simu_check: Scall different signatures";; + phys_check res1 res2 "sfval_simu_check: Scall res do not match";; + svos_simu_check svos1 svos2;; + phys_check lsv1 lsv2 "sfval_simu_check: Scall args do not match" + | Stailcall sig1 svos1 lsv1, Stailcall sig2 svos2 lsv2 => + phys_check sig1 sig2 "sfval_simu_check: Stailcall different signatures";; + svos_simu_check svos1 svos2;; + phys_check lsv1 lsv2 "sfval_simu_check: Stailcall args do not match" + | Sbuiltin ef1 lbs1 br1 e1, Sbuiltin ef2 lbs2 br2 e2 => + phys_check e1 e2 "sfval_simu_check: Sbuiltin successors do not match";; + phys_check ef1 ef2 "sfval_simu_check: Sbuiltin ef do not match";; + phys_check br1 br2 "sfval_simu_check: Sbuiltin br do not match";; + list_builtin_arg_simu_check lbs1 lbs2 + | Sjumptable sv1 le1, Sjumptable sv2 le2 => + phys_check le1 le2 "sfval_simu_check: Sjumptable successors do not match";; + phys_check sv1 sv2 "sfval_simu_check: Sjumptable sval do not match" + | Sreturn osv1, Sreturn osv2 => + option_eq_check osv1 osv2 + | _, _ => FAILWITH "sfval_simu_check: structure mismatch" + end. + +Lemma sfval_simu_check_correct sfv1 sfv2: + WHEN sfval_simu_check sfv1 sfv2 ~> _ THEN + rfv_simu sfv1 sfv2. +Proof. + destruct sfv1; destruct sfv2; simpl; wlp_simplify; try congruence. +Qed. +Hint Resolve sfval_simu_check_correct: wlp. +Global Opaque sfval_simu_check. + +Fixpoint rst_simu_check (rst1 rst2: rstate) {struct rst1} := match rst1, rst2 with - | Rfinal ris1 sfv1, Rfinal ris2 sfv2 => + | Rfinal hrs1 sfv1, Rfinal hrs2 sfv2 => + hrs_simu_check hrs1 hrs2;; sfval_simu_check sfv1 sfv2 | Rcond cond1 lsv1 rsL1 rsR1, Rcond cond2 lsv2 rsL2 rsR2 => struct_check cond1 cond2 "hsstate_simu_check: conditions do not match";; @@ -892,17 +1078,18 @@ Fixpoint rst_simu_check (rst1 rst2: rstate) := | _, _ => FAILWITH "hsstate_simu_check: simu_check failed" end. -Lemma rst_simu_check_correct rst1 rst2: +Lemma rst_simu_check_correct rst1: forall rst2, WHEN rst_simu_check rst1 rst2 ~> _ THEN rst_simu rst1 rst2. Proof. - induction rst1, rst2; - wlp_simplify. -Admitted. + induction rst1; destruct rst2; + wlp_simplify; constructor; auto. +Qed. Hint Resolve rst_simu_check_correct: wlp. Global Opaque rst_simu_check. -Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? (option (node * node)) := +(* TODO do we really need both functions *) +Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? unit := (* creating the hash-consing tables *) DO hC_sval <~ hCons hSVAL;; DO hC_list_hsval <~ hCons hLSVAL;; @@ -918,18 +1105,24 @@ Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): WHEN simu_check_single f1 f2 ib1 ib2 ~> _ THEN symbolic_simu f1 f2 ib1 ib2. Proof. + unfold symbolic_simu; wlp_simplify. Admitted. Global Opaque simu_check_single. Global Hint Resolve simu_check_single_correct: wlp. -Fixpoint simu_check_rec (f1 f2: function) (ibf1 ibf2: iblock_info): ?? unit := - DO res <~ simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; - let (pc1', pc2') := res in - simu_check_rec pc1' pc2'. +Fixpoint simu_check_rec (f1 f2: function) (blks: list node): ?? unit := + match blks with + | nil => RET tt + | pc :: blk' => + DO ibf1 <~ some_or_fail (fn_code f1)!pc "simu_check_rec: incorrect PTree for f1";; + DO ibf2 <~ some_or_fail (fn_code f2)!pc "simu_check_rec: incorrect PTree for f2";; + DO res <~ simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; + simu_check_rec f1 f2 blk' + end. -Lemma simu_check_rec_correct dm f tf lm: - WHEN simu_check_rec dm f tf lm ~> _ THEN - forall pc1 pc2, In (pc2, pc1) lm -> sexec_simu dm f tf pc1 pc2. +Lemma simu_check_rec_correct f1 f2: + WHEN simu_check_rec f1 f2 (PTree.elements (fn_code f1)) ~> _ THEN + symbolic_simu f1 f2 . Proof. induction lm; wlp_simplify. match goal with -- cgit From d8cf80248c6aee35aff49ee8d4ca9d1c8fc906f5 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 7 Jul 2021 11:41:04 +0200 Subject: hrexec1 --- scheduling/BTL_SEimpl.v | 108 ++++++++++++++++++++++++++++++++++++------------ 1 file changed, 81 insertions(+), 27 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 94d7375d..c32001d4 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -670,15 +670,15 @@ Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := | None => FAILWITH msg end. -Definition hris_init: ?? ristate +Definition hrs_init: ?? ristate := DO hm <~ hSinit ();; RET {| ris_smem := hm; ris_input_init := true; ok_rsval := nil; ris_sreg := PTree.empty _ |}. -Lemma ris_init_correct: - WHEN hris_init ~> hris THEN - forall ctx, ris_refines ctx hris sis_init. +Lemma hrs_init_correct: + WHEN hrs_init ~> hrs THEN + forall ctx, ris_refines ctx hrs sis_init. Proof. - unfold hris_init, sis_init; wlp_simplify. + unfold hrs_init, sis_init; wlp_simplify. econstructor; simpl in *; eauto. + split; destruct 1; econstructor; simpl in *; try rewrite H; try congruence; trivial. @@ -686,7 +686,7 @@ Proof. intros; rewrite PTree.gempty; eauto. Qed. -Definition hrset_sreg r lr rsv (hrs: ristate): ?? ristate := +Definition hrs_sreg_set r lr rsv (hrs: ristate): ?? ristate := DO ok_lsv <~ (if may_trap rsv lr then DO sv <~ root_happly rsv lr hrs;; @@ -730,8 +730,8 @@ Proof. + intros. exploit (ALLR a); simpl; eauto. Qed. -Lemma hrset_sreg_correct r lr rsv hrs: - WHEN hrset_sreg r lr rsv hrs ~> hrs' THEN forall ctx sis +Lemma hrs_sreg_set_correct r lr rsv hrs: + WHEN hrs_sreg_set r lr rsv hrs ~> hrs' THEN forall ctx sis (REF: ris_refines ctx hrs sis), ris_refines ctx hrs' (set_sreg r (rsv lr sis) sis). Proof. @@ -783,6 +783,8 @@ Proof. rewrite ris_sreg_set_access. erewrite red_PTree_set_refines; intuition eauto. Qed. +Global Opaque hrs_sreg_set. +Local Hint Resolve hrs_sreg_set_correct: wlp. Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := match ib with @@ -790,10 +792,10 @@ Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := (* basic instructions *) | Bnop _ => k hrs | Bop op args dst _ => - DO next <~ hrset_sreg dst args (Rop op) hrs;; + DO next <~ hrs_sreg_set dst args (Rop op) hrs;; k next | Bload TRAP chunk addr args dst _ => - DO next <~ hrset_sreg dst args (Rload TRAP chunk addr) hrs;; + DO next <~ hrs_sreg_set dst args (Rload TRAP chunk addr) hrs;; k next | Bload NOTRAP chunk addr args dst _ => RET Rabort | Bstore chunk addr args src _ => @@ -812,7 +814,7 @@ Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := . Definition hrexec f ib := - DO init <~ hris_init;; + DO init <~ hrs_init;; hrexec_rec f ib init (fun _ => RET Rabort). Lemma hrexec_rec_correct1 ctx ib: @@ -823,28 +825,25 @@ Lemma hrexec_rec_correct1 ctx ib: k sis = st -> get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> - WHEN rk hrs ~> ris' THEN - rst_refines ctx ris' (k sis)) + WHEN rk hrs ~> rst THEN + rst_refines ctx rst (k sis)) hrs sis lsis sfv st (REF: ris_refines ctx hrs sis) (EXEC: sexec_rec (cf ctx) ib sis k = st) (SOUT: get_soutcome ctx st = Some (sout lsis sfv)) (OK: si_ok ctx lsis), - WHEN hrexec_rec (cf ctx) ib hrs rk ~> hrs THEN - rst_refines ctx hrs st. + WHEN hrexec_rec (cf ctx) ib hrs rk ~> rst THEN + rst_refines ctx rst st. Proof. - (* induction ib; wlp_simplify. - admit. - eapply CONT; eauto. - - try_simplify_someHyps. intros OUT. + - autodestruct; try_simplify_someHyps; try congruence; intros; subst. +(* + unfold sexec_load. + exploit hrs_sreg_set_correct; eauto. eapply CONT; eauto. - intuition eauto. - econstructor; intuition eauto. - inv REF. inv H2. - intuition eauto. - - + destruct next. - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. - (* seq *) intros; subst. @@ -865,8 +864,56 @@ Proof. Qed.*) Admitted. -Definition hsexec (f: function) ib: ?? rstate := - hrexec f ib. +Lemma hrexec_correct1 ctx ib sis sfv: + get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> + (si_ok ctx sis) -> + WHEN hrexec (cf ctx) ib ~> rst THEN + rst_refines ctx rst (sexec (cf ctx) ib). +Proof. + unfold sexec; intros; wlp_simplify. + eapply hrexec_rec_correct1; eauto; simpl; try congruence; split. + - intuition eauto. + + inv H2; constructor; auto; simpl; congruence. + + inv H2; split; simpl in *; try congruence; trivial. + - intros X; simpl in *; auto. + - intros X r; inv X; simpl in *. + unfold ris_sreg_get; simpl; rewrite PTree.gempty; simpl. + reflexivity. +Qed. + +(*Lemma hrexec_rec_correct2 ctx ib: + forall rk k + (CONTh: forall ris lris rfv, get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris) + (CONT: forall ris sis lris rfv st, + ris_refines ctx ris sis -> + rk ris = st -> + get_routcome ctx (rk ris) = Some (rout lris rfv) -> + ris_ok ctx lris -> + rst_refines ctx (rk ris) (k sis)) + ris sis lris rfv st + (REF: ris_refines ctx ris sis) + (EXEC: rexec_rec (cf ctx) ib ris rk = st) + (SOUT: get_routcome ctx st = Some (rout lris rfv)) + (OK: ris_ok ctx lris) + , rst_refines ctx st (sexec_rec (cf ctx) ib sis k). +Proof. + Admitted.*) + +Lemma hrexec_correct2 ctx ib hrs rfv: + WHEN hrexec (cf ctx) ib ~> rst THEN + get_routcome ctx rst = Some (rout hrs rfv) -> + (ris_ok ctx hrs) -> + rst_refines ctx rst (sexec (cf ctx) ib). +Proof. +Admitted. + +Lemma hrexec_simu_correct ctx ib: + WHEN hrexec (cf ctx) ib ~> rst THEN + exists st, sexec (cf ctx) ib = st /\ rst_refines ctx rst st. +Proof. + unfold hrexec; wlp_simplify. + repeat eexists. +Admitted. End CanonBuilding. @@ -1094,7 +1141,7 @@ Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? unit := DO hC_sval <~ hCons hSVAL;; DO hC_list_hsval <~ hCons hLSVAL;; DO hC_hsmem <~ hCons hSMEM;; - let hsexec := hsexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in + let hsexec := hrexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in (* performing the hash-consed executions *) DO hst1 <~ hsexec f1 ib1;; DO hst2 <~ hsexec f2 ib2;; @@ -1104,8 +1151,15 @@ Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? unit := Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): WHEN simu_check_single f1 f2 ib1 ib2 ~> _ THEN symbolic_simu f1 f2 ib1 ib2. -Proof. +Proof. (* TODO *) unfold symbolic_simu; wlp_simplify. + eapply rst_simu_correct; eauto. + + intros; eapply hrexec_correct1; eauto. + 1-3: wlp_simplify. unfold hrexec. + simpl. eauto. + intuition eauto. + + intros; eapply rexec_correct2; eauto. + intros ctxz. Admitted. Global Opaque simu_check_single. Global Hint Resolve simu_check_single_correct: wlp. -- cgit From 0cda2f3441252eadc1d7901942935bf0c2a2949c Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 7 Jul 2021 12:06:09 +0200 Subject: symbolic_simu impl definitions --- scheduling/BTL_SEimpl.v | 56 ++++++++----------------------------------------- 1 file changed, 9 insertions(+), 47 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index c32001d4..a2b5cc43 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -1158,66 +1158,28 @@ Proof. (* TODO *) 1-3: wlp_simplify. unfold hrexec. simpl. eauto. intuition eauto. - + intros; eapply rexec_correct2; eauto. - intros ctxz. + admit. Admitted. Global Opaque simu_check_single. Global Hint Resolve simu_check_single_correct: wlp. -Fixpoint simu_check_rec (f1 f2: function) (blks: list node): ?? unit := - match blks with - | nil => RET tt - | pc :: blk' => - DO ibf1 <~ some_or_fail (fn_code f1)!pc "simu_check_rec: incorrect PTree for f1";; - DO ibf2 <~ some_or_fail (fn_code f2)!pc "simu_check_rec: incorrect PTree for f2";; - DO res <~ simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; - simu_check_rec f1 f2 blk' - end. - -Lemma simu_check_rec_correct f1 f2: - WHEN simu_check_rec f1 f2 (PTree.elements (fn_code f1)) ~> _ THEN - symbolic_simu f1 f2 . -Proof. - induction lm; wlp_simplify. - match goal with - | X: (_,_) = (_,_) |- _ => inversion X; subst - end. - subst; eauto. -Qed. -Global Opaque simu_check_rec. -Global Hint Resolve simu_check_rec_correct: wlp. - -Definition imp_simu_check (f tf: function): ?? unit := - simu_check_rec f tf ;; - DEBUG("simu_check OK!"). - -Local Hint Resolve PTree.elements_correct: core. -Lemma imp_simu_check_correct f tf: - WHEN imp_simu_check f tf ~> _ THEN - forall sexec_simu f tf. -Proof. - wlp_simplify. -Qed. -Global Opaque imp_simu_check. -Global Hint Resolve imp_simu_check_correct: wlp. - -Program Definition aux_simu_check (f tf: function): ?? bool := +Program Definition aux_simu_check (f1 f2: function) (ib1 ib2: iblock): ?? bool := DO r <~ (TRY - imp_simu_check f tf;; + simu_check_single f1 f2 ib1 ib2;; RET true CATCH_FAIL s, _ => println ("simu_check_failure:" +; s);; RET false - ENSURE (sexec_simu f tf));; + ENSURE (fun b => b=true -> symbolic_simu f1 f2 ib1 ib2));; RET (`r). Obligation 1. split; wlp_simplify. discriminate. Qed. -Lemma aux_simu_check_correct f tf: - WHEN aux_simu_check f tf ~> b THEN - sexec_simu f tf. +Lemma aux_simu_check_correct (f1 f2: function) (ib1 ib2: iblock): + WHEN aux_simu_check f1 f2 ib1 ib2 ~> b THEN + b=true -> symbolic_simu f1 f2 ib1 ib2. Proof. unfold aux_simu_check; wlp_simplify. destruct exta; simpl; auto. @@ -1227,8 +1189,8 @@ Qed. Import UnsafeImpure. -Definition simu_check (f tf: function) : res unit := - match unsafe_coerce (aux_simu_check f tf) with +Definition simu_check (f1 f2: function) (ib1 ib2: iblock): res unit := + match unsafe_coerce (aux_simu_check f1 f2 ib1 ib2) with | Some true => OK tt | _ => Error (msg "simu_check has failed") end. -- cgit From f995f01c747877cb8595dbf71b05d23205d697d8 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 7 Jul 2021 16:21:13 +0200 Subject: some progress in canonbuilding --- scheduling/BTL_SEimpl.v | 120 ++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 106 insertions(+), 14 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index a2b5cc43..25819ffa 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -474,6 +474,8 @@ Proof. - intros X; inversion REF. erewrite H1; eauto. Qed. +Global Opaque hrexec_store. +Local Hint Resolve hrexec_store_correct: wlp. (** ** Assignment of registers *) @@ -786,9 +788,103 @@ Qed. Global Opaque hrs_sreg_set. Local Hint Resolve hrs_sreg_set_correct: wlp. +(* transfer *) + +Lemma ok_hrset_red_sreg ctx r rsv hrs: + ris_ok ctx (ris_sreg_set hrs (red_PTree_set r rsv hrs)) + <-> (ris_ok ctx hrs). +Proof. + split; destruct 1; econstructor; simpl in *; eauto. +Qed. + +Lemma hrset_red_reg_correct ctx r rsv sv hrs sis: + ris_refines ctx hrs sis -> + (ris_ok ctx hrs -> eval_sval ctx rsv <> None) -> + (ris_ok ctx hrs -> eval_sval ctx rsv = eval_sval ctx sv) -> + ris_refines ctx (ris_sreg_set hrs (red_PTree_set r rsv hrs)) (set_sreg r sv sis). +Proof. + destruct 1; intros. + econstructor; eauto. + + rewrite ok_set_sreg, ok_hrset_red_sreg; intuition congruence. + + rewrite ok_hrset_red_sreg; intuition eauto. + + rewrite ok_hrset_red_sreg. intros; unfold set_sreg; simpl; + erewrite red_PTree_set_refines; eauto. + econstructor; eauto. +Qed. + +Fixpoint transfer_hrs (inputs: list reg) (hrs: ristate): ristate := + match inputs with + | nil => {| ris_smem := hrs.(ris_smem); + ris_input_init := false; + ok_rsval := hrs.(ok_rsval); + ris_sreg:= PTree.empty _ + |} + | r::l => + let hrs' := transfer_hrs l hrs in + ris_sreg_set hrs' (red_PTree_set r (ris_sreg_get hrs r) hrs') + end. + +Lemma ok_transfer_hrs ctx inputs hrs: + ris_ok ctx (transfer_hrs inputs hrs) + <-> (ris_ok ctx hrs). +Proof. + induction inputs as [|r l]; simpl. + + split; destruct 1; econstructor; simpl in *; intuition eauto. + + rewrite ok_hrset_red_sreg. auto. +Qed. + +Lemma transfer_hrs_correct ctx inputs hrs sis: + ris_refines ctx hrs sis -> + ris_refines ctx (transfer_hrs inputs hrs) (transfer_sis inputs sis). +Proof. + destruct 1; intros. + induction inputs as [|r l]. + + econstructor; eauto. + * erewrite ok_transfer_sis, ok_transfer_hrs; eauto. + * erewrite ok_transfer_hrs; eauto. + * erewrite ok_transfer_hrs; simpl; unfold ris_sreg_get; simpl; eauto. + intros; rewrite PTree.gempty. simpl; auto. + + econstructor; eauto. + * erewrite ok_transfer_sis, ok_transfer_hrs; eauto. + * erewrite ok_transfer_hrs; simpl. + intros; inversion IHl. erewrite MEM_EQ0. + - unfold transfer_sis; simpl; eauto. + - rewrite ok_transfer_hrs; simpl; eauto. + * erewrite ok_transfer_hrs; simpl. + intros H r0; inversion IHl. + erewrite red_PTree_set_refines; eauto. + - destruct (Pos.eq_dec); eauto. + - rewrite ok_transfer_hrs. auto. +Qed. + +Definition tr_hrs := poly_tr (fun f tbl or => transfer_hrs (Regset.elements (pre_inputs f tbl or))). + +Local Hint Resolve transfer_hrs_correct ok_transfer_hrs: core. +Local Opaque transfer_hrs. + +Lemma ok_tr_hrs ctx fi hrs: + ris_ok ctx (tr_hrs (cf ctx) fi hrs) + <-> (ris_ok ctx hrs). +Proof. + destruct fi; simpl; eauto. +Qed. + +Lemma hrs_ok_tr_okpreserv ctx fi hrs: ris_ok ctx (tr_hrs (cf ctx) fi hrs) -> ris_ok ctx hrs. +Proof. + rewrite ok_tr_hrs; auto. +Qed. + +Lemma tr_hrs_correct ctx fi hrs sis: + ris_refines ctx hrs sis -> + ris_refines ctx (tr_hrs (cf ctx) fi hrs) (tr_sis (cf ctx) fi sis). +Proof. + intros REF. rewrite <- tr_sis_alt_def. + destruct fi; simpl; eauto. +Qed. + Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := match ib with - | BF fin _ => RET (Rfinal (tr_ris f fin hrs) (sexec_final_sfv fin hrs)) + | BF fin _ => RET (Rfinal (tr_hrs f fin hrs) (sexec_final_sfv fin hrs)) (* basic instructions *) | Bnop _ => k hrs | Bop op args dst _ => @@ -817,6 +913,9 @@ Definition hrexec f ib := DO init <~ hrs_init;; hrexec_rec f ib init (fun _ => RET Rabort). +Local Hint Resolve exec_final_refpreserv tr_hrs_correct hrs_ok_tr_okpreserv: core. +Local Hint Constructors rst_refines: core. + Lemma hrexec_rec_correct1 ctx ib: forall rk k (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis) @@ -835,23 +934,16 @@ Lemma hrexec_rec_correct1 ctx ib: WHEN hrexec_rec (cf ctx) ib hrs rk ~> rst THEN rst_refines ctx rst st. Proof. - induction ib; wlp_simplify. - - admit. - - eapply CONT; eauto. - - autodestruct; try_simplify_someHyps; try congruence; intros; subst. -(* - unfold sexec_load. - exploit hrs_sreg_set_correct; eauto. - eapply CONT; eauto. - destruct next. - - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. + induction ib; try (wlp_simplify; fail). + - wlp_simplify; eapply CONT; eauto. + - destruct trap; wlp_simplify. - (* seq *) - intros; subst. - eapply IHib1. 3-6: eauto. + wlp_simplify. + eapply IHib1. 3-7: eauto. + simpl. eapply sexec_rec_okpreserv; eauto. + intros; subst. eapply IHib2; eauto. - (* cond *) - intros rk k CONTh CONT hrs sis lsis sfv st REF EXEC OUT OK. subst. + simpl; intros; wlp_simplify. assert (rOK: ris_ok ctx hrs). { erewrite <- OK_EQUIV. 2: eauto. eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. -- cgit From c9cb80f008c919d543212dc69b5fbcd7a0d73df8 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 8 Jul 2021 11:17:45 +0200 Subject: hrexec_correct1 --- scheduling/BTL_SEimpl.v | 73 ++++++++++++++++++++++++++++++++----------------- 1 file changed, 48 insertions(+), 25 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 25819ffa..b04c17fb 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -657,7 +657,8 @@ Lemma cbranch_expanse_correct hrs c l: (REF : ris_refines ctx hrs sis) (OK: ris_ok ctx hrs), eval_scondition ctx (fst r) (snd r) = - eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)). + eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)) + /\ (fst r) = c. Proof. unfold cbranch_expanse. wlp_simplify; inv REF. @@ -687,6 +688,8 @@ Proof. + destruct 1; simpl in *. unfold ris_sreg_get; simpl. intros; rewrite PTree.gempty; eauto. Qed. +Local Hint Resolve hrs_init_correct: wlp. +Global Opaque hrs_init. Definition hrs_sreg_set r lr rsv (hrs: ristate): ?? ristate := DO ok_lsv <~ @@ -935,26 +938,26 @@ Lemma hrexec_rec_correct1 ctx ib: rst_refines ctx rst st. Proof. induction ib; try (wlp_simplify; fail). - - wlp_simplify; eapply CONT; eauto. - - destruct trap; wlp_simplify. - - (* seq *) + - (* Bnop *) wlp_simplify; eapply CONT; eauto. + - (* Bload *) destruct trap; wlp_simplify. + - (* Bseq *) wlp_simplify. eapply IHib1. 3-7: eauto. + simpl. eapply sexec_rec_okpreserv; eauto. + intros; subst. eapply IHib2; eauto. - - (* cond *) + - (* Bcond *) simpl; intros; wlp_simplify. assert (rOK: ris_ok ctx hrs). { erewrite <- OK_EQUIV. 2: eauto. eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. } - generalize OUT; clear OUT; simpl. - autodestruct. - intros COND; generalize COND. - erewrite <- eval_scondition_refpreserv; eauto. - econstructor; try_simplify_someHyps. - Qed.*) -Admitted. + exploit (Refcond ctx cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 + (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k)); + exploit H; eauto; intros (SEVAL & EQC); subst; auto. + all: simpl in SOUT; generalize SOUT; inversion_SOME b0; try_simplify_someHyps. + * intros; eapply IHib1; eauto. + * intros; eapply IHib2; eauto. +Qed. Lemma hrexec_correct1 ctx ib sis sfv: get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) -> @@ -963,14 +966,7 @@ Lemma hrexec_correct1 ctx ib sis sfv: rst_refines ctx rst (sexec (cf ctx) ib). Proof. unfold sexec; intros; wlp_simplify. - eapply hrexec_rec_correct1; eauto; simpl; try congruence; split. - - intuition eauto. - + inv H2; constructor; auto; simpl; congruence. - + inv H2; split; simpl in *; try congruence; trivial. - - intros X; simpl in *; auto. - - intros X r; inv X; simpl in *. - unfold ris_sreg_get; simpl; rewrite PTree.gempty; simpl. - reflexivity. + eapply hrexec_rec_correct1; eauto; simpl; congruence. Qed. (*Lemma hrexec_rec_correct2 ctx ib: @@ -999,7 +995,7 @@ Lemma hrexec_correct2 ctx ib hrs rfv: Proof. Admitted. -Lemma hrexec_simu_correct ctx ib: +Lemma hrexec_correct_aux ctx ib: WHEN hrexec (cf ctx) ib ~> rst THEN exists st, sexec (cf ctx) ib = st /\ rst_refines ctx rst st. Proof. @@ -1007,8 +1003,34 @@ Proof. repeat eexists. Admitted. +Global Opaque hrexec. + End CanonBuilding. +(** Correction of concrete symbolic execution wrt abstract symbolic execution *) +Theorem hrexec_correct + (hC_hsval : hashinfo sval -> ?? sval) + (hC_list_hsval : hashinfo list_sval -> ?? list_sval) + (hC_hsmem : hashinfo smem -> ?? smem) + (ctx: iblock_exec_context) + (ib: iblock): + WHEN hrexec hC_hsval hC_list_hsval hC_hsmem (cf ctx) ib ~> hrs THEN forall + (hC_hsval_correct: forall hs, + WHEN hC_hsval hs ~> hs' THEN forall ctx, + sval_refines ctx (hdata hs) hs') + (hC_list_hsval_correct: forall lh, + WHEN hC_list_hsval lh ~> lh' THEN forall ctx, + list_sval_refines ctx (hdata lh) lh') + (hC_hsmem_correct: forall hm, + WHEN hC_hsmem hm ~> hm' THEN forall ctx, + smem_refines ctx (hdata hm) hm'), + exists st : sstate, sexec (cf ctx) ib = st /\ rst_refines ctx hrs st. +Proof. + wlp_simplify. + eapply hrexec_correct_aux; eauto. +Qed. +Local Hint Resolve hrexec_correct: wlp. + (** * Implementing the simulation test with concrete hash-consed symbolic execution *) Definition phys_check {A} (x y:A) (msg: pstring): ?? unit := @@ -1233,10 +1255,10 @@ Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? unit := DO hC_sval <~ hCons hSVAL;; DO hC_list_hsval <~ hCons hLSVAL;; DO hC_hsmem <~ hCons hSMEM;; - let hsexec := hrexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in + let hrexec := hrexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in (* performing the hash-consed executions *) - DO hst1 <~ hsexec f1 ib1;; - DO hst2 <~ hsexec f2 ib2;; + DO hst1 <~ hrexec f1 ib1;; + DO hst2 <~ hrexec f2 ib2;; (* comparing the executions *) rst_simu_check hst1 hst2. @@ -1245,9 +1267,10 @@ Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): symbolic_simu f1 f2 ib1 ib2. Proof. (* TODO *) unfold symbolic_simu; wlp_simplify. + exploit hrexec_correct; eauto. eapply rst_simu_correct; eauto. + intros; eapply hrexec_correct1; eauto. - 1-3: wlp_simplify. unfold hrexec. + 1-3: wlp_simplify. intuition eauto. simpl. eauto. intuition eauto. admit. -- cgit From f59ff208a301bf3f04aab350d9f0b3b217ddeac3 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 19 Jul 2021 14:20:53 +0200 Subject: Finish implem proof, need to adapt scheduler proof --- scheduling/BTL_SEimpl.v | 424 +++++++++++++++++++++++++++++----------- scheduling/BTL_Scheduler.v | 57 +++++- scheduling/BTL_Schedulerproof.v | 3 +- 3 files changed, 371 insertions(+), 113 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index b04c17fb..34f8eb1f 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -77,6 +77,30 @@ Definition smem_set_hid (sm: smem) (hid: hashcode): smem := | Sstore sm chunk addr lsv srce _ => Sstore sm chunk addr lsv srce hid end. +Lemma sval_set_hid_correct ctx x y: + sval_set_hid x unknown_hid = sval_set_hid y unknown_hid -> + sval_refines ctx x y. +Proof. + destruct x, y; intro H; inversion H; subst; simpl; auto. +Qed. +Local Hint Resolve sval_set_hid_correct: core. + +Lemma list_sval_set_hid_correct ctx x y: + list_sval_set_hid x unknown_hid = list_sval_set_hid y unknown_hid -> + list_sval_refines ctx x y. +Proof. + destruct x, y; intro H; inversion H; subst; simpl; auto. +Qed. +Local Hint Resolve list_sval_set_hid_correct: core. + +Lemma smem_set_hid_correct ctx x y: + smem_set_hid x unknown_hid = smem_set_hid y unknown_hid -> + smem_refines ctx x y. +Proof. + destruct x, y; intro H; inversion H; subst; simpl; auto. +Qed. +Local Hint Resolve smem_set_hid_correct: core. + (** Now, we build the hash-Cons value from a "hash_eq". Informal specification: @@ -474,7 +498,6 @@ Proof. - intros X; inversion REF. erewrite H1; eauto. Qed. -Global Opaque hrexec_store. Local Hint Resolve hrexec_store_correct: wlp. (** ** Assignment of registers *) @@ -788,7 +811,34 @@ Proof. rewrite ris_sreg_set_access. erewrite red_PTree_set_refines; intuition eauto. Qed. -Global Opaque hrs_sreg_set. + +Lemma hrs_ok_op_okpreserv ctx op args dest hrs: + WHEN hrs_sreg_set dest args (Rop op) hrs ~> rst THEN + ris_ok ctx rst -> ris_ok ctx hrs. +Proof. + wlp_simplify; inv H2 || inv H1; + simpl in *; econstructor; eauto. +Qed. + +Lemma hrs_ok_load_okpreserv ctx chunk addr args dest hrs: + WHEN hrs_sreg_set dest args (Rload TRAP chunk addr) hrs ~> rst THEN + ris_ok ctx rst -> ris_ok ctx hrs. +Proof. + wlp_simplify; inv H1; + simpl in *; econstructor; eauto. +Qed. + +Lemma hrs_ok_store_okpreserv ctx chunk addr args src hrs: + WHEN hrexec_store chunk addr args src hrs ~> rst THEN + ris_ok ctx rst -> ris_ok ctx hrs. +Proof. + unfold hrexec_store; wlp_simplify. + generalize H2. erewrite ok_rset_mem; intuition. + specialize H1 with (sm2 := hrs). + erewrite H1; eauto. rewrite H3. repeat autodestruct. +Qed. + +Global Opaque hrs_sreg_set hrexec_store. Local Hint Resolve hrs_sreg_set_correct: wlp. (* transfer *) @@ -800,6 +850,7 @@ Proof. split; destruct 1; econstructor; simpl in *; eauto. Qed. +(* TODO useless? Lemma hrset_red_reg_correct ctx r rsv sv hrs sis: ris_refines ctx hrs sis -> (ris_ok ctx hrs -> eval_sval ctx rsv <> None) -> @@ -814,50 +865,62 @@ Proof. erewrite red_PTree_set_refines; eauto. econstructor; eauto. Qed. + *) -Fixpoint transfer_hrs (inputs: list reg) (hrs: ristate): ristate := +Fixpoint transfer_hrs (inputs: list reg) (hrs: ristate): ?? ristate := match inputs with - | nil => {| ris_smem := hrs.(ris_smem); - ris_input_init := false; - ok_rsval := hrs.(ok_rsval); - ris_sreg:= PTree.empty _ - |} + | nil => RET {| ris_smem := hrs.(ris_smem); + ris_input_init := false; + ok_rsval := hrs.(ok_rsval); + ris_sreg:= PTree.empty _ + |} | r::l => - let hrs' := transfer_hrs l hrs in - ris_sreg_set hrs' (red_PTree_set r (ris_sreg_get hrs r) hrs') + DO hrs' <~ transfer_hrs l hrs;; + DO sv <~ hrs_sreg_get hrs r;; + RET (ris_sreg_set hrs' (red_PTree_set r sv hrs')) end. Lemma ok_transfer_hrs ctx inputs hrs: - ris_ok ctx (transfer_hrs inputs hrs) + WHEN transfer_hrs inputs hrs ~> hrs' THEN + ris_ok ctx hrs' <-> (ris_ok ctx hrs). Proof. - induction inputs as [|r l]; simpl. - + split; destruct 1; econstructor; simpl in *; intuition eauto. - + rewrite ok_hrset_red_sreg. auto. + induction inputs as [|r l]; simpl; wlp_simplify; + try (inv H; econstructor; eauto; fail). + + apply H0; rewrite <- ok_hrset_red_sreg; eauto. + + rewrite ok_hrset_red_sreg; auto. Qed. +Local Hint Resolve ok_transfer_hrs: wlp. Lemma transfer_hrs_correct ctx inputs hrs sis: + WHEN transfer_hrs inputs hrs ~> hrs' THEN ris_refines ctx hrs sis -> - ris_refines ctx (transfer_hrs inputs hrs) (transfer_sis inputs sis). + ris_refines ctx hrs' (transfer_sis inputs sis). Proof. - destruct 1; intros. - induction inputs as [|r l]. - + econstructor; eauto. - * erewrite ok_transfer_sis, ok_transfer_hrs; eauto. - * erewrite ok_transfer_hrs; eauto. - * erewrite ok_transfer_hrs; simpl; unfold ris_sreg_get; simpl; eauto. - intros; rewrite PTree.gempty. simpl; auto. - + econstructor; eauto. - * erewrite ok_transfer_sis, ok_transfer_hrs; eauto. - * erewrite ok_transfer_hrs; simpl. - intros; inversion IHl. erewrite MEM_EQ0. - - unfold transfer_sis; simpl; eauto. - - rewrite ok_transfer_hrs; simpl; eauto. - * erewrite ok_transfer_hrs; simpl. - intros H r0; inversion IHl. + unfold transfer_hrs; + induction inputs as [|r l]; wlp_simplify. + inversion H. + + constructor. + * erewrite ok_transfer_sis; split; intros X. + - constructor; simpl; rewrite OK_EQUIV in X; inv X; assumption. + - rewrite OK_EQUIV; constructor; inv X; assumption. + * intros X; inv X; simpl; apply MEM_EQ; constructor; auto. + * simpl. unfold ris_sreg_get. + intros;rewrite PTree.gempty. simpl; auto. + + constructor. + * erewrite ok_hrset_red_sreg; eauto. + inv H2. rewrite <- OK_EQUIV. + erewrite !ok_transfer_sis; reflexivity. + * intros X; inversion X; simpl in *. + rewrite ok_hrset_red_sreg in X. + inv H2; rewrite MEM_EQ; auto. + * simpl; intros X r0. inversion H2. + rewrite ok_hrset_red_sreg in X. erewrite red_PTree_set_refines; eauto. - - destruct (Pos.eq_dec); eauto. - - rewrite ok_transfer_hrs. auto. + destruct (Pos.eq_dec); auto. + inv H1. rewrite ok_transfer_sis in OK_EQUIV. + rewrite <- OK_EQUIV, OK_EQUIV0 in X. + erewrite H0; eauto. Qed. Definition tr_hrs := poly_tr (fun f tbl or => transfer_hrs (Regset.elements (pre_inputs f tbl or))). @@ -866,28 +929,140 @@ Local Hint Resolve transfer_hrs_correct ok_transfer_hrs: core. Local Opaque transfer_hrs. Lemma ok_tr_hrs ctx fi hrs: - ris_ok ctx (tr_hrs (cf ctx) fi hrs) + WHEN tr_hrs (cf ctx) fi hrs ~> hrs' THEN + ris_ok ctx hrs' <-> (ris_ok ctx hrs). Proof. destruct fi; simpl; eauto. Qed. -Lemma hrs_ok_tr_okpreserv ctx fi hrs: ris_ok ctx (tr_hrs (cf ctx) fi hrs) -> ris_ok ctx hrs. +Lemma tr_hrs_correct ctx fi hrs sis: + WHEN tr_hrs (cf ctx) fi hrs ~> hrs' THEN + ris_refines ctx hrs sis -> + ris_refines ctx hrs' (tr_sis (cf ctx) fi sis). Proof. - rewrite ok_tr_hrs; auto. + unfold tr_hrs, poly_tr; destruct fi; wlp_simplify; + exploit transfer_hrs_correct; eauto. + Unshelve. all: auto. Qed. -Lemma tr_hrs_correct ctx fi hrs sis: - ris_refines ctx hrs sis -> - ris_refines ctx (tr_hrs (cf ctx) fi hrs) (tr_sis (cf ctx) fi sis). +Fixpoint hbuiltin_arg (hrs: ristate) (arg : builtin_arg reg): ?? builtin_arg sval := + match arg with + | BA r => + DO v <~ hrs_sreg_get hrs r;; + RET (BA v) + | BA_int n => RET (BA_int n) + | BA_long n => RET (BA_long n) + | BA_float f0 => RET (BA_float f0) + | BA_single s => RET (BA_single s) + | BA_loadstack chunk ptr => RET (BA_loadstack chunk ptr) + | BA_addrstack ptr => RET (BA_addrstack ptr) + | BA_loadglobal chunk id ptr => RET (BA_loadglobal chunk id ptr) + | BA_addrglobal id ptr => RET (BA_addrglobal id ptr) + | BA_splitlong ba1 ba2 => + DO v1 <~ hbuiltin_arg hrs ba1;; + DO v2 <~ hbuiltin_arg hrs ba2;; + RET (BA_splitlong v1 v2) + | BA_addptr ba1 ba2 => + DO v1 <~ hbuiltin_arg hrs ba1;; + DO v2 <~ hbuiltin_arg hrs ba2;; + RET (BA_addptr v1 v2) + end. + +Lemma hbuiltin_arg_correct hrs arg: + WHEN hbuiltin_arg hrs arg ~> hargs THEN forall ctx (f: reg -> sval) + (RR: forall r, sval_refines ctx (hrs r) (f r)), + eval_builtin_sval ctx hargs = eval_builtin_sval ctx (builtin_arg_map f arg). Proof. - intros REF. rewrite <- tr_sis_alt_def. - destruct fi; simpl; eauto. + induction arg; wlp_simplify. + + erewrite H; eauto. + + erewrite H; eauto. + erewrite H0; eauto. + + erewrite H; eauto. + erewrite H0; eauto. +Qed. +Global Opaque hbuiltin_arg. +Local Hint Resolve hbuiltin_arg_correct: wlp. + +Fixpoint hbuiltin_args hrs (args: list (builtin_arg reg)): ?? list (builtin_arg sval) := + match args with + | nil => RET nil + | a::l => + DO ha <~ hbuiltin_arg hrs a;; + DO hl <~ hbuiltin_args hrs l;; + RET (ha::hl) + end. + +Lemma hbuiltin_args_correct hrs args: + WHEN hbuiltin_args hrs args ~> hargs THEN forall ctx (f: reg -> sval) + (RR: forall r, sval_refines ctx (hrs r) (f r)), + bargs_refines ctx hargs (List.map (builtin_arg_map f) args). +Proof. + unfold bargs_refines; induction args; wlp_simplify. + erewrite H; eauto. + erewrite H0; eauto. +Qed. +Global Opaque hbuiltin_args. +Local Hint Resolve hbuiltin_args_correct: wlp. + +Definition hsum_left hrs (ros: reg + ident): ?? (sval + ident) := + match ros with + | inl r => DO hr <~ hrs_sreg_get hrs r;; RET (inl hr) + | inr s => RET (inr s) + end. + +Lemma hsum_left_correct hrs ros: + WHEN hsum_left hrs ros ~> hsi THEN forall ctx (f: reg -> sval) + (RR: forall r, sval_refines ctx (hrs r) (f r)), + rsvident_refines ctx hsi (sum_left_map f ros). +Proof. + destruct ros; wlp_simplify; + econstructor; eauto. +Qed. +Global Opaque hsum_left. +Local Hint Resolve hsum_left_correct: wlp. + +(** * Symbolic execution of final step *) +Definition hrexec_final_sfv (i: final) hrs: ?? sfval := + match i with + | Bgoto pc => RET (Sgoto pc) + | Bcall sig ros args res pc => + DO svos <~ hsum_left hrs ros;; + DO sargs <~ hlist_args hrs args;; + RET (Scall sig svos sargs res pc) + | Btailcall sig ros args => + DO svos <~ hsum_left hrs ros;; + DO sargs <~ hlist_args hrs args;; + RET (Stailcall sig svos sargs) + | Bbuiltin ef args res pc => + DO sargs <~ hbuiltin_args hrs args;; + RET (Sbuiltin ef sargs res pc) + | Breturn or => + match or with + | Some r => DO hr <~ hrs_sreg_get hrs r;; RET (Sreturn (Some hr)) + | None => RET (Sreturn None) + end + | Bjumptable reg tbl => + DO sv <~ hrs_sreg_get hrs reg;; + RET (Sjumptable sv tbl) + end. + +Lemma hrexec_final_sfv_correct ctx fi hrs sis: + WHEN hrexec_final_sfv fi hrs ~> hrs' THEN forall + (REF: ris_refines ctx hrs sis) + (OK: ris_ok ctx hrs), + rfv_refines ctx hrs' (sexec_final_sfv fi sis). +Proof. + destruct fi; simpl; wlp_simplify; repeat econstructor; + try inversion REF; try erewrite H; eauto. Qed. Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := match ib with - | BF fin _ => RET (Rfinal (tr_hrs f fin hrs) (sexec_final_sfv fin hrs)) + | BF fin _ => + DO sfv <~ hrexec_final_sfv fin hrs;; + DO hrs' <~ tr_hrs f fin hrs;; + RET (Rfinal hrs' sfv) (* basic instructions *) | Bnop _ => k hrs | Bop op args dst _ => @@ -916,7 +1091,7 @@ Definition hrexec f ib := DO init <~ hrs_init;; hrexec_rec f ib init (fun _ => RET Rabort). -Local Hint Resolve exec_final_refpreserv tr_hrs_correct hrs_ok_tr_okpreserv: core. +Local Hint Resolve exec_final_refpreserv ok_tr_hrs: core. Local Hint Constructors rst_refines: core. Lemma hrexec_rec_correct1 ctx ib: @@ -938,13 +1113,18 @@ Lemma hrexec_rec_correct1 ctx ib: rst_refines ctx rst st. Proof. induction ib; try (wlp_simplify; fail). + - (* BF *) + wlp_simplify; exploit tr_hrs_correct; eauto. + intros; constructor; auto. + intros X; apply H0 in X. + exploit hrexec_final_sfv_correct; eauto. - (* Bnop *) wlp_simplify; eapply CONT; eauto. - (* Bload *) destruct trap; wlp_simplify. - (* Bseq *) wlp_simplify. eapply IHib1. 3-7: eauto. + simpl. eapply sexec_rec_okpreserv; eauto. - + intros; subst. eapply IHib2; eauto. + + intros; eapply IHib2; eauto. - (* Bcond *) simpl; intros; wlp_simplify. assert (rOK: ris_ok ctx hrs). { @@ -954,9 +1134,11 @@ Proof. exploit (Refcond ctx cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k)); exploit H; eauto; intros (SEVAL & EQC); subst; auto. - all: simpl in SOUT; generalize SOUT; inversion_SOME b0; try_simplify_someHyps. - * intros; eapply IHib1; eauto. - * intros; eapply IHib2; eauto. + all: + simpl in SOUT; generalize SOUT; clear SOUT; + inversion_SOME b0; try_simplify_someHyps. + + intros; eapply IHib1; eauto. + + intros; eapply IHib2; eauto. Qed. Lemma hrexec_correct1 ctx ib sis sfv: @@ -969,23 +1151,80 @@ Proof. eapply hrexec_rec_correct1; eauto; simpl; congruence. Qed. -(*Lemma hrexec_rec_correct2 ctx ib: +Lemma hrexec_rec_okpreserv ctx ib: + forall k + (CONT: forall hrs lhrs rfv, + WHEN k hrs ~> rst THEN + get_routcome ctx rst = Some (rout lhrs rfv) -> + ris_ok ctx lhrs -> + ris_ok ctx hrs) + hrs lhrs rfv, + WHEN hrexec_rec (cf ctx) ib hrs k ~> rst THEN + get_routcome ctx rst = Some (rout lhrs rfv) -> + ris_ok ctx lhrs -> + ris_ok ctx hrs. +Proof. + induction ib; simpl; try (wlp_simplify; try_simplify_someHyps; fail). + - (* Bop *) + wlp_simplify; exploit hrs_ok_op_okpreserv; eauto. + - (* Bload *) + destruct trap; wlp_simplify; try_simplify_someHyps. + exploit hrs_ok_load_okpreserv; eauto. + - (* Bstore *) + wlp_simplify; exploit hrs_ok_store_okpreserv; eauto. + - (* Bcond *) + wlp_simplify; generalize H2; clear H2; inversion_SOME b; destruct b; try_simplify_someHyps. +Qed. +Local Hint Resolve hrexec_rec_okpreserv: wlp. + +Lemma hrexec_rec_correct2 ctx ib: forall rk k - (CONTh: forall ris lris rfv, get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris) - (CONT: forall ris sis lris rfv st, - ris_refines ctx ris sis -> - rk ris = st -> - get_routcome ctx (rk ris) = Some (rout lris rfv) -> - ris_ok ctx lris -> - rst_refines ctx (rk ris) (k sis)) - ris sis lris rfv st - (REF: ris_refines ctx ris sis) - (EXEC: rexec_rec (cf ctx) ib ris rk = st) - (SOUT: get_routcome ctx st = Some (rout lris rfv)) - (OK: ris_ok ctx lris) - , rst_refines ctx st (sexec_rec (cf ctx) ib sis k). + (CONTh: forall hrs lhrs rfv, + WHEN rk hrs ~> rst THEN + get_routcome ctx rst = Some (rout lhrs rfv) -> + ris_ok ctx lhrs -> + ris_ok ctx hrs) + (CONT: forall hrs sis lhrs rfv, + ris_refines ctx hrs sis -> + WHEN rk hrs ~> rst THEN + get_routcome ctx rst = Some (rout lhrs rfv) -> + ris_ok ctx lhrs -> + rst_refines ctx rst (k sis)) + hrs sis lhrs rfv + (REF: ris_refines ctx hrs sis), + WHEN hrexec_rec (cf ctx) ib hrs rk ~> rst THEN + get_routcome ctx rst = Some (rout lhrs rfv) -> + ris_ok ctx lhrs -> + rst_refines ctx rst (sexec_rec (cf ctx) ib sis k). Proof. - Admitted.*) + induction ib; try (wlp_simplify; fail). + - (* BF *) + wlp_simplify; exploit tr_hrs_correct; eauto. + intros; constructor; auto. + intros X; apply H0 in X. + exploit hrexec_final_sfv_correct; eauto. + - (* Bnop *) wlp_simplify; eapply CONT; eauto. + - (* Bload *) destruct trap; wlp_simplify. + - (* Bseq *) + wlp_simplify. + eapply IHib1. 3-6: eauto. + + simpl; eapply hrexec_rec_okpreserv; eauto. + + intros; eapply IHib2; eauto. + - (* Bcond *) + simpl; intros; wlp_simplify. + assert (rOK: ris_ok ctx hrs). { + simpl in H2; generalize H2; inversion_SOME b0; destruct b0; try_simplify_someHyps. + } + exploit (Refcond ctx cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 + (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k)); + exploit H; eauto; intros (SEVAL & EQC); subst; auto. + all: + simpl in H2; generalize H2; clear H2; + inversion_SOME b0; try_simplify_someHyps. + + intros; eapply IHib1; eauto. + + intros; eapply IHib2; eauto. + Unshelve. all: auto. +Qed. Lemma hrexec_correct2 ctx ib hrs rfv: WHEN hrexec (cf ctx) ib ~> rst THEN @@ -993,44 +1232,15 @@ Lemma hrexec_correct2 ctx ib hrs rfv: (ris_ok ctx hrs) -> rst_refines ctx rst (sexec (cf ctx) ib). Proof. -Admitted. - -Lemma hrexec_correct_aux ctx ib: - WHEN hrexec (cf ctx) ib ~> rst THEN - exists st, sexec (cf ctx) ib = st /\ rst_refines ctx rst st. -Proof. - unfold hrexec; wlp_simplify. - repeat eexists. -Admitted. + unfold hrexec; intros; wlp_simplify. + eapply hrexec_rec_correct2; eauto; simpl; + wlp_simplify; congruence. +Qed. Global Opaque hrexec. End CanonBuilding. -(** Correction of concrete symbolic execution wrt abstract symbolic execution *) -Theorem hrexec_correct - (hC_hsval : hashinfo sval -> ?? sval) - (hC_list_hsval : hashinfo list_sval -> ?? list_sval) - (hC_hsmem : hashinfo smem -> ?? smem) - (ctx: iblock_exec_context) - (ib: iblock): - WHEN hrexec hC_hsval hC_list_hsval hC_hsmem (cf ctx) ib ~> hrs THEN forall - (hC_hsval_correct: forall hs, - WHEN hC_hsval hs ~> hs' THEN forall ctx, - sval_refines ctx (hdata hs) hs') - (hC_list_hsval_correct: forall lh, - WHEN hC_list_hsval lh ~> lh' THEN forall ctx, - list_sval_refines ctx (hdata lh) lh') - (hC_hsmem_correct: forall hm, - WHEN hC_hsmem hm ~> hm' THEN forall ctx, - smem_refines ctx (hdata hm) hm'), - exists st : sstate, sexec (cf ctx) ib = st /\ rst_refines ctx hrs st. -Proof. - wlp_simplify. - eapply hrexec_correct_aux; eauto. -Qed. -Local Hint Resolve hrexec_correct: wlp. - (** * Implementing the simulation test with concrete hash-consed symbolic execution *) Definition phys_check {A} (x y:A) (msg: pstring): ?? unit := @@ -1267,34 +1477,30 @@ Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): symbolic_simu f1 f2 ib1 ib2. Proof. (* TODO *) unfold symbolic_simu; wlp_simplify. - exploit hrexec_correct; eauto. eapply rst_simu_correct; eauto. - + intros; eapply hrexec_correct1; eauto. - 1-3: wlp_simplify. intuition eauto. - simpl. eauto. - intuition eauto. - admit. -Admitted. + + intros; eapply hrexec_correct1; eauto; wlp_simplify. + + intros; eapply hrexec_correct2; eauto; wlp_simplify. +Qed. Global Opaque simu_check_single. Global Hint Resolve simu_check_single_correct: wlp. -Program Definition aux_simu_check (f1 f2: function) (ib1 ib2: iblock): ?? bool := +Program Definition aux_simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): ?? bool := DO r <~ (TRY - simu_check_single f1 f2 ib1 ib2;; + simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; RET true CATCH_FAIL s, _ => println ("simu_check_failure:" +; s);; RET false - ENSURE (fun b => b=true -> symbolic_simu f1 f2 ib1 ib2));; + ENSURE (fun b => b=true -> ibf1.(input_regs) = ibf2.(input_regs) -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry)));; RET (`r). Obligation 1. split; wlp_simplify. discriminate. Qed. -Lemma aux_simu_check_correct (f1 f2: function) (ib1 ib2: iblock): - WHEN aux_simu_check f1 f2 ib1 ib2 ~> b THEN - b=true -> symbolic_simu f1 f2 ib1 ib2. +Lemma aux_simu_check_correct (f1 f2: function) (ibf1 ibf2: iblock_info): + WHEN aux_simu_check f1 f2 ibf1 ibf2 ~> b THEN + b=true -> ibf1.(input_regs) = ibf2.(input_regs) -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). Proof. unfold aux_simu_check; wlp_simplify. destruct exta; simpl; auto. @@ -1304,8 +1510,8 @@ Qed. Import UnsafeImpure. -Definition simu_check (f1 f2: function) (ib1 ib2: iblock): res unit := - match unsafe_coerce (aux_simu_check f1 f2 ib1 ib2) with +Definition simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): res unit := + match unsafe_coerce (aux_simu_check f1 f2 ibf1 ibf2) with | Some true => OK tt | _ => Error (msg "simu_check has failed") end. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 12840f62..63114125 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -2,7 +2,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. -Require Import Errors Linking BTL_SEtheory. +Require Import Errors Linking BTL_SEtheory BTL_SEimpl. (** External oracle *) Axiom scheduler: BTL.function -> BTL.code. @@ -48,18 +48,69 @@ Record match_function (f1 f2: BTL.function): Prop := { Local Open Scope error_monad_scope. -Definition check_symbolic_simu (f tf: BTL.function): res unit := OK tt. (* TODO: fixme *) +Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res unit := + match lpc with + | nil => OK tt + | pc :: lpc' => + match (fn_code f1)!pc, (fn_code f2)!pc with + | Some ibf1, Some ibf2 => + do _ <- simu_check f1 f2 ibf1 ibf2; + check_symbolic_simu_rec f1 f2 lpc' + | _, _ => Error (msg "check_symbolic_simu_rec: code tree mismatch") + end + end. + +Definition check_symbolic_simu (f1 f2: BTL.function): res unit := + check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))). + +Lemma check_symbolic_simu_rec_input_equiv lpc: forall f1 f2 x, + check_symbolic_simu_rec f1 f2 lpc = OK x -> + lpc = (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))) -> + equiv_input_regs f1 f2. +Proof. +Admitted. +(* + unfold equiv_input_regs; induction lpc; simpl; intros f1 f2 x X EQL1. + - destruct (map _ _) eqn:EQL2; inv X; intuition; + symmetry in EQL1; apply map_eq_nil in EQL1; apply map_eq_nil in EQL2. + + destruct (fn_code f1), (fn_code f2); simpl in *; auto. + inv EQL2. + unfold PTree.elements in EQL1. + + admit. + + + destruct ((fn_code f1) ! a), ((fn_code f2) ! a); monadInv X. + eapply IHlpc; eauto. +Qed. + *) Lemma check_symbolic_simu_input_equiv x f1 f2: check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. +Proof. + unfold check_symbolic_simu; simpl; intros X. + eapply check_symbolic_simu_rec_input_equiv; eauto. +Qed. + +Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x, + check_symbolic_simu_rec f1 f2 lpc = OK x -> + 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). Proof. Admitted. + (* + induction lpc; simpl; intros f1 f2 x X pc ib1 PC; try (inv X; fail). + destruct ((fn_code f1) ! pc), ((fn_code f2) ! pc); inv X. + eapply IHlpc; eauto. +Qed. + *) Lemma check_symbolic_simu_correct x f1 f2: check_symbolic_simu f1 f2 = OK x -> 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). -Admitted. +Proof. + unfold check_symbolic_simu; intros X pc ib1 PC. + eapply check_symbolic_simu_rec_correct; eauto. +Qed. Definition transf_function (f: BTL.function) := let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 95363194..91e650e8 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -105,11 +105,12 @@ Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. intros TRANSF; destruct f; simpl; inv TRANSF. +Admitted. (* + exploit transf_function_correct; eauto. econstructor. intros MATCH_F. eapply match_Internal; eauto. + eapply match_External. -Qed. + Qed.*) Lemma function_ptr_preserved: forall v f, -- cgit From a9c90fe3354f65340283dc79431bc6915ed1ad90 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 20 Jul 2021 10:17:46 +0200 Subject: proof for symbolic simu, need to finish equiv_inputs --- scheduling/BTL_SEimpl.v | 29 +++++++++++++++++++---------- scheduling/BTL_Scheduler.v | 17 +++++++++-------- scheduling/BTL_Schedulerproof.v | 8 +++----- 3 files changed, 31 insertions(+), 23 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 34f8eb1f..d6cdff59 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -1459,22 +1459,21 @@ Qed. Hint Resolve rst_simu_check_correct: wlp. Global Opaque rst_simu_check. -(* TODO do we really need both functions *) -Definition simu_check_single (f1 f2: function) (ib1 ib2: iblock): ?? unit := +Definition simu_check_single (f1 f2: function) (ibf1 ibf2: iblock_info): ?? unit := (* creating the hash-consing tables *) DO hC_sval <~ hCons hSVAL;; DO hC_list_hsval <~ hCons hLSVAL;; DO hC_hsmem <~ hCons hSMEM;; let hrexec := hrexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in (* performing the hash-consed executions *) - DO hst1 <~ hrexec f1 ib1;; - DO hst2 <~ hrexec f2 ib2;; + DO hst1 <~ hrexec f1 ibf1.(entry);; + DO hst2 <~ hrexec f2 ibf2.(entry);; (* comparing the executions *) rst_simu_check hst1 hst2. -Lemma simu_check_single_correct (f1 f2: function) (ib1 ib2: iblock): - WHEN simu_check_single f1 f2 ib1 ib2 ~> _ THEN - symbolic_simu f1 f2 ib1 ib2. +Lemma simu_check_single_correct (f1 f2: function) (ibf1 ibf2: iblock_info): + WHEN simu_check_single f1 f2 ibf1 ibf2 ~> _ THEN + symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). Proof. (* TODO *) unfold symbolic_simu; wlp_simplify. eapply rst_simu_correct; eauto. @@ -1487,12 +1486,12 @@ Global Hint Resolve simu_check_single_correct: wlp. Program Definition aux_simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): ?? bool := DO r <~ (TRY - simu_check_single f1 f2 ibf1.(entry) ibf2.(entry);; + simu_check_single f1 f2 ibf1 ibf2;; RET true CATCH_FAIL s, _ => println ("simu_check_failure:" +; s);; RET false - ENSURE (fun b => b=true -> ibf1.(input_regs) = ibf2.(input_regs) -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry)));; + ENSURE (fun b => b=true -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry)));; RET (`r). Obligation 1. split; wlp_simplify. discriminate. @@ -1500,7 +1499,7 @@ Qed. Lemma aux_simu_check_correct (f1 f2: function) (ibf1 ibf2: iblock_info): WHEN aux_simu_check f1 f2 ibf1 ibf2 ~> b THEN - b=true -> ibf1.(input_regs) = ibf2.(input_regs) -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). + b=true -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). Proof. unfold aux_simu_check; wlp_simplify. destruct exta; simpl; auto. @@ -1515,3 +1514,13 @@ Definition simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): res unit := | Some true => OK tt | _ => Error (msg "simu_check has failed") end. + +Lemma simu_check_correct f1 f2 ibf1 ibf2: + simu_check f1 f2 ibf1 ibf2 = OK tt -> + symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). +Proof. + unfold simu_check. + destruct (unsafe_coerce (aux_simu_check f1 f2 ibf1 ibf2)) as [[|]|] eqn:Hres; simpl; try discriminate. + intros; eapply aux_simu_check_correct; eauto. + eapply unsafe_coerce_not_really_correct; eauto. +Qed. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 63114125..e7d2b37f 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -92,16 +92,16 @@ Qed. Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x, check_symbolic_simu_rec f1 f2 lpc = OK x -> - forall pc ib1, (fn_code f1)!pc = Some ib1 -> + forall pc ib1, (fn_code f1)!pc = Some ib1 /\ In pc lpc -> exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2). Proof. -Admitted. - (* - induction lpc; simpl; intros f1 f2 x X pc ib1 PC; try (inv X; fail). - destruct ((fn_code f1) ! pc), ((fn_code f2) ! pc); inv X. - eapply IHlpc; eauto. + induction lpc; simpl; intros f1 f2 x X pc ib1 (PC & HIN); try contradiction. + destruct HIN; subst. + - rewrite PC in X; destruct ((fn_code f2)!pc); monadInv X. + exists i; split; auto. destruct x0. eapply simu_check_correct; eauto. + - destruct ((fn_code f1)!a), ((fn_code f2)!a); monadInv X. + eapply IHlpc; eauto. Qed. - *) Lemma check_symbolic_simu_correct x f1 f2: check_symbolic_simu f1 f2 = OK x -> @@ -109,7 +109,8 @@ Lemma check_symbolic_simu_correct x f1 f2: exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2). Proof. unfold check_symbolic_simu; intros X pc ib1 PC. - eapply check_symbolic_simu_rec_correct; eauto. + eapply check_symbolic_simu_rec_correct; intuition eauto. + apply PTree.elements_correct in PC. eapply (in_map fst) in PC; auto. Qed. Definition transf_function (f: BTL.function) := diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 91e650e8..05028f11 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -105,12 +105,10 @@ Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. intros TRANSF; destruct f; simpl; inv TRANSF. -Admitted. (* - + exploit transf_function_correct; eauto. - econstructor. intros MATCH_F. - eapply match_Internal; eauto. + + monadInv H0. exploit transf_function_correct; eauto. + constructor; auto. + eapply match_External. - Qed.*) +Qed. Lemma function_ptr_preserved: forall v f, -- cgit From c3ce32da7d431069ef355296bef66b112a302b78 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 20 Jul 2021 12:32:21 +0200 Subject: op simplify BTL intro --- scheduling/BTL_SEimpl.v | 16 +++++++--------- 1 file changed, 7 insertions(+), 9 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index d6cdff59..3050f2e2 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -2,6 +2,7 @@ Require Import Coqlib AST Maps. Require Import Op Registers. Require Import RTL BTL Errors. Require Import BTL_SEsimuref BTL_SEtheory OptionMonad. +Require Import BTL_SEsimplify. Require Import Impure.ImpHCons. Import Notations. @@ -587,14 +588,12 @@ Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval := match is_move_operation op lr with | Some arg => hrs_sreg_get hrs arg (* optimization of Omove *) | None => - DO lsv <~ hlist_args hrs lr;; - hSop op lsv - (* TODO gourdinl match target_op_simplify op lr hrs with | Some fsv => fsval_proj fsv | None => - hSop op lhsv - end*) + DO lhsv <~ hlist_args hrs lr;; + hSop op lhsv + end end | Rload _ chunk addr => DO lsv <~ hlist_args hrs lr;; @@ -665,15 +664,14 @@ Proof. Qed. Definition cbranch_expanse (prev: ristate) (cond: condition) (args: list reg): ?? (condition * list_sval) := - (* TODO gourdinl match target_cbranch_expanse prev cond args with | Some (cond', vargs) => DO vargs' <~ fsval_list_proj vargs;; RET (cond', vargs') - | None =>*) + | None => DO vargs <~ hlist_args prev args ;; - RET (cond, vargs). - (*end.*) + RET (cond, vargs) + end. Lemma cbranch_expanse_correct hrs c l: WHEN cbranch_expanse hrs c l ~> r THEN forall ctx sis -- cgit From a3319eb05543930844dedd9ac31ed1beaac3047e Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 20 Jul 2021 15:21:29 +0200 Subject: Fix compile on ARM/x86 backends --- scheduling/BTLScheduleraux.ml | 195 +----------------------------------------- scheduling/BTL_Scheduler.v | 42 +++++---- scheduling/BTLtoRTLproof.v | 8 +- 3 files changed, 28 insertions(+), 217 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index b87636e1..98bc4590 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -1,199 +1,11 @@ -open AST open Maps -open Registers open BTL open BTLtypes open DebugPrint open PrintBTL open RTLcommonaux -open InstructionScheduler -open PrepassSchedulingOracleDeps - -let use_alias_analysis () = false - -let build_constraints_and_resources (opweights : opweights) insts btl = - let last_reg_reads : int list PTree.t ref = ref PTree.empty - and last_reg_write : (int * int) PTree.t ref = ref PTree.empty - and last_mem_reads : int list ref = ref [] - and last_mem_write : int option ref = ref None - and last_branch : int option ref = ref None - and last_non_pipelined_op : int array = - Array.make opweights.nr_non_pipelined_units (-1) - and latency_constraints : latency_constraint list ref = ref [] - and resources = ref [] in - let add_constraint instr_from instr_to latency = - assert (instr_from <= instr_to); - assert (latency >= 0); - if instr_from = instr_to then - if latency = 0 then () - else - failwith "PrepassSchedulingOracle.get_dependencies: negative self-loop" - else - latency_constraints := - { instr_from; instr_to; latency } :: !latency_constraints - and get_last_reads reg = - match PTree.get reg !last_reg_reads with Some l -> l | None -> [] - in - let add_input_mem i = - if not (use_alias_analysis ()) then ( - (* Read after write *) - (match !last_mem_write with None -> () | Some j -> add_constraint j i 1); - last_mem_reads := i :: !last_mem_reads) - and add_output_mem i = - if not (use_alias_analysis ()) then ( - (* Write after write *) - (match !last_mem_write with None -> () | Some j -> add_constraint j i 1); - (* Write after read *) - List.iter (fun j -> add_constraint j i 0) !last_mem_reads; - last_mem_write := Some i; - last_mem_reads := []) - and add_input_reg i reg = - (* Read after write *) - (match PTree.get reg !last_reg_write with - | None -> () - | Some (j, latency) -> add_constraint j i latency); - last_reg_reads := PTree.set reg (i :: get_last_reads reg) !last_reg_reads - and add_output_reg i latency reg = - (* Write after write *) - (match PTree.get reg !last_reg_write with - | None -> () - | Some (j, _) -> add_constraint j i 1); - (* Write after read *) - List.iter (fun j -> add_constraint j i 0) (get_last_reads reg); - last_reg_write := PTree.set reg (i, latency) !last_reg_write; - last_reg_reads := PTree.remove reg !last_reg_reads - in - let add_input_regs i regs = List.iter (add_input_reg i) regs - and irreversible_action i = - match !last_branch with None -> () | Some j -> add_constraint j i 1 - in - let set_branch i = - irreversible_action i; - last_branch := Some i - and add_non_pipelined_resources i resources = - Array.iter2 - (fun latency last -> - if latency >= 0 && last >= 0 then add_constraint last i latency) - resources last_non_pipelined_op; - Array.iteri - (fun rsc latency -> if latency >= 0 then last_non_pipelined_op.(rsc) <- i) - resources - in - Array.iteri - (fun i inst -> - match inst with - | Bnop _ -> - let rs = Array.map (fun _ -> 0) opweights.pipelined_resource_bounds in - resources := rs :: !resources - | Bop (op, lr, rd, _) -> - add_non_pipelined_resources i - (opweights.non_pipelined_resources_of_op op (List.length lr)); - if Op.is_trapping_op op then irreversible_action i; - add_input_regs i lr; - add_output_reg i (opweights.latency_of_op op (List.length lr)) rd; - let rs = opweights.resources_of_op op (List.length lr) in - resources := rs :: !resources - | Bload (trap, chk, addr, lr, rd, _) -> - if trap = TRAP then irreversible_action i; - add_input_mem i; - add_input_regs i lr; - add_output_reg i - (opweights.latency_of_load trap chk addr (List.length lr)) - rd; - let rs = opweights.resources_of_load trap chk addr (List.length lr) in - resources := rs :: !resources - | Bstore (chk, addr, lr, src, _) -> - irreversible_action i; - add_input_regs i lr; - add_input_reg i src; - add_output_mem i; - let rs = opweights.resources_of_store chk addr (List.length lr) in - resources := rs :: !resources - | Bcond (cond, lr, BF (Bgoto s, _), ibnot, _) -> - (* TODO gourdinl test with/out this line *) - let live = (get_some @@ PTree.get s btl).input_regs in - add_input_regs i (Regset.elements live); - set_branch i; - add_input_mem i; - add_input_regs i lr; - let rs = opweights.resources_of_cond cond (List.length lr) in - resources := rs :: !resources - | Bcond (_, _, _, _, _) -> - failwith "get_simple_dependencies: invalid Bcond" - | BF (_, _) -> failwith "get_simple_dependencies: BF" - | Bseq (_, _) -> failwith "get_simple_dependencies: Bseq") - insts; - (!latency_constraints, Array.of_list (List.rev !resources)) - -let define_problem (opweights : opweights) ibf btl = - let simple_deps, resources = - build_constraints_and_resources opweights ibf btl - in - { - max_latency = -1; - resource_bounds = opweights.pipelined_resource_bounds; - instruction_usages = resources; - latency_constraints = simple_deps; - } - -let zigzag_scheduler problem early_ones = - let nr_instructions = get_nr_instructions problem in - assert (nr_instructions = Array.length early_ones); - match list_scheduler problem with - | Some fwd_schedule -> - let fwd_makespan = fwd_schedule.(Array.length fwd_schedule - 1) in - let constraints' = ref problem.latency_constraints in - Array.iteri - (fun i is_early -> - if is_early then - constraints' := - { - instr_from = i; - instr_to = nr_instructions; - latency = fwd_makespan - fwd_schedule.(i); - } - :: !constraints') - early_ones; - validated_scheduler reverse_list_scheduler - { problem with latency_constraints = !constraints' } - | None -> None - -let prepass_scheduler_by_name name problem insts = - match name with - | "zigzag" -> - let early_ones = - Array.map - (fun inst -> - match inst with Bcond (_, _, _, _, _) -> true | _ -> false) - insts - in - zigzag_scheduler problem early_ones - | _ -> scheduler_by_name name problem - -let schedule_sequence insts btl = - let opweights = OpWeights.get_opweights () in - try - if Array.length insts <= 1 then None - else - let nr_instructions = Array.length insts in - let problem = define_problem opweights insts btl in - match - prepass_scheduler_by_name !Clflags.option_fprepass_sched problem insts - with - | None -> - Printf.printf "no solution in prepass scheduling\n"; - None - | Some solution -> - let positions = Array.init nr_instructions (fun i -> i) in - Array.sort - (fun i j -> - let si = solution.(i) and sj = solution.(j) in - if si < sj then -1 else if si > sj then 1 else i - j) - positions; - Some positions - with Failure s -> - Printf.printf "failure in prepass scheduling: %s\n" s; - None +open ExpansionOracle +open PrepassSchedulingOracle let flatten_blk_basics ibf = let ib = ibf.entry in @@ -241,7 +53,8 @@ let schedule_blk n ibf btl = let rec do_schedule btl = function | [] -> btl | (n, ibf) :: blks -> - let btl' = schedule_blk n ibf btl in + let code_exp = expanse n ibf btl in + let btl' = schedule_blk n ibf code_exp in do_schedule btl' blks let btl_scheduler f = diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index e7d2b37f..a3053add 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -60,35 +60,33 @@ Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res uni end end. -Definition check_symbolic_simu (f1 f2: BTL.function): res unit := - check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))). - -Lemma check_symbolic_simu_rec_input_equiv lpc: forall f1 f2 x, - check_symbolic_simu_rec f1 f2 lpc = OK x -> - lpc = (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))) -> - equiv_input_regs f1 f2. +Definition erase_input_regs f1 f2 := + let code := PTree.map (fun pc ibf => + let oibf := match (fn_code f1)!pc with + | None => ibf + | Some oibf => oibf + end in + let regs := oibf.(input_regs) in + {| entry:= ibf.(entry); input_regs := regs; binfo := ibf.(binfo) |}) (fn_code f2) in + BTL.mkfunction (fn_sig f2) (fn_params f2) (fn_stacksize f2) code (fn_entrypoint f2). + +Lemma erase_input_regs_correct f1 f2 f3: + erase_input_regs f1 f2 = f3 -> + equiv_input_regs f1 f3. Proof. + unfold erase_input_regs; destruct f3; simpl; intros. + unfold equiv_input_regs; intuition auto. simpl. Admitted. -(* - unfold equiv_input_regs; induction lpc; simpl; intros f1 f2 x X EQL1. - - destruct (map _ _) eqn:EQL2; inv X; intuition; - symmetry in EQL1; apply map_eq_nil in EQL1; apply map_eq_nil in EQL2. - + destruct (fn_code f1), (fn_code f2); simpl in *; auto. - inv EQL2. - unfold PTree.elements in EQL1. - + admit. - + - destruct ((fn_code f1) ! a), ((fn_code f2) ! a); monadInv X. - eapply IHlpc; eauto. -Qed. - *) + +Definition check_symbolic_simu (f1 f2: BTL.function): res unit := + (*let f3 := erase_input_regs f1 f2 in*) + check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))). Lemma check_symbolic_simu_input_equiv x f1 f2: check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. Proof. unfold check_symbolic_simu; simpl; intros X. - eapply check_symbolic_simu_rec_input_equiv; eauto. -Qed. +Admitted. Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x, check_symbolic_simu_rec f1 f2 lpc = OK x -> diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 75f67d51..9b8dabab 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -231,13 +231,13 @@ Proof. - (* exec_load *) inv MIB. exists pc'; split; auto; constructor. apply plus_one. eapply exec_Iload; eauto. - erewrite <- eval_addressing_preserved; eauto. - intros; rewrite symbols_preserved; trivial. + all: erewrite <- eval_addressing_preserved; eauto; + intros; rewrite symbols_preserved; trivial. - (* exec_store *) inv MIB. exists pc'; split; auto; constructor. apply plus_one. eapply exec_Istore; eauto. - erewrite <- eval_addressing_preserved; eauto. - intros; rewrite symbols_preserved; trivial. + all: erewrite <- eval_addressing_preserved; eauto; + intros; rewrite symbols_preserved; trivial. - (* exec_seq_stop *) inv MIB; eauto. - (* exec_seq_continue *) -- cgit From 23c01485970efa11a7207ac2124f5922a011b0d4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 20 Jul 2021 20:01:12 +0200 Subject: new expansion oracle for BTL --- scheduling/BTLScheduleraux.ml | 4 +++- scheduling/BTL_Scheduler.v | 2 +- scheduling/BTL_Schedulerproof.v | 6 ++++-- 3 files changed, 8 insertions(+), 4 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 98bc4590..6a114b74 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -60,7 +60,9 @@ let rec do_schedule btl = function let btl_scheduler f = let btl = f.fn_code in (*debug_flag := true;*) - let btl' = do_schedule btl (PTree.elements btl) in + let elts = PTree.elements btl in + find_last_reg elts; + let btl' = do_schedule btl elts in debug "Scheduled BTL Code:\n"; print_btl_code stderr btl'; (*debug_flag := false;*) diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index a3053add..39672749 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -113,7 +113,7 @@ Qed. Definition transf_function (f: BTL.function) := let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in - do _ <- check_symbolic_simu f tf; + (*do _ <- check_symbolic_simu f tf;*) OK tf. Definition transf_fundef (f: fundef) : res fundef := diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 05028f11..9b0383ae 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -98,17 +98,19 @@ Proof. unfold transf_function. intros H. monadInv H. econstructor; eauto. eapply check_symbolic_simu_input_equiv; eauto. +Admitted. (* eapply check_symbolic_simu_correct; eauto. -Qed. + Qed.*) Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. intros TRANSF; destruct f; simpl; inv TRANSF. +Admitted. (* + monadInv H0. exploit transf_function_correct; eauto. constructor; auto. + eapply match_External. -Qed. + Qed.*) Lemma function_ptr_preserved: forall v f, -- cgit From 9ce1b23f6361d8070490df1269a7bc53aa215efb Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 21 Jul 2021 12:16:27 +0200 Subject: expansions btl proofs --- scheduling/BTL_SEimpl.v | 6 +++++- scheduling/BTL_Scheduler.v | 2 +- scheduling/BTL_Schedulerproof.v | 6 ++---- scheduling/BTLmatchRTL.v | 4 ++-- 4 files changed, 10 insertions(+), 8 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 3050f2e2..7711c72d 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -613,7 +613,11 @@ Proof. { wlp_simplify; exploit is_move_operation_correct; eauto. intros (Hop & Hlsv); subst; simpl in *. inv REF. simplify_SOME z; erewrite H; eauto. } - wlp_simplify; inv REF. erewrite H0; eauto. + destruct (target_op_simplify _ _ _) eqn: Htarget_op_simp; wlp_simplify. + { destruct (eval_list_sval _ _) eqn: OKlist; try congruence. + rewrite <- H; exploit target_op_simplify_correct; eauto. } + inv REF; clear Htarget_op_simp. + intuition eauto. - (* Rload *) destruct trap; wlp_simplify; inv REF. + erewrite H0; eauto. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 39672749..a3053add 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -113,7 +113,7 @@ Qed. Definition transf_function (f: BTL.function) := let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in - (*do _ <- check_symbolic_simu f tf;*) + do _ <- check_symbolic_simu f tf; OK tf. Definition transf_fundef (f: fundef) : res fundef := diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 9b0383ae..05028f11 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -98,19 +98,17 @@ Proof. unfold transf_function. intros H. monadInv H. econstructor; eauto. eapply check_symbolic_simu_input_equiv; eauto. -Admitted. (* eapply check_symbolic_simu_correct; eauto. - Qed.*) +Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. intros TRANSF; destruct f; simpl; inv TRANSF. -Admitted. (* + monadInv H0. exploit transf_function_correct; eauto. constructor; auto. + eapply match_External. - Qed.*) +Qed. Lemma function_ptr_preserved: forall v f, diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index 439ba9cc..a59c847e 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -412,7 +412,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different r in Bload") else Error (msg "BTL.verify_block: different lr in Bload") else Error (msg "BTL.verify_block: different addressing in Bload") - else Error (msg "BTL.verify_block: different mchunk in Bload") + else Error (msg "BTL.verify_block: different chunk in Bload") else Error (msg "BTL.verify_block: NOTRAP trapping_mode unsupported in Bload") | _ => Error (msg "BTL.verify_block: incorrect cfg Bload") end @@ -426,7 +426,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different r in Bstore") else Error (msg "BTL.verify_block: different lr in Bstore") else Error (msg "BTL.verify_block: different addressing in Bstore") - else Error (msg "BTL.verify_block: different mchunk in Bstore") + else Error (msg "BTL.verify_block: different chunk in Bstore") | _ => Error (msg "BTL.verify_block: incorrect cfg Bstore") end | Bseq b1 b2 => -- cgit From 6adfc52c130ee73cf6ee2ee8b85ed8b5e5024a4a Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 22 Jul 2021 16:43:14 +0200 Subject: branches expansions support --- scheduling/BTL_SEimpl.v | 23 ++++++++++++----------- scheduling/BTL_SEsimuref.v | 10 +++++----- 2 files changed, 17 insertions(+), 16 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 7711c72d..305c3cfa 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -682,12 +682,13 @@ Lemma cbranch_expanse_correct hrs c l: (REF : ris_refines ctx hrs sis) (OK: ris_ok ctx hrs), eval_scondition ctx (fst r) (snd r) = - eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)) - /\ (fst r) = c. + eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)). Proof. unfold cbranch_expanse. - wlp_simplify; inv REF. - unfold eval_scondition; erewrite <- H; eauto. + destruct (target_cbranch_expanse _ _ _) eqn: TARGET; wlp_simplify; + unfold eval_scondition; try erewrite <- H; inversion REF; eauto. + destruct p as [c' l']; simpl. + exploit target_cbranch_expanse_correct; eauto. Qed. Local Hint Resolve cbranch_expanse_correct: wlp. Global Opaque cbranch_expanse. @@ -1133,9 +1134,9 @@ Proof. erewrite <- OK_EQUIV. 2: eauto. eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto. } - exploit (Refcond ctx cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 + exploit (Refcond ctx (fst exta) cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k)); - exploit H; eauto; intros (SEVAL & EQC); subst; auto. + exploit H; eauto; intros SEVAL; auto. all: simpl in SOUT; generalize SOUT; clear SOUT; inversion_SOME b0; try_simplify_someHyps. @@ -1217,9 +1218,9 @@ Proof. assert (rOK: ris_ok ctx hrs). { simpl in H2; generalize H2; inversion_SOME b0; destruct b0; try_simplify_someHyps. } - exploit (Refcond ctx cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 + exploit (Refcond ctx (fst exta) cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1 (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k)); - exploit H; eauto; intros (SEVAL & EQC); subst; auto. + exploit H; eauto; intros SEVAL; auto. all: simpl in H2; generalize H2; clear H2; inversion_SOME b0; try_simplify_someHyps. @@ -1444,11 +1445,11 @@ Fixpoint rst_simu_check (rst1 rst2: rstate) {struct rst1} := hrs_simu_check hrs1 hrs2;; sfval_simu_check sfv1 sfv2 | Rcond cond1 lsv1 rsL1 rsR1, Rcond cond2 lsv2 rsL2 rsR2 => - struct_check cond1 cond2 "hsstate_simu_check: conditions do not match";; - phys_check lsv1 lsv2 "hsstate_simu_check: args do not match";; + struct_check cond1 cond2 "rst_simu_check: conditions do not match";; + phys_check lsv1 lsv2 "rst_simu_check: args do not match";; rst_simu_check rsL1 rsL2;; rst_simu_check rsR1 rsR2 - | _, _ => FAILWITH "hsstate_simu_check: simu_check failed" + | _, _ => FAILWITH "rst_simu_check: simu_check failed" end. Lemma rst_simu_check_correct rst1: forall rst2, diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index c812c607..7af1af62 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -260,11 +260,11 @@ Inductive rst_refines ctx: rstate -> sstate -> Prop := (RIS: ris_refines ctx ris sis) (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv) : rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv) - | Refcond cond rargs args rifso rifnot ifso ifnot - (RCOND: eval_scondition ctx cond rargs = eval_scondition ctx cond args) - (REFso: eval_scondition ctx cond rargs = Some true -> rst_refines ctx rifso ifso) - (REFnot: eval_scondition ctx cond rargs = Some false -> rst_refines ctx rifnot ifnot) - : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args ifso ifnot) + | Refcond rcond cond rargs args rifso rifnot ifso ifnot + (RCOND: eval_scondition ctx rcond rargs = eval_scondition ctx cond args) + (REFso: eval_scondition ctx rcond rargs = Some true -> rst_refines ctx rifso ifso) + (REFnot: eval_scondition ctx rcond rargs = Some false -> rst_refines ctx rifnot ifnot) + : rst_refines ctx (Rcond rcond rargs rifso rifnot) (Scond cond args ifso ifnot) | Refabort : rst_refines ctx Rabort Sabort . -- cgit From 8de8dc6616c49018c6151887f76ea08c8f1ff04e Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Thu, 22 Jul 2021 17:54:54 +0200 Subject: nothing admitted !! --- scheduling/BTL_Scheduler.v | 39 +++++++++++++++++++++++---------------- scheduling/BTL_Schedulerproof.v | 3 ++- 2 files changed, 25 insertions(+), 17 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index a3053add..0662a357 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -62,31 +62,35 @@ Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res uni Definition erase_input_regs f1 f2 := let code := PTree.map (fun pc ibf => - let oibf := match (fn_code f1)!pc with + let oibf := match (fn_code f2)!pc with | None => ibf | Some oibf => oibf end in - let regs := oibf.(input_regs) in - {| entry:= ibf.(entry); input_regs := regs; binfo := ibf.(binfo) |}) (fn_code f2) in + {| entry:= oibf.(entry); input_regs := ibf.(input_regs); binfo := oibf.(binfo) |}) (fn_code f1) in BTL.mkfunction (fn_sig f2) (fn_params f2) (fn_stacksize f2) code (fn_entrypoint f2). Lemma erase_input_regs_correct f1 f2 f3: erase_input_regs f1 f2 = f3 -> equiv_input_regs f1 f3. Proof. - unfold erase_input_regs; destruct f3; simpl; intros. - unfold equiv_input_regs; intuition auto. simpl. -Admitted. + unfold erase_input_regs; simpl; intros. + unfold equiv_input_regs; intuition auto. + - subst; simpl. rewrite PTree.gmap, H0; simpl; auto. + - subst; simpl in H0. rewrite PTree.gmap in H0. + destruct ((fn_code f1)!pc) eqn:EQF; simpl in H0; inv H0; auto. + - subst; simpl in H1. rewrite PTree.gmap in H1. + rewrite H0 in H1; simpl in H1; inv H1; simpl; auto. +Qed. +Global Opaque erase_input_regs. Definition check_symbolic_simu (f1 f2: BTL.function): res unit := - (*let f3 := erase_input_regs f1 f2 in*) check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))). -Lemma check_symbolic_simu_input_equiv x f1 f2: - check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. -Proof. - unfold check_symbolic_simu; simpl; intros X. -Admitted. +Definition transf_function (f: BTL.function) := + let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in + let tf' := erase_input_regs f tf in + do _ <- check_symbolic_simu f tf'; + OK tf'. Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x, check_symbolic_simu_rec f1 f2 lpc = OK x -> @@ -111,10 +115,13 @@ Proof. apply PTree.elements_correct in PC. eapply (in_map fst) in PC; auto. Qed. -Definition transf_function (f: BTL.function) := - let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in - do _ <- check_symbolic_simu f tf; - OK tf. +Lemma check_symbolic_simu_input_equiv x f1 f2: + transf_function f1 = OK f2 -> + check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. +Proof. + unfold transf_function; intros TRANSF X; monadInv TRANSF. + exploit erase_input_regs_correct; eauto. +Qed. Definition transf_fundef (f: fundef) : res fundef := transf_partial_fundef (fun f => transf_function f) f. diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 05028f11..27efe56f 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -95,7 +95,8 @@ Qed. Lemma transf_function_correct f f': transf_function f = OK f' -> match_function f f'. Proof. - unfold transf_function. intros H. monadInv H. + unfold transf_function. intros H. + assert (OH: transf_function f = OK f') by auto. monadInv H. econstructor; eauto. eapply check_symbolic_simu_input_equiv; eauto. eapply check_symbolic_simu_correct; eauto. -- cgit From b063fb03af84483671833d40491f4fa8d2c8b4c9 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Fri, 23 Jul 2021 18:25:00 +0200 Subject: non trapping loads --- scheduling/BTL.v | 51 ++++++++++++++++++++++++++++++----------- scheduling/BTL_Livecheck.v | 14 ++++++----- scheduling/BTL_SEimpl.v | 15 +++++------- scheduling/BTL_SEsimuref.v | 9 +++----- scheduling/BTL_SEtheory.v | 44 +++++++++++++++++++++-------------- scheduling/BTL_Schedulerproof.v | 2 +- scheduling/BTLmatchRTL.v | 14 +++++++---- scheduling/BTLtoRTLproof.v | 3 ++- scheduling/RTLtoBTLproof.v | 15 +++++++----- 9 files changed, 103 insertions(+), 64 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 6536addb..96377943 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -205,7 +205,7 @@ Definition find_function (ros: reg + ident) (rs: regset) : option fundef := Local Open Scope option_monad_scope. -(* (* TODO: a new (hopefully simpler) scheme to support "NOTRAP" wrt current scheme of RTL *) +(* TODO: a new (hopefully simpler) scheme to support "NOTRAP" wrt current scheme of RTL *) Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop := | has_loaded_normal a trap @@ -217,11 +217,10 @@ Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop := (DEFAULT: v = default_notrap_load_value chunk) : has_loaded sp rs m chunk addr args v NOTRAP . +Local Hint Constructors has_loaded: core. (* TODO: move this scheme in "Memory" module if this scheme is useful ! *) -*) - (** internal big-step execution of one iblock *) Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option final -> Prop := | exec_final rs m fin iinfo: iblock_istep sp rs m (BF fin iinfo) rs m (Some fin) @@ -229,15 +228,9 @@ Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option fi | exec_op rs m op args res v iinfo (EVAL: eval_operation ge sp op rs##args m = Some v) : iblock_istep sp rs m (Bop op args res iinfo) (rs#res <- v) m None - | exec_load_TRAP rs m chunk addr args dst a v iinfo - (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 iinfo) (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 + | exec_load rs m trap chunk addr args dst v iinfo (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 -*) + : iblock_istep sp rs m (Bload trap chunk addr args dst iinfo) (rs#dst <- v) m None | exec_store rs m chunk addr args src a m' iinfo (EVAL: eval_addressing ge sp addr rs##args = Some a) (STORE: Mem.storev chunk m a rs#src = Some m') @@ -274,7 +267,18 @@ Fixpoint iblock_istep_run sp ib rs m: option outcome := SOME v <- Mem.loadv chunk m a IN Some {| _rs := rs#dst <- v; _m:= m; _fin := None |} | Bload NOTRAP chunk addr args dst _ => - None (* TODO *) + match eval_addressing ge sp addr rs##args with + | Some a => + match Mem.loadv chunk m a with + | Some v => Some {| _rs := rs#dst <- v; _m:= m; _fin := None |} + | None => + let v := default_notrap_load_value chunk in + Some {| _rs := rs#dst <- v; _m:= m; _fin := None |} + end + | None => + let v := default_notrap_load_value chunk in + Some {| _rs := rs#dst <- v; _m:= m; _fin := None |} + end | 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 @@ -291,13 +295,34 @@ Fixpoint iblock_istep_run sp ib rs m: option outcome := iblock_istep_run sp (if b then ifso else ifnot) rs m end. +Lemma iblock_istep_run_equiv_load sp ib v rs rs' m trap chunk addr args dst iinfo ofin: + ib = (Bload trap chunk addr args dst iinfo) -> + rs' = rs # dst <- v -> + has_loaded sp rs m chunk addr args v trap -> + iblock_istep sp rs m ib rs' m ofin <-> + iblock_istep_run sp ib rs m = Some {| _rs := rs'; _m := m; _fin := ofin |}. +Proof. + split; subst; inv H1; simpl in *; intros. + - destruct trap; inv H; simpl in *; repeat autodestruct. + - inv H; autodestruct; rewrite LOAD; auto. + - destruct trap; inv H; simpl in *; + rewrite EVAL, LOAD in H1; inv H1; repeat econstructor; eauto. + - generalize H; autodestruct. + + rewrite LOAD in *; inv H; auto; constructor. + eapply has_loaded_default; eauto; try_simplify_someHyps. + + inv H; constructor. eapply has_loaded_default; eauto; try_simplify_someHyps. +Qed. +Local Hint Resolve iblock_istep_run_equiv_load: core. + Lemma iblock_istep_run_equiv sp rs m ib rs' m' ofin: iblock_istep sp rs m ib rs' m' ofin <-> iblock_istep_run sp ib rs m = Some {| _rs := rs'; _m:= m'; _fin := ofin |}. Proof. constructor. - - induction 1; simpl; try autodestruct; try_simplify_someHyps. + - induction 1; try (eapply iblock_istep_run_equiv_load; eauto; fail); + simpl; try autodestruct; try_simplify_someHyps. - generalize rs m rs' m' ofin; clear rs m rs' m' ofin. induction ib; simpl; repeat (try autodestruct; try_simplify_someHyps). + 1,2: constructor; eapply has_loaded_default; try_simplify_someHyps. destruct o; try_simplify_someHyps; subst; eauto. Qed. diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v index 9f96e74e..d200b9bd 100644 --- a/scheduling/BTL_Livecheck.v +++ b/scheduling/BTL_Livecheck.v @@ -546,12 +546,14 @@ Proof. - (* Bload *) erewrite <- eqlive_reg_listmem; eauto. try_simplify_someHyps; intros. - rewrite LOAD; eauto. - repeat econstructor. - apply eqlive_reg_update. - eapply eqlive_reg_monotonic; eauto. - intros r0; rewrite regset_add_spec. - intuition. + destruct trap; inv LOAD; + rewrite EVAL, LOAD0 || (autodestruct; try rewrite LOAD0; auto). + all: + repeat econstructor; + apply eqlive_reg_update; + eapply eqlive_reg_monotonic; eauto; + intros r0; rewrite regset_add_spec; + intuition. - (* Bstore *) erewrite <- eqlive_reg_listmem; eauto. rewrite <- (REGS src); auto. diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index 305c3cfa..aa27bd19 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -823,11 +823,11 @@ Proof. simpl in *; econstructor; eauto. Qed. -Lemma hrs_ok_load_okpreserv ctx chunk addr args dest hrs: - WHEN hrs_sreg_set dest args (Rload TRAP chunk addr) hrs ~> rst THEN +Lemma hrs_ok_load_okpreserv ctx chunk addr args dest hrs trap: + WHEN hrs_sreg_set dest args (Rload trap chunk addr) hrs ~> rst THEN ris_ok ctx rst -> ris_ok ctx hrs. Proof. - wlp_simplify; inv H1; + wlp_simplify; inv H2 || inv H1; simpl in *; econstructor; eauto. Qed. @@ -1071,10 +1071,9 @@ Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate := | Bop op args dst _ => DO next <~ hrs_sreg_set dst args (Rop op) hrs;; k next - | Bload TRAP chunk addr args dst _ => - DO next <~ hrs_sreg_set dst args (Rload TRAP chunk addr) hrs;; + | Bload trap chunk addr args dst _ => + DO next <~ hrs_sreg_set dst args (Rload trap chunk addr) hrs;; k next - | Bload NOTRAP chunk addr args dst _ => RET Rabort | Bstore chunk addr args src _ => DO next <~ hrexec_store chunk addr args src hrs;; k next @@ -1122,7 +1121,6 @@ Proof. intros X; apply H0 in X. exploit hrexec_final_sfv_correct; eauto. - (* Bnop *) wlp_simplify; eapply CONT; eauto. - - (* Bload *) destruct trap; wlp_simplify. - (* Bseq *) wlp_simplify. eapply IHib1. 3-7: eauto. @@ -1171,7 +1169,7 @@ Proof. - (* Bop *) wlp_simplify; exploit hrs_ok_op_okpreserv; eauto. - (* Bload *) - destruct trap; wlp_simplify; try_simplify_someHyps. + destruct trap; wlp_simplify; try_simplify_someHyps; exploit hrs_ok_load_okpreserv; eauto. - (* Bstore *) wlp_simplify; exploit hrs_ok_store_okpreserv; eauto. @@ -1207,7 +1205,6 @@ Proof. intros X; apply H0 in X. exploit hrexec_final_sfv_correct; eauto. - (* Bnop *) wlp_simplify; eapply CONT; eauto. - - (* Bload *) destruct trap; wlp_simplify. - (* Bseq *) wlp_simplify. eapply IHib1. 3-6: eauto. diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v index 7af1af62..d47e53b8 100644 --- a/scheduling/BTL_SEsimuref.v +++ b/scheduling/BTL_SEsimuref.v @@ -348,7 +348,7 @@ Proof. unfold sexec_op. rewrite ok_set_sreg. intuition. Qed. -Lemma si_ok_load_okpreserv ctx chunk addr args dest sis: si_ok ctx (sexec_load TRAP chunk addr args dest sis) -> si_ok ctx sis. +Lemma si_ok_load_okpreserv ctx chunk addr args dest sis trap: si_ok ctx (sexec_load trap chunk addr args dest sis) -> si_ok ctx sis. Proof. unfold sexec_load. rewrite ok_set_sreg. intuition. Qed. @@ -705,8 +705,7 @@ Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate := (* basic instructions *) | Bnop _ => k ris | Bop op args res _ => k (rexec_op op args res ris) - | Bload TRAP chunk addr args dst _ => k (rexec_load TRAP chunk addr args dst ris) - | Bload NOTRAP chunk addr args dst _ => Rabort + | Bload trap chunk addr args dst _ => k (rexec_load trap chunk addr args dst ris) | Bstore chunk addr args src _ => k (rexec_store chunk addr args src ris) (* composed instructions *) | Bseq ib1 ib2 => @@ -739,7 +738,6 @@ Lemma rexec_rec_correct1 ctx ib: , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st. Proof. induction ib; simpl; try (intros; subst; eauto; fail). - - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. - (* seq *) intros; subst. eapply IHib1. 3-6: eauto. @@ -774,7 +772,7 @@ Proof. unfold rexec_op. rewrite ok_rset_sreg. intuition. Qed. -Lemma ris_ok_load_okpreserv ctx chunk addr args dest ris: ris_ok ctx (rexec_load TRAP chunk addr args dest ris) -> ris_ok ctx ris. +Lemma ris_ok_load_okpreserv ctx chunk addr args dest ris trap: ris_ok ctx (rexec_load trap chunk addr args dest ris) -> ris_ok ctx ris. Proof. unfold rexec_load. rewrite ok_rset_sreg. intuition. Qed. @@ -818,7 +816,6 @@ Lemma rexec_rec_correct2 ctx ib: , rst_refines ctx st (sexec_rec (cf ctx) ib sis k). Proof. induction ib; simpl; try (intros; subst; eauto; fail). - - (* load *) intros; subst; autodestruct; simpl in *; subst; eauto. - (* seq *) intros; subst. eapply IHib1. 3-6: eauto. diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 2a335790..e2e6be6b 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -626,18 +626,22 @@ Definition sexec_load trap chunk addr args dst sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in set_sreg dst (fSload sis.(si_smem) trap chunk addr args) sis. -Lemma sexec_load_TRAP_correct ctx chunk addr args dst sis rs m a v - (EVAL: eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a) - (LOAD: Mem.loadv chunk m a = Some v) +Lemma sexec_load_correct ctx chunk addr args dst sis rs m v trap + (HLOAD: has_loaded (cge ctx) (csp ctx) rs m chunk addr args v trap) (SIS: sem_sistate ctx sis rs m) - :(sem_sistate ctx (sexec_load TRAP chunk addr args dst sis) (rs#dst <- v) m). + :(sem_sistate ctx (sexec_load trap chunk addr args dst sis) (rs#dst <- v) m). Proof. - eapply set_sreg_correct; eauto. + inv HLOAD; eapply set_sreg_correct; eauto. + 2,4: intros; rewrite Regmap.gso; auto. + - simpl. destruct SIS as (PRE&MEM®). + destruct trap; rewrite Regmap.gss; simpl; auto; + erewrite eval_list_sval_inj; simpl; auto; + try_simplify_someHyps. + intros. rewrite LOAD; auto. - simpl. destruct SIS as (PRE&MEM®). rewrite Regmap.gss; simpl; auto. erewrite eval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - - intros; rewrite Regmap.gso; auto. + autodestruct; rewrite MEM, LOAD; auto. Qed. Definition sexec_store chunk addr args src sis: sistate := @@ -812,8 +816,7 @@ Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := (* basic instructions *) | Bnop _ => k sis | Bop op args res _ => k (sexec_op op args res sis) - | Bload TRAP chunk addr args dst _ => k (sexec_load TRAP chunk addr args dst sis) - | Bload NOTRAP chunk addr args dst _ => Sabort (* TODO *) + | Bload trap chunk addr args dst _ => k (sexec_load trap chunk addr args dst sis) | Bstore chunk addr args src _ => k (sexec_store chunk addr args src sis) (* composed instructions *) | Bseq ib1 ib2 => @@ -830,7 +833,7 @@ Definition sexec f ib := sexec_rec f ib sis_init (fun _ => Sabort). Local Hint Constructors sem_sstate: core. Local Hint Resolve sexec_op_correct sexec_final_sfv_correct tr_sis_regs_correct_aux tr_sis_regs_correct - sexec_load_TRAP_correct sexec_store_correct sis_init_correct: core. + sexec_load_correct sexec_store_correct sis_init_correct: core. Lemma sexec_rec_correct ctx stk t s ib rs m rs1 m1 ofin (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): forall sis k @@ -921,9 +924,9 @@ Proof. intros; eapply set_sreg_preserves_abort; eauto. Qed. -Lemma sexec_load_preserves_abort ctx chunk addr args dest sis: +Lemma sexec_load_preserves_abort ctx chunk addr args dest sis trap: abort_sistate ctx sis - -> abort_sistate ctx (sexec_load TRAP chunk addr args dest sis). + -> abort_sistate ctx (sexec_load trap chunk addr args dest sis). Proof. intros; eapply set_sreg_preserves_abort; eauto. Qed. @@ -962,8 +965,6 @@ Lemma sexec_exclude_abort ctx stk ib t s1: forall sis k Proof. induction ib; simpl; intros; eauto. - (* final *) inversion SEXEC; subst; eauto. - - (* load *) destruct trap; eauto. - inversion SEXEC. - (* seq *) eapply IHib1; eauto. simpl. eauto. @@ -1052,11 +1053,20 @@ Proof. * intros; rewrite REG, str_tr_regs_equiv; auto. * intros (s2 & EQUIV & SFV'); eauto. - (* Bop *) autodestruct; eauto. - - destruct trap; [| inv SEXEC ]. - repeat autodestruct; eauto. - all: intros; eapply CONT; eauto; + - destruct trap. + + repeat autodestruct. + { eexists; split; eauto. + eapply sexec_load_correct; eauto. + econstructor; eauto. } + all: + intros; eapply CONT; eauto; eapply sexec_load_TRAP_abort; eauto; intros; try_simplify_someHyps. + + repeat autodestruct; + eexists; split; eauto; + eapply sexec_load_correct; eauto; + try (econstructor; eauto; fail). + all: eapply has_loaded_default; auto; try_simplify_someHyps. - repeat autodestruct; eauto. all: intros; eapply CONT; eauto; eapply sexec_store_abort; eauto; diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 27efe56f..40ad0d88 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -354,7 +354,7 @@ Proof. - repeat econstructor; eauto. - repeat econstructor; eauto. erewrite args_eqregs; eauto. - - repeat econstructor; eauto. + - inv LOAD; repeat econstructor; eauto; erewrite args_eqregs; eauto. - repeat econstructor; eauto. erewrite args_eqregs; eauto. diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index a59c847e..c271ae02 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -48,11 +48,15 @@ Proof. do 2 eexists; rewrite EVAL'. repeat (split; eauto). eapply set_reg_lessdef; eauto. - (* Bload *) - exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto. - intros (v2 & EVAL' & LESSDEF). exploit Mem.loadv_extends; eauto. - intros (v3 & LOAD' & LESSDEF'). - do 2 eexists; rewrite EVAL', LOAD'. repeat (split; eauto). - eapply set_reg_lessdef; eauto. + inv LOAD. + + exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto. + intros (v2 & EVAL' & LESSDEF). exploit Mem.loadv_extends; eauto. + intros (v3 & LOAD' & LESSDEF'); autodestruct; + do 2 eexists; rewrite EVAL', LOAD'; + repeat (split; eauto); eapply set_reg_lessdef; eauto. + + destruct (eval_addressing ge sp addr rs ## args) eqn:EQA; + repeat autodestruct; do 2 eexists; + repeat (split; eauto); eapply set_reg_lessdef; eauto. - (* Bstore *) exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto. intros (v2 & EVAL' & LESSDEF). exploit Mem.storev_extends; eauto. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 9b8dabab..324f1d14 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -230,7 +230,8 @@ Proof. intros; rewrite symbols_preserved; trivial. - (* exec_load *) inv MIB. exists pc'; split; auto; constructor. - apply plus_one. eapply exec_Iload; eauto. + apply plus_one. inv LOAD. + eapply exec_Iload; eauto. all: erewrite <- eval_addressing_preserved; eauto; intros; rewrite symbols_preserved; trivial. - (* exec_store *) diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 8a352434..13ba5a29 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -123,12 +123,15 @@ Lemma normRTLrec_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin1: iblock_istep tge sp rs0 m0 (normRTLrec ib k) rs2 m2 ofin2. Proof. induction 1; simpl; intuition subst; eauto. - { (* Bnop *) autodestruct; eauto. } - 1-3: (* Bop, Bload, Bstore *) - intros; repeat econstructor; eauto. - (* Bcond *) - destruct ofin; intuition subst; - destruct b; eapply IHISTEP; eauto. + - (* Bnop *) autodestruct; eauto. + - (* Bop *) repeat econstructor; eauto. + - (* Bload *) inv LOAD. + + repeat econstructor; eauto. + + do 2 (econstructor; eauto). + eapply has_loaded_default; eauto. + - (* Bcond *) repeat econstructor; eauto. + destruct ofin; intuition subst; + destruct b; eapply IHISTEP; eauto. Qed. Lemma normRTL_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin: -- cgit From 0b076ef6eb5553be43ce81c27e438f632b17cb32 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 27 Jul 2021 11:22:30 +0200 Subject: prepass act --- scheduling/BTLScheduleraux.ml | 26 ++++++++++++++------------ 1 file changed, 14 insertions(+), 12 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 6a114b74..9c8f6ab5 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -37,18 +37,20 @@ let apply_schedule bseq olast positions = build_iblock ibl let schedule_blk n ibf btl = - let bseq, olast = flatten_blk_basics ibf in - match schedule_sequence bseq btl with - | Some positions -> - debug "%d," (p2i n); - Array.iter (fun p -> debug "%d " p) positions; - debug "\n"; - let new_ib = apply_schedule bseq olast positions in - let new_ibf = - { entry = new_ib; binfo = ibf.binfo; input_regs = ibf.input_regs } - in - PTree.set n new_ibf btl - | None -> btl + if not !Clflags.option_fprepass then btl + else + let bseq, olast = flatten_blk_basics ibf in + match schedule_sequence bseq btl with + | Some positions -> + debug "%d," (p2i n); + Array.iter (fun p -> debug "%d " p) positions; + debug "\n"; + let new_ib = apply_schedule bseq olast positions in + let new_ibf = + { entry = new_ib; binfo = ibf.binfo; input_regs = ibf.input_regs } + in + PTree.set n new_ibf btl + | None -> btl let rec do_schedule btl = function | [] -> btl -- cgit From b3e47d62f708777248e5c630abd3afa8ddfdefc4 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 27 Jul 2021 17:07:46 +0200 Subject: test non-trapping loads using CI... --- scheduling/BTLScheduleraux.ml | 95 ++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 93 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 6a114b74..ebd4089b 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -1,6 +1,7 @@ open Maps open BTL open BTLtypes +open Machine open DebugPrint open PrintBTL open RTLcommonaux @@ -27,13 +28,80 @@ let flatten_blk_basics ibf = let ibl = traverse_blk ib in (Array.of_list ibl, !last) +let is_a_cb = function Bcond _ -> true | _ -> false + +let is_a_load = function Bload _ -> true | _ -> false + +module SI = Set.Make (Int) + +let find_array arr n = + let index = ref None in + (try + Array.iteri + (fun i n' -> + match !index with + | Some _ -> raise Exit + | None -> if n = n' then index := Some i) + arr + with Exit -> ()); + get_some @@ !index + +let count_cbs bseq olast indexes = + debug "count_cbs\n"; + let current_cbs = ref SI.empty in + let cbs_above = Hashtbl.create 100 in + let update_cbs n ib = + print_btl_inst stderr ib; + if is_a_cb ib then ( + debug "n is %d, add cb at: %d\n" n indexes.(n); + current_cbs := SI.add indexes.(n) !current_cbs) + else if is_a_load ib then ( + debug "n is %d, add load at: %d\n" n indexes.(n); + Hashtbl.add cbs_above indexes.(n) !current_cbs) + in + Array.iteri (fun n ib -> update_cbs n ib) bseq; + (match olast with + | Some last -> + debug "last\n"; + update_cbs (Array.length bseq) last + | None -> ()); + cbs_above + let apply_schedule bseq olast positions = - let ibl = Array.to_list (Array.map (fun i -> bseq.(i)) positions) in + let fmap n = find_array positions n in + let seq = (Array.init (Array.length positions) (fun i -> i)) in + let fseq = Array.map fmap seq in + debug_flag := true; + Array.iter (fun i -> debug "%d " i) positions; + debug "\n"; + Array.iter (fun i -> debug "%d " i) fseq; + debug "\n"; + Array.iter + (fun i -> debug "%d " i) + (Array.init (Array.length positions) (fun i -> i)); + debug "\n"; + let cbs_above_old = count_cbs bseq olast fseq in + let bseq_new = Array.map (fun i -> bseq.(i)) positions in + let cbs_above_new = count_cbs bseq_new olast seq in + Array.iteri + (fun n ib -> + let n' = fseq.(n) in + match ib with + | Bload (t, a, b, c, d, e) -> + let set_old = Hashtbl.find cbs_above_old n' in + let set_new = Hashtbl.find cbs_above_new n' in + if SI.subset set_old set_new then + bseq_new.(n') <- Bload (AST.TRAP, a, b, c, d, e) + else assert !config.has_non_trapping_loads + | _ -> ()) + bseq; + let ibl = Array.to_list bseq_new in let rec build_iblock = function | [] -> failwith "build_iblock: empty list" | [ ib ] -> ( match olast with Some last -> Bseq (ib, last) | None -> ib) | ib1 :: ib2 :: k -> Bseq (ib1, build_iblock (ib2 :: k)) in + debug_flag := false; build_iblock ibl let schedule_blk n ibf btl = @@ -50,11 +118,34 @@ let schedule_blk n ibf btl = PTree.set n new_ibf btl | None -> btl +let turn_all_loads_nontrap n ibf btl = + if not !config.has_non_trapping_loads then btl + else + let rec traverse_rec ib = + match ib with + | Bseq (ib1, ib2) -> Bseq (traverse_rec ib1, traverse_rec ib2) + | Bload (t, a, b, c, d, e) -> Bload (AST.NOTRAP, a, b, c, d, e) + | _ -> ib + in + let ib' = traverse_rec ibf.entry in + let ibf' = + { entry = ib'; input_regs = ibf.input_regs; binfo = ibf.binfo } + in + PTree.set n ibf' btl + let rec do_schedule btl = function | [] -> btl | (n, ibf) :: blks -> let code_exp = expanse n ibf btl in - let btl' = schedule_blk n ibf code_exp in + let code_nt = turn_all_loads_nontrap n ibf btl in + let btl' = schedule_blk n ibf code_nt in + (*debug_flag := true;*) + if btl != code_exp then ( + debug "#######################################################\n"; + print_btl_code stderr btl; + debug "~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~\n"; + print_btl_code stderr code_exp); + (*debug_flag := false;*) do_schedule btl' blks let btl_scheduler f = -- cgit From 77ee161826e24e87f801cbbeb797fb3a4a4a0fe9 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 27 Jul 2021 17:28:16 +0200 Subject: test ci2 --- scheduling/BTLScheduleraux.ml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 268ee7a5..5ebc4144 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -71,7 +71,7 @@ let apply_schedule bseq olast positions = let fmap n = find_array positions n in let seq = (Array.init (Array.length positions) (fun i -> i)) in let fseq = Array.map fmap seq in - debug_flag := true; + (*debug_flag := true;*) Array.iter (fun i -> debug "%d " i) positions; debug "\n"; Array.iter (fun i -> debug "%d " i) fseq; @@ -92,7 +92,7 @@ let apply_schedule bseq olast positions = let set_new = Hashtbl.find cbs_above_new n' in if SI.subset set_old set_new then bseq_new.(n') <- Bload (AST.TRAP, a, b, c, d, e) - else assert !config.has_non_trapping_loads + else (Printf.eprintf "\nTEST_GOURDINL_OK\n"; assert !config.has_non_trapping_loads) | _ -> ()) bseq; let ibl = Array.to_list bseq_new in @@ -101,7 +101,7 @@ let apply_schedule bseq olast positions = | [ ib ] -> ( match olast with Some last -> Bseq (ib, last) | None -> ib) | ib1 :: ib2 :: k -> Bseq (ib1, build_iblock (ib2 :: k)) in - debug_flag := false; + (*debug_flag := false;*) build_iblock ibl let schedule_blk n ibf btl = -- cgit From 056658bd2986d9e12ac07a54d25c08eb8a62ff60 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 28 Jul 2021 10:32:09 +0200 Subject: remove todos, clean --- scheduling/BTLScheduleraux.ml | 70 +++++++++++++---------------------------- scheduling/BTL_SEimpl.v | 20 +----------- scheduling/BTL_Scheduler.v | 46 +++++++++++++++++++++++++++ scheduling/BTL_Schedulerproof.v | 47 --------------------------- scheduling/PrintBTL.ml | 13 -------- scheduling/RTLpathSE_impl.v | 14 --------- scheduling/RTLtoBTLaux.ml | 2 -- 7 files changed, 69 insertions(+), 143 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index 5ebc4144..c8c4e0d3 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -47,61 +47,45 @@ let find_array arr n = get_some @@ !index let count_cbs bseq olast indexes = - debug "count_cbs\n"; let current_cbs = ref SI.empty in let cbs_above = Hashtbl.create 100 in let update_cbs n ib = print_btl_inst stderr ib; - if is_a_cb ib then ( - debug "n is %d, add cb at: %d\n" n indexes.(n); - current_cbs := SI.add indexes.(n) !current_cbs) - else if is_a_load ib then ( - debug "n is %d, add load at: %d\n" n indexes.(n); - Hashtbl.add cbs_above indexes.(n) !current_cbs) + if is_a_cb ib then current_cbs := SI.add indexes.(n) !current_cbs + else if is_a_load ib then Hashtbl.add cbs_above indexes.(n) !current_cbs in Array.iteri (fun n ib -> update_cbs n ib) bseq; (match olast with - | Some last -> - debug "last\n"; - update_cbs (Array.length bseq) last + | Some last -> update_cbs (Array.length bseq) last | None -> ()); cbs_above let apply_schedule bseq olast positions = - let fmap n = find_array positions n in - let seq = (Array.init (Array.length positions) (fun i -> i)) in - let fseq = Array.map fmap seq in - (*debug_flag := true;*) - Array.iter (fun i -> debug "%d " i) positions; - debug "\n"; - Array.iter (fun i -> debug "%d " i) fseq; - debug "\n"; - Array.iter - (fun i -> debug "%d " i) - (Array.init (Array.length positions) (fun i -> i)); - debug "\n"; - let cbs_above_old = count_cbs bseq olast fseq in let bseq_new = Array.map (fun i -> bseq.(i)) positions in - let cbs_above_new = count_cbs bseq_new olast seq in - Array.iteri - (fun n ib -> - let n' = fseq.(n) in - match ib with - | Bload (t, a, b, c, d, e) -> - let set_old = Hashtbl.find cbs_above_old n' in - let set_new = Hashtbl.find cbs_above_new n' in - if SI.subset set_old set_new then - bseq_new.(n') <- Bload (AST.TRAP, a, b, c, d, e) - else (Printf.eprintf "\nTEST_GOURDINL_OK\n"; assert !config.has_non_trapping_loads) - | _ -> ()) - bseq; + (if !config.has_non_trapping_loads then + let fmap n = find_array positions n in + let seq = Array.init (Array.length positions) (fun i -> i) in + let fseq = Array.map fmap seq in + let cbs_above_old = count_cbs bseq olast fseq in + let cbs_above_new = count_cbs bseq_new olast seq in + Array.iteri + (fun n ib -> + let n' = fseq.(n) in + match ib with + | Bload (t, a, b, c, d, e) -> + let set_old = Hashtbl.find cbs_above_old n' in + let set_new = Hashtbl.find cbs_above_new n' in + if SI.subset set_old set_new then + bseq_new.(n') <- Bload (AST.TRAP, a, b, c, d, e) + else assert !config.has_non_trapping_loads + | _ -> ()) + bseq); let ibl = Array.to_list bseq_new in let rec build_iblock = function | [] -> failwith "build_iblock: empty list" | [ ib ] -> ( match olast with Some last -> Bseq (ib, last) | None -> ib) | ib1 :: ib2 :: k -> Bseq (ib1, build_iblock (ib2 :: k)) in - (*debug_flag := false;*) build_iblock ibl let schedule_blk n ibf btl = @@ -110,9 +94,6 @@ let schedule_blk n ibf btl = let bseq, olast = flatten_blk_basics ibf in match schedule_sequence bseq btl with | Some positions -> - debug "%d," (p2i n); - Array.iter (fun p -> debug "%d " p) positions; - debug "\n"; let new_ib = apply_schedule bseq olast positions in let new_ibf = { entry = new_ib; binfo = ibf.binfo; input_regs = ibf.input_regs } @@ -139,15 +120,8 @@ let rec do_schedule btl = function | [] -> btl | (n, ibf) :: blks -> let code_exp = expanse n ibf btl in - let code_nt = turn_all_loads_nontrap n ibf btl in + let code_nt = turn_all_loads_nontrap n ibf code_exp in let btl' = schedule_blk n ibf code_nt in - (*debug_flag := true;*) - if btl != code_exp then ( - debug "#######################################################\n"; - print_btl_code stderr btl; - debug "~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~\n"; - print_btl_code stderr code_exp); - (*debug_flag := false;*) do_schedule btl' blks let btl_scheduler f = diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v index aa27bd19..65d16d01 100644 --- a/scheduling/BTL_SEimpl.v +++ b/scheduling/BTL_SEimpl.v @@ -745,7 +745,6 @@ Proof. rewrite Heq; eauto. Qed. -(* TODO gourdinl move this in BTL_SEtheory? *) Lemma eval_list_sval_inj_not_none ctx st: forall l, (forall r, List.In r l -> eval_sval ctx (si_sreg st r) = None -> False) -> eval_list_sval ctx (list_sval_inj (map (si_sreg st) l)) = None -> False. @@ -853,23 +852,6 @@ Proof. split; destruct 1; econstructor; simpl in *; eauto. Qed. -(* TODO useless? -Lemma hrset_red_reg_correct ctx r rsv sv hrs sis: - ris_refines ctx hrs sis -> - (ris_ok ctx hrs -> eval_sval ctx rsv <> None) -> - (ris_ok ctx hrs -> eval_sval ctx rsv = eval_sval ctx sv) -> - ris_refines ctx (ris_sreg_set hrs (red_PTree_set r rsv hrs)) (set_sreg r sv sis). -Proof. - destruct 1; intros. - econstructor; eauto. - + rewrite ok_set_sreg, ok_hrset_red_sreg; intuition congruence. - + rewrite ok_hrset_red_sreg; intuition eauto. - + rewrite ok_hrset_red_sreg. intros; unfold set_sreg; simpl; - erewrite red_PTree_set_refines; eauto. - econstructor; eauto. -Qed. - *) - Fixpoint transfer_hrs (inputs: list reg) (hrs: ristate): ?? ristate := match inputs with | nil => RET {| ris_smem := hrs.(ris_smem); @@ -1474,7 +1456,7 @@ Definition simu_check_single (f1 f2: function) (ibf1 ibf2: iblock_info): ?? unit Lemma simu_check_single_correct (f1 f2: function) (ibf1 ibf2: iblock_info): WHEN simu_check_single f1 f2 ibf1 ibf2 ~> _ THEN symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry). -Proof. (* TODO *) +Proof. unfold symbolic_simu; wlp_simplify. eapply rst_simu_correct; eauto. + intros; eapply hrexec_correct1; eauto; wlp_simplify. diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v index 0662a357..78c597e0 100644 --- a/scheduling/BTL_Scheduler.v +++ b/scheduling/BTL_Scheduler.v @@ -46,6 +46,52 @@ Record match_function (f1 f2: BTL.function): Prop := { exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2); }. +Inductive match_fundef: fundef -> fundef -> Prop := + | match_Internal f f': match_function f f' -> match_fundef (Internal f) (Internal f') + | match_External ef: match_fundef (External ef) (External ef). + +Inductive match_stackframes: stackframe -> stackframe -> Prop := + | match_stackframe_intro + res f sp pc rs rs' f' + (EQREGS: forall r, rs # r = rs' # r) + (TRANSF: match_function f f') + : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs'). + +Inductive match_states: state -> state -> Prop := + | match_states_intro + st f sp pc rs rs' m st' f' + (EQREGS: forall r, rs # r = rs' # r) + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function f f') + : match_states (State st f sp pc rs m) (State st' f' sp pc rs' m) + | match_states_call + st st' f f' args m + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_fundef f f') + : match_states (Callstate st f args m) (Callstate st' f' args m) + | match_states_return + st st' v m + (STACKS: list_forall2 match_stackframes st st') + : match_states (Returnstate st v m) (Returnstate st' v m) + . + +Lemma match_stack_equiv stk1 stk2: + list_forall2 match_stackframes stk1 stk2 -> + forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> + list_forall2 match_stackframes stk1 stk3. +Proof. + induction 1; intros stk3 EQ; inv EQ; constructor; eauto. + inv H3; inv H; econstructor; eauto. + intros; rewrite <- EQUIV; auto. +Qed. + +Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3. +Proof. + Local Hint Resolve match_stack_equiv: core. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; rewrite <- EQUIV; auto. +Qed. + Local Open Scope error_monad_scope. Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res unit := diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v index 40ad0d88..9a73d278 100644 --- a/scheduling/BTL_Schedulerproof.v +++ b/scheduling/BTL_Schedulerproof.v @@ -45,53 +45,6 @@ Proof. + eapply linkorder_refl. Qed. -(* TODO gourdinl move this to BTL_Scheduler.v? *) -Inductive match_fundef: fundef -> fundef -> Prop := - | match_Internal f f': match_function f f' -> match_fundef (Internal f) (Internal f') - | match_External ef: match_fundef (External ef) (External ef). - -Inductive match_stackframes: stackframe -> stackframe -> Prop := - | match_stackframe_intro - res f sp pc rs rs' f' - (EQREGS: forall r, rs # r = rs' # r) - (TRANSF: match_function f f') - : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs'). - -Inductive match_states: state -> state -> Prop := - | match_states_intro - st f sp pc rs rs' m st' f' - (EQREGS: forall r, rs # r = rs' # r) - (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_function f f') - : match_states (State st f sp pc rs m) (State st' f' sp pc rs' m) - | match_states_call - st st' f f' args m - (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_fundef f f') - : match_states (Callstate st f args m) (Callstate st' f' args m) - | match_states_return - st st' v m - (STACKS: list_forall2 match_stackframes st st') - : match_states (Returnstate st v m) (Returnstate st' v m) - . - -Lemma match_stack_equiv stk1 stk2: - list_forall2 match_stackframes stk1 stk2 -> - forall stk3, list_forall2 equiv_stackframe stk2 stk3 -> - list_forall2 match_stackframes stk1 stk3. -Proof. - induction 1; intros stk3 EQ; inv EQ; constructor; eauto. - inv H3; inv H; econstructor; eauto. - intros; rewrite <- EQUIV; auto. -Qed. - -Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3. -Proof. - Local Hint Resolve match_stack_equiv: core. - destruct 1; intros EQ; inv EQ; econstructor; eauto. - intros; rewrite <- EQUIV; auto. -Qed. - Lemma transf_function_correct f f': transf_function f = OK f' -> match_function f f'. Proof. diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml index 52178064..89254301 100644 --- a/scheduling/PrintBTL.ml +++ b/scheduling/PrintBTL.ml @@ -116,16 +116,3 @@ let print_btl_code pp btl = (PTree.elements btl); fprintf pp "\n") else () - -(* TODO gourdinl remove or adapt this? -let print_function pp id f = - fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; - let instrs = List.map (fun (n, i) -> i.entry) (PTree.elements f.fn_code) in - List.iter (print_iblock pp true "") instrs; - fprintf pp "}\n\n" - -let print_globdef pp (id, gd) = - match gd with Gfun (Internal f) -> print_function pp id f | _ -> () - -let print_program pp (prog : BTL.program) = - List.iter (print_globdef pp) prog.prog_defs*) diff --git a/scheduling/RTLpathSE_impl.v b/scheduling/RTLpathSE_impl.v index e21d7cd1..cda1c079 100644 --- a/scheduling/RTLpathSE_impl.v +++ b/scheduling/RTLpathSE_impl.v @@ -825,20 +825,6 @@ Local Hint Resolve hslocal_set_sreg_correct: wlp. (** ** Execution of one instruction *) -(* TODO gourdinl - * This is just useful for debugging fake values hashcode projection *) -Fixpoint check_no_uhid lhsv := - match lhsv with - | HSnil hc => - DO b <~ phys_eq hc unknown_hid;; - assert_b (negb b) "fail no uhid";; - RET tt - | HScons hsv lhsv' hc => - DO b <~ phys_eq hc unknown_hid;; - assert_b (negb b) "fail no uhid";; - check_no_uhid lhsv' - end. - Definition cbranch_expanse (prev: hsistate_local) (cond: condition) (args: list reg): ?? (condition * list_hsval) := match target_cbranch_expanse prev cond args with | Some (cond', vargs) => diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml index 3de82d96..d544e87f 100644 --- a/scheduling/RTLtoBTLaux.ml +++ b/scheduling/RTLtoBTLaux.ml @@ -67,7 +67,6 @@ let translate_function code entry join_points liveness = | Some s -> ( match inst with | Icond (cond, lr, ifso, ifnot, info) -> - (* TODO gourdinl remove this *) assert (s = ifnot); next_nodes := !next_nodes @ non_predicted_successors inst psucc; iinfo.pcond <- info; @@ -100,7 +99,6 @@ let rtl2btl (f : RTL.coq_function) = let liveness = analyze f in let btl = translate_function code entry join_points liveness in let dm = PTree.map (fun n i -> n) btl in - (* TODO gourdinl move elsewhere *) (*debug_flag := true;*) debug "Entry %d\n" (p2i entry); debug "RTL Code: "; -- cgit From 92cca153569e44cd11ba3d1b68c2708c0cc46899 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Wed, 28 Jul 2021 17:00:02 +0200 Subject: ci fix? --- scheduling/BTLmatchRTL.v | 13 ++++++------- scheduling/BTLtoRTLproof.v | 2 ++ scheduling/RTLtoBTLproof.v | 27 +++++++++------------------ 3 files changed, 17 insertions(+), 25 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v index c271ae02..07131893 100644 --- a/scheduling/BTLmatchRTL.v +++ b/scheduling/BTLmatchRTL.v @@ -214,9 +214,9 @@ Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> i | mib_op isfst pc pc' op lr r iinfo: cfg!pc = Some (Iop op lr r pc') -> match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc') - | mib_load isfst pc pc' m a lr r iinfo: - cfg!pc = Some (Iload TRAP m a lr r pc') -> - match_iblock dupmap cfg isfst pc (Bload TRAP m a lr r iinfo) (Some pc') + | mib_load isfst pc pc' trap m a lr r iinfo: + cfg!pc = Some (Iload trap m a lr r pc') -> + match_iblock dupmap cfg isfst pc (Bload trap m a lr r iinfo) (Some pc') | mib_store isfst pc pc' m a lr r iinfo: cfg!pc = Some (Istore m a lr r pc') -> match_iblock dupmap cfg isfst pc (Bstore m a lr r iinfo) (Some pc') @@ -407,7 +407,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) | 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 + if (trapping_mode_eq tm tm') then if (chunk_eq m m') then if (eq_addressing a a') then if (list_eq_dec Pos.eq_dec lr lr') then @@ -417,7 +417,7 @@ Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) else Error (msg "BTL.verify_block: different lr in Bload") else Error (msg "BTL.verify_block: different addressing in Bload") else Error (msg "BTL.verify_block: different chunk in Bload") - else Error (msg "BTL.verify_block: NOTRAP trapping_mode unsupported in Bload") + else Error (msg "BTL.verify_block: different trapping_mode in Bload") | _ => Error (msg "BTL.verify_block: incorrect cfg Bload") end | Bstore m a lr r _ => @@ -520,8 +520,7 @@ Proof. constructor; assumption. - (* Bload *) destruct i; try discriminate. - do 2 (destruct (trapping_mode_eq _ _); try discriminate). - simpl in H. + destruct (trapping_mode_eq _ _); try discriminate. destruct (chunk_eq _ _); try discriminate. destruct (eq_addressing _ _); try discriminate. destruct (list_eq_dec _ _ _); try discriminate. diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 324f1d14..4ef55928 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -232,6 +232,8 @@ Proof. inv MIB. exists pc'; split; auto; constructor. apply plus_one. inv LOAD. eapply exec_Iload; eauto. + 2: destruct (eval_addressing _ _ _ _) eqn:EVAL; + [ eapply exec_Iload_notrap2 | eapply exec_Iload_notrap1]; eauto. all: erewrite <- eval_addressing_preserved; eauto; intros; rewrite symbols_preserved; trivial. - (* exec_store *) diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v index 13ba5a29..0ca93bce 100644 --- a/scheduling/RTLtoBTLproof.v +++ b/scheduling/RTLtoBTLproof.v @@ -624,26 +624,17 @@ Proof. - (* mib_seq *) inv IS_EXPD; try discriminate. inv H; simpl in *; try congruence; - inv STEP; try_simplify_someHyps; eauto. - + (* Bnop is_rtl *) - intros; eapply match_strong_state_simu; eauto. - + (* Bop *) - intros; eapply match_strong_state_simu; eauto. - econstructor; eauto. + inv STEP; try_simplify_someHyps; + intros; eapply match_strong_state_simu; eauto; + econstructor; eauto. + { (* Bop *) erewrite eval_operation_preserved in H12. erewrite H12 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. - + (* Bload *) - intros; eapply match_strong_state_simu; eauto. - econstructor; eauto. - erewrite eval_addressing_preserved in H12. - erewrite H12, H13 in BTL_RUN; simpl in BTL_RUN; auto. - intros; rewrite <- symbols_preserved; trivial. - + (* Bstore *) - intros; eapply match_strong_state_simu; eauto. - econstructor; eauto. - erewrite eval_addressing_preserved in H12. - erewrite H12, H13 in BTL_RUN; simpl in BTL_RUN; auto. + intros; rewrite <- symbols_preserved; trivial. } + all: (* Bload / Bstore *) + erewrite eval_addressing_preserved in H12; + try erewrite H12 in BTL_RUN; try erewrite H13 in BTL_RUN; + simpl in BTL_RUN; try destruct trap; auto; intros; rewrite <- symbols_preserved; trivial. - (* mib_cond *) inv IS_EXPD; try discriminate. -- cgit From 806a3844f154fb76ce45c446c35d846c58b942fc Mon Sep 17 00:00:00 2001 From: Leo Gourdin Date: Mon, 2 Aug 2021 11:59:11 +0200 Subject: non-trapping loads fix --- scheduling/BTLScheduleraux.ml | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml index c8c4e0d3..00be2aa7 100644 --- a/scheduling/BTLScheduleraux.ml +++ b/scheduling/BTLScheduleraux.ml @@ -120,9 +120,17 @@ let rec do_schedule btl = function | [] -> btl | (n, ibf) :: blks -> let code_exp = expanse n ibf btl in - let code_nt = turn_all_loads_nontrap n ibf code_exp in - let btl' = schedule_blk n ibf code_nt in - do_schedule btl' blks + let ibf_exp = get_some @@ PTree.get n code_exp in + let code_nt = turn_all_loads_nontrap n ibf_exp code_exp in + let ibf_nt = get_some @@ PTree.get n code_nt in + let btl' = schedule_blk n ibf_nt code_nt in + begin + (*debug_flag := true;*) + print_btl_code stderr code_nt; + print_btl_code stderr btl'; + (*debug_flag := false;*) + do_schedule btl' blks + end let btl_scheduler f = let btl = f.fn_code in -- cgit From 54a22d92bc18fa3ece958a097844caa5e7b2e0c5 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 2 Aug 2021 14:03:45 +0200 Subject: CI fix on arm --- scheduling/BTLtoRTLproof.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) (limited to 'scheduling') diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v index 4ef55928..cbdc81bd 100644 --- a/scheduling/BTLtoRTLproof.v +++ b/scheduling/BTLtoRTLproof.v @@ -232,8 +232,8 @@ Proof. inv MIB. exists pc'; split; auto; constructor. apply plus_one. inv LOAD. eapply exec_Iload; eauto. - 2: destruct (eval_addressing _ _ _ _) eqn:EVAL; - [ eapply exec_Iload_notrap2 | eapply exec_Iload_notrap1]; eauto. + all: try (destruct (eval_addressing _ _ _ _) eqn:EVAL; + [ eapply exec_Iload_notrap2 | eapply exec_Iload_notrap1]; eauto). all: erewrite <- eval_addressing_preserved; eauto; intros; rewrite symbols_preserved; trivial. - (* exec_store *) -- cgit