From 841ee5b298bb80d8bc3d821deb2f03d8ba0a5e61 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 5 Sep 2019 18:52:39 +0200 Subject: A variant of RTL with a path-step semantics --- backend/RTLpath.v | 476 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 476 insertions(+) create mode 100644 backend/RTLpath.v diff --git a/backend/RTLpath.v b/backend/RTLpath.v new file mode 100644 index 00000000..bc7f5a03 --- /dev/null +++ b/backend/RTLpath.v @@ -0,0 +1,476 @@ +(** We introduce a data-structure for partitioning the nodes of a RTL CFG into a set of "traces" (in the sense of "trace-scheduling") + Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert: + a "path" is simply a list of successive nodes in the CFG. + + Here, we extend syntactically the notion of RTL programs with a structure of "path_map". + + Our goal is to provide two bisimulable semantics: + - one is simply the RTL semantics + - the other is based on a notion of "path-step": each path is executed in a single step. + + Remark that all analyses on RTL programs should thus be appliable for "free" also for RTLpath programs ! +*) + +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL Linking. + +Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end) + (at level 200, X ident, A at level 100, B at level 200) + : option_monad_scope. + +Local Open Scope option_monad_scope. + +(** * Syntax of RTLpath programs *) + +(** Internal instruction = instruction with a default successor in a path. *) + +Definition default_succ (i: instruction): option node := + match i with + | Inop s => Some s + | Iop op args res s => Some s + | Iload chunk addr args dst s => Some s + | Istore chunk addr args src s => Some s + | Icond cond args ifso ifnot => Some ifnot + | _ => None (* TODO: we could choose a successor for jumptable ? *) + end. + +Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *) + match i with + | Icond cond args ifso ifnot => Some ifso + | _ => None + end. + +(** Our notion of path. + + We do not formally require that the set of path is a partition of the CFG. + path may have intersections ! + + Moreover, we do not formally require that path have a single entry-point (a superblock structure) + + But, in practice, these properties are probably necessary in order to ensure the success of dynamic verification of scheduling. + + Here: we only require that each exit-point of a path is the entry-point of a path + (and that internal node of a path are internal instructions) +*) + +Definition path_map: Type := PTree.t nat. + +(* +Definition path_entry (c:code) (pm: path_map) (n: node): Prop + := pm!n = None -> c!n=None. +*) + +Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := + | wf_last_node i pc: + c!pc = Some i -> + (forall n, List.In n (successors_instr i) -> pm!n = None -> c!n=None) -> + wellformed_path c pm 0 pc + | wf_internal_node path i pc pc': + c!pc = Some i -> + default_succ i = Some pc' -> + (forall n, early_exit i = Some n -> pm!n = None -> c!n=None) -> + wellformed_path c pm path pc' -> + wellformed_path c pm (S path) pc. + +(* all paths defined from the path_map are wellformed *) +Definition wellformed_path_map (c:code) (pm: path_map): Prop := + forall n path, pm!n = Some path -> wellformed_path c pm path n. + +(** We "extend" the notion of RTL program with the additional structure for path. + + There is thus a trivial "forgetful functor" from RTLpath programs to RTL ones. +*) + +Record function : Type := + { fn_RTL:> RTL.function; + fn_path: path_map; + (* condition 1 below: the entry-point of the code is an entry-point of a path *) + fn_entry_point_wf: fn_path!(fn_RTL.(fn_entrypoint)) = None -> fn_RTL.(fn_code)!(fn_RTL.(fn_entrypoint)) = None; + (* condition 2 below: the path_map is well-formed *) + fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path + }. + + +Definition fundef := AST.fundef function. +Definition program := AST.program fundef unit. +Definition genv := Genv.t fundef unit. + +Definition fundef_RTL (fu: fundef) : RTL.fundef := + match fu with + | Internal f => Internal f.(fn_RTL) + | External ef => External ef + end. +Coercion fundef_RTL: fundef >-> RTL.fundef. + +Definition program_RTL (p: program) : RTL.program := transform_program fundef_RTL p. +Coercion program_RTL: program >-> RTL.program. + +(** * Path-step semantics of RTLpath programs *) + +(* Semantics of internal instructions (mimicking RTL semantics) *) + +Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. + +Inductive stackframe : Type := + | Stackframe + (res: reg) (**r where to store the result *) + (f: function) (**r calling function *) + (sp: val) (**r stack pointer in calling function *) + (pc: node) (**r program point in calling function *) + (rs: regset) (**r register state in calling function *) + . + +Definition stf_RTL (st: stackframe): RTL.stackframe := + match st with + | Stackframe res f sp pc rs => RTL.Stackframe res f sp pc rs + end. + +Fixpoint stack_RTL (stack: list stackframe): list RTL.stackframe := + match stack with + | nil => nil + | cons stf stack' => cons (stf_RTL stf) (stack_RTL stack') + end. + +Inductive state : Type := + | State + (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 *) + | Callstate + (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 *) + | Returnstate + (stack: list stackframe) (**r call stack *) + (v: val) (**r return value for the call *) + (m: mem) (**r memory state *) + . + +Definition state_RTL (s: state): RTL.state := + match s with + | State stack f sp pc rs m => RTL.State (stack_RTL stack) f sp pc rs m + | Callstate stack f args m => RTL.Callstate (stack_RTL stack) f args m + | Returnstate stack v m => RTL.Returnstate (stack_RTL stack) v m + end. +Coercion state_RTL: state >-> RTL.state. + +Definition step_istate (ge: RTL.genv) (i: instruction) (stack: list stackframe) (f: function) (sp: val) (rs: regset) (m: mem): option istate := + match i with + | Inop pc' => Some (ist true pc' rs m) + | Iop op args res pc' => + SOME v <- eval_operation ge sp op rs##args m IN + Some (ist true pc' (rs#res <- v) m) + | Iload chunk addr args dst pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME v <- Mem.loadv chunk m a IN + Some (ist true pc' (rs#dst <- v) m) + | Istore chunk addr args src pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME m' <- Mem.storev chunk m a rs#src IN + Some (ist true pc' rs m') + | Icond cond args ifso ifnot => + SOME b <- eval_condition cond rs##args m IN + Some (ist (negb b) (if b then ifso else ifnot) rs m) + | _ => None (* TODO jumptable ? *) + end. + +(** Execution of a path in a single step *) + +Fixpoint path_step_istate ge (path:nat) stack (f: function) sp rs m pc: option istate := + match path with + | O => Some (ist true pc rs m) + | S p => + SOME i <- (fn_code f)!pc IN + SOME st <- step_istate ge i stack f sp rs m IN + if st.(continue) then + path_step_istate ge p stack f sp st.(the_rs) st.(the_mem) st.(the_pc) + else + Some st + end. + +Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef := + match ros with + | inl r => Genv.find_funct pge rs#r + | inr symb => + match Genv.find_symbol pge symb with + | None => None + | Some b => Genv.find_funct_ptr pge b + end + end. + +Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := + | exec_istate i sp pc rs m st: + (fn_code f)!pc = Some i -> + step_istate ge i stack f sp rs m = Some st -> + path_last_step ge pge stack f sp pc rs m + E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + | exec_Icall sp pc rs m sig ros args res pc' fd: + (fn_code f)!pc = Some(Icall sig ros args res pc') -> + find_function pge ros rs = Some fd -> + funsig fd = sig -> + path_last_step ge pge stack f sp pc rs m + E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) + | exec_Itailcall stk pc rs m sig ros args fd m': + (fn_code f)!pc = Some(Itailcall sig ros args) -> + find_function pge ros rs = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Callstate stack fd rs##args m') + | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': + (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> + eval_builtin_args ge (fun r => rs#r) sp m args vargs -> + external_call ef ge vargs m t vres m' -> + path_last_step ge pge stack f sp pc rs m + t (State stack f sp pc' (regmap_setres res vres rs) m') + | exec_Ijumptable sp pc rs m arg tbl n pc': + (fn_code f)!pc = Some(Ijumptable arg tbl) -> + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + path_last_step ge pge stack f sp pc rs m + E0 (State stack f sp pc' rs m) + | exec_Ireturn stk pc rs m or m': + (fn_code f)!pc = Some(Ireturn or) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Returnstate stack (regmap_optget or Vundef rs) m'). + +Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := + | exec_early_exit st: + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = false -> + path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + | exec_normal_exit st t s: + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = true -> + path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> + path_step ge pge path stack f sp rs m pc t s. + +Inductive step ge pge: state -> trace -> state -> Prop := + | exec_path path stack f sp rs m pc t s: + (fn_path f)!pc = Some path -> + path_step ge pge path stack f sp rs m pc t s -> + step ge pge (State stack f sp pc rs m) t s + | exec_function_internal s f args m m' stk: + Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) -> + step ge pge (Callstate s (Internal f) args m) + E0 (State s + f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + m') + | exec_function_external s ef args res t m m': + external_call ef ge args m t res m' -> + step ge pge (Callstate s (External ef) args m) + t (Returnstate s res m') + | exec_return res f sp pc rs s vres m: + step ge pge (Returnstate (Stackframe res f sp pc rs :: s) vres m) + E0 (State s f sp pc (rs#res <- vres) m). + +Definition initial_state (p: program) (st:state): Prop + := RTL.initial_state p st. + +Definition final_state (st: state) (i:int): Prop + := RTL.final_state st i. + +Definition semantics (p: program) := + Semantics (step (Genv.globalenv (program_RTL p))) (initial_state p) final_state (Genv.globalenv p). + + +(** * Proving the bisimulation between (semantics p) and (RTL.semantics p). *) + +(** ** Preliminaries: simple tactics for option-monad *) + +Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B): + (forall x, e = Some x -> P (f x)) -> (P None) -> (P (SOME x <- e IN f x)). +Proof. + intros; destruct e; simpl; auto. +Qed. + +Local Ltac inversion_SOME x := + try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]). + +Ltac simplify_someHyp := + match goal with + | H: None = Some _ |- _ => inversion H; subst + | H: Some _ = None |- _ => inversion H; subst + | H: ?t = ?t |- _ => clear H + | H: Some _ = Some _ |- _ => inversion H; subst + | H: Some _ <> None |- _ => clear H + | H: None <> Some _ |- _ => clear H + | H: _ = Some _ |- _ => (rewrite !H in * |- * || generalize H); clear H + end. + +Ltac simplify_someHyps := + repeat (simplify_someHyp; simpl in * |- *). + +Ltac try_simplify_someHyps := + try (intros; simplify_someHyps; eauto). + +Ltac simplify_SOME x := + (repeat inversion_SOME x); try_simplify_someHyps. + + +Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. + +(** ** The easy way: Forward simulation of RTLpath by RTL *) + +Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. + +Lemma step_istate_correct ge i stack f sp rs m st : + step_istate ge i stack f sp rs m = Some st -> + forall pc, (fn_code f)!pc = Some i -> + RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. +Qed. + +(* FIXME: useful ? +Lemma step_istate_internal ge i stack f sp rs m : + step_istate ge i stack f sp rs m <> None -> + default_succ i <> None. +Proof. + destruct i; simpl; try congruence; simplify_SOME x. +Qed. +*) + +(* FIXME: useful ? +Lemma step_istate_false_continue ge i stack f sp rs m st : + step_istate ge i stack f sp rs m = Some st -> + st.(continue) = false -> + exists cond, exists args, exists ifnot, i = Icond cond args st.(the_pc) ifnot /\ st.(the_rs) = rs /\ st.(the_mem) = m. +Proof. + destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. + destruct b; simpl in * |- *; try congruence. + repeat eexists; eauto. +Qed. +*) + +Local Hint Resolve star_refl. + +Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + induction path; simpl; try_simplify_someHyps. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:cont. + + intros; eapply star_step. + - eapply step_istate_correct; eauto. + - simpl; eauto. + - auto. + + intros; simplify_someHyp; eapply star_step. + - eapply step_istate_correct; eauto. + - simpl; eauto. + - auto. +Qed. + +Lemma path_step_istate_correct_early_exit ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = false -> + plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + destruct path; simpl; try_simplify_someHyps; try congruence. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:cont. + + intros; eapply plus_left. + - eapply step_istate_correct; eauto. + - eapply path_step_istate_correct; eauto. + - auto. + + intros X; inversion X; subst. + eapply plus_one. + eapply step_istate_correct; eauto. +Qed. + + +Section PRESERVATION. + +Variable p: program. + +Lemma match_prog_RTL: match_program (fun _ f tf => tf = fundef_RTL f) eq p (program_RTL p). +Proof. + eapply match_transform_program; eauto. +Qed. + +Let pge := Genv.globalenv p. +Let ge := Genv.globalenv (program_RTL p). + +Lemma senv_preserved: Senv.equiv pge ge. +Proof (Genv.senv_match match_prog_RTL). + +Lemma symbols_preserved s: Genv.find_symbol ge s = Genv.find_symbol pge s. +Proof (Genv.find_symbol_match match_prog_RTL s). + +Lemma find_function_RTL_match ros rs fd: + find_function pge ros rs = Some fd -> RTL.find_function ge ros rs = Some (fundef_RTL fd). +Proof. + destruct ros; simpl. + + destruct (rs#r); simpl; try congruence. + destruct (Ptrofs.eq_dec i Ptrofs.zero); simpl; try congruence. + intros H; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. + destruct 1 as (cuint & tf & H1 & H2 & H3); subst; auto. + + rewrite symbols_preserved. + destruct (Genv.find_symbol pge i); simpl; try congruence. + intros H; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. + destruct 1 as (cuint & tf & H1 & H2 & H3); subst; auto. +Qed. + +Local Hint Resolve step_istate_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. + +Lemma path_last_step_correct stack f sp pc rs m t s: + path_last_step ge pge stack f sp pc rs m t s -> + RTL.step ge (State stack f sp pc rs m) t s. +Proof. + destruct 1; try (eapply step_istate_correct); simpl; eauto. +Qed. + +Lemma path_step_correct path stack f sp pc rs m t s: + path_step ge pge path stack f sp rs m pc t s -> + plus RTL.step ge (State stack f sp pc rs m) t s. +Proof. + destruct 1. + + eapply path_step_istate_correct_early_exit; eauto. + + eapply plus_right. + eapply path_step_istate_correct; eauto. + eapply path_last_step_correct; eauto. + auto. +Qed. + +Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return. + +Lemma step_correct s t s': step ge pge s t s' -> plus RTL.step ge s t s'. +Proof. + destruct 1; try (eapply path_step_correct); simpl; eauto. +Qed. + +Theorem RTLpath_correct: forward_simulation (semantics p) (RTL.semantics p). +Proof. + eapply forward_simulation_plus with (match_states := fun s1 s2 => s2 = state_RTL s1); simpl; auto. + - apply senv_preserved. + - unfold initial_state; eauto. + - unfold final_state; intros; subst; eauto. + - intros; subst. admit. +Abort. + +End PRESERVATION. + +(** The hard way: Forward simulation of RTL by RTLpath *) + +Lemma step_istate_complete ge t i stack f sp rs m pc s': + RTL.step ge (State stack f sp pc rs m) t s' -> + (fn_code f)!pc = Some i -> + default_succ i <> None -> + t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; + (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. +Qed. -- cgit From b48eb874413ff789bc97b60bdfe6a34acedd5abe Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 6 Sep 2019 07:36:05 +0200 Subject: RTLpath: schema de simulation pour completeness --- backend/RTLpath.v | 63 +++++++++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 57 insertions(+), 6 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index bc7f5a03..885f8f53 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -320,7 +320,11 @@ Ltac simplify_SOME x := Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. -(** ** The easy way: Forward simulation of RTLpath by RTL *) +(** ** The easy way: Forward simulation of RTLpath by RTL + +This way can be viewed as a correctness property: all transitions in RTLpath are valid RTL transitions ! + +*) Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. @@ -392,7 +396,7 @@ Proof. Qed. -Section PRESERVATION. +Section CORRECTNESS. Variable p: program. @@ -458,12 +462,32 @@ Proof. - apply senv_preserved. - unfold initial_state; eauto. - unfold final_state; intros; subst; eauto. - - intros; subst. admit. -Abort. + - intros; subst. eexists; intuition. + eapply step_correct; eauto. +Qed. + +End CORRECTNESS. -End PRESERVATION. -(** The hard way: Forward simulation of RTL by RTLpath *) +(** The hard way: Forward simulation of RTL by RTLpath + +This way can be viewed as a completeness property: all transitions in RTL can be represented as RTLpath transitions ! + +*) + +(* This lemma is probably needed to compose a pass from RTL -> RTLpath with other passes.*) +Lemma match_RTL_prog {LA: Linker fundef} {LV: Linker unit} p: match_program (fun _ f tf => f = fundef_RTL tf) eq (program_RTL p) p. +Proof. + unfold match_program, match_program_gen; intuition. + unfold program_RTL at 2; simpl. + generalize (prog_defs p). + induction l as [|a l]; simpl; eauto. + destruct a; simpl. + intros; eapply list_forall2_cons; eauto. + unfold match_ident_globdef; simpl; intuition; destruct g as [f|v]; simpl; eauto. + eapply match_globdef_var. destruct v; eauto. +Qed. + Lemma step_istate_complete ge t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> @@ -474,3 +498,30 @@ Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. Qed. + + +Section COMPLETENESS. + +Variable p: program. + +Let pge := Genv.globalenv p. +Let ge := Genv.globalenv (program_RTL p). + +(* stub *) +Variable match_states: nat -> RTL.state -> state -> Prop. + +Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). +Proof. + eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). + constructor 1. + - apply lt_wf. +(* + - unfold initial_state; eauto. + - unfold final_state; intros; subst; eauto. + - intros; subst. eexists; intuition. + eapply step_correct; eauto. + - apply senv_preserved. +*) +Abort. + +End COMPLETENESS. \ No newline at end of file -- cgit From 2292a56e3fda3d7a391a12ab0460dad466bac45a Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 6 Sep 2019 10:42:23 +0200 Subject: RTLpath: a first idea of match_states --- backend/RTLpath.v | 42 +++++++++++++++++++++++++++++++----------- 1 file changed, 31 insertions(+), 11 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index 885f8f53..21a03ca2 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -507,21 +507,41 @@ Variable p: program. Let pge := Genv.globalenv p. Let ge := Genv.globalenv (program_RTL p). -(* stub *) -Variable match_states: nat -> RTL.state -> state -> Prop. +Definition is_inst (s: RTL.state): bool := + match s with + | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => true | None => false end + | _ => false + end. + +Inductive match_inst_states (n: nat) (s1:RTL.state): state -> Prop := + | State_match path stack f sp pc rs m s2: + (fn_path f)!pc = Some path -> (n <= path)%nat -> + path_step_istate ge (path-n) stack f sp rs m pc = Some s2 -> + s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> + match_inst_states n s1 (State stack f sp pc rs m). + +Definition match_states (n: nat) (s1:RTL.state) (s2:state): Prop := + if is_inst s1 + then match_inst_states n s1 s2 + else s1 = state_RTL s2. + +Lemma find_funct_ptr_RTL_preserv b f: + Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0). +Admitted. Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). Proof. eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). - constructor 1. + constructor 1; simpl. - apply lt_wf. -(* - - unfold initial_state; eauto. - - unfold final_state; intros; subst; eauto. - - intros; subst. eexists; intuition. - eapply step_correct; eauto. - - apply senv_preserved. -*) -Abort. + - unfold initial_state, match_states. destruct 1; simpl; exists O. + destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. + exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. + - unfold final_state, match_states. intros i s1 s2 r H1 H2; destruct H2. + destruct s2; simpl in * |- *; inversion H1; subst. + constructor. + - admit. + - intros id; destruct (senv_preserved p); simpl in * |-. intuition. +Admitted. End COMPLETENESS. \ No newline at end of file -- cgit From 7d695d0ee45f722bfae9a5bafd182bb1212c7c76 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 6 Sep 2019 15:26:25 +0200 Subject: RTLpath: beginning of the proof (in backward style) --- backend/RTLpath.v | 113 ++++++++++++++++++++++++++++++++++++++++++++++++------ 1 file changed, 102 insertions(+), 11 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index 21a03ca2..f3b05220 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -55,28 +55,36 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, (and that internal node of a path are internal instructions) *) + +(* By convention, we say that node [n] of the CFG is an entry-point of a path + if it is an entry of the path_map. + + Such a path of entry [n] is defined from a natural [path] representing the [path] default-successors of [n]. + + Remark: a path can loop several times in the CFG. + +*) + Definition path_map: Type := PTree.t nat. -(* Definition path_entry (c:code) (pm: path_map) (n: node): Prop := pm!n = None -> c!n=None. -*) Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := | wf_last_node i pc: c!pc = Some i -> - (forall n, List.In n (successors_instr i) -> pm!n = None -> c!n=None) -> + (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> wellformed_path c pm 0 pc | wf_internal_node path i pc pc': c!pc = Some i -> default_succ i = Some pc' -> - (forall n, early_exit i = Some n -> pm!n = None -> c!n=None) -> + (forall n, early_exit i = Some n -> path_entry c pm n) -> wellformed_path c pm path pc' -> wellformed_path c pm (S path) pc. (* all paths defined from the path_map are wellformed *) Definition wellformed_path_map (c:code) (pm: path_map): Prop := - forall n path, pm!n = Some path -> wellformed_path c pm path n. + forall n path, pm!n = Some path -> wellformed_path c pm path n. (** We "extend" the notion of RTL program with the additional structure for path. @@ -87,7 +95,7 @@ Record function : Type := { fn_RTL:> RTL.function; fn_path: path_map; (* condition 1 below: the entry-point of the code is an entry-point of a path *) - fn_entry_point_wf: fn_path!(fn_RTL.(fn_entrypoint)) = None -> fn_RTL.(fn_code)!(fn_RTL.(fn_entrypoint)) = None; + fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); (* condition 2 below: the path_map is well-formed *) fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path }. @@ -499,6 +507,12 @@ Proof. (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. Qed. +Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. +Proof. + unfold path_entry; destruct pm!n; eauto. intuition congruence. +Qed. +Local Hint Resolve fn_entry_point_wf fn_path_wf. + Section COMPLETENESS. @@ -507,6 +521,14 @@ Variable p: program. Let pge := Genv.globalenv p. Let ge := Genv.globalenv (program_RTL p). +Lemma find_funct_ptr_RTL_preserv b f: + Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0). +Proof. + intros; exploit (Genv.find_funct_ptr_match (match_RTL_prog p)); eauto. + destruct 1 as (cunit & tf & X & Y & Z); subst. + eauto. +Qed. + Definition is_inst (s: RTL.state): bool := match s with | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => true | None => false end @@ -515,7 +537,8 @@ Definition is_inst (s: RTL.state): bool := Inductive match_inst_states (n: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: - (fn_path f)!pc = Some path -> (n <= path)%nat -> + (fn_path f)!pc = Some path -> + (n <= path)%nat -> path_step_istate ge (path-n) stack f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> match_inst_states n s1 (State stack f sp pc rs m). @@ -525,10 +548,77 @@ Definition match_states (n: nat) (s1:RTL.state) (s2:state): Prop := then match_inst_states n s1 s2 else s1 = state_RTL s2. -Lemma find_funct_ptr_RTL_preserv b f: - Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0). +Lemma path_step_complete stack f sp rs m pc t s1' i path st: + (fn_path f)!pc = Some path -> + (i <= path)%nat -> + path_step_istate ge (path-i) stack f sp rs m pc = Some st -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + exists i' s2', + (path_step ge pge path stack f sp rs m pc t s2' + \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (i' < i)%nat)) /\ match_states i' s1' s2'. Admitted. +Lemma step_noninst_complete s1 t s1' s2: + is_inst s1 = false -> + s1 = state_RTL s2 -> + RTL.step ge s1 t s1' -> + exists s2', step ge pge s2 t s2' /\ exists i', match_states i' s1' s2'. +Proof. + intros H0 H1 H2; destruct s2; subst; simpl in * |- *; try congruence; + inversion H2; clear H2; subst; try_simplify_someHyps; try congruence. + + (* exec_function_internal *) + destruct f; simpl in H3; inversion H3; subst; clear H3. + eexists; constructor 1. + * eapply exec_function_internal; eauto. + * unfold match_states; simpl. + destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X; [|exists O; auto]. + exploit initialize_path; eauto. + destruct 1 as (path & Hpath). + exists path. + econstructor; eauto. + - cutrewrite (path-path=O)%nat; simpl; eauto. + omega. + - simpl; auto. + + (* exec_function_external *) + destruct f; simpl in H3 |-; inversion H3; subst; clear H3. + eexists; constructor 1. + * apply exec_function_external; eauto. + * unfold match_states; simpl. exists O; auto. + + (* exec_return *) + destruct stack eqn: Hstack; simpl in H1; inversion H1; clear H1; subst. + destruct s0 eqn: Hs0; simpl in H0; inversion H0; clear H0; subst. + eexists; constructor 1. + * apply exec_return. + * unfold match_states; simpl. + admit. + (* FIXME: il nous faut dans le match_states (et donc en hypothèse de ce lemme) + un invariant additionnel qui dit que tous les points de retour + dans le stack_frame sont des points d'entrée valides ! + *) +Admitted. + + +Lemma step_complete s1 t s1' i s2: + match_states i s1 s2 -> + RTL.step ge s1 t s1' -> + exists i' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (i' < i)%nat)) /\ match_states i' s1' s2'. +Proof. + unfold match_states at 1. destruct (is_inst s1) eqn: His1. + - destruct 1 as [path stack f sp pc rs m s2 X X0 X1 X2]; subst. +(* FIXME: USELESS HERE ? + simpl in His1. + destruct ((fn_code f) ! (the_pc s2)) eqn:Hpc; simpl; try congruence. +*) + intros; exploit path_step_complete; eauto. + destruct 1 as (i' & s2' & H0 & H1). + repeat eexists; eauto. + destruct H0. + + exploit exec_path; eauto. + + intuition auto. + - intros; exploit step_noninst_complete; eauto. + firstorder. +Qed. + Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). Proof. eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). @@ -540,8 +630,9 @@ Proof. - unfold final_state, match_states. intros i s1 s2 r H1 H2; destruct H2. destruct s2; simpl in * |- *; inversion H1; subst. constructor. - - admit. + - Local Hint Resolve plus_one star_refl. + intros; exploit step_complete; eauto. firstorder (subst; eauto). - intros id; destruct (senv_preserved p); simpl in * |-. intuition. -Admitted. +Qed. End COMPLETENESS. \ No newline at end of file -- cgit From 38f306593a03fb291ac267398b521b7c4dc68f6a Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 6 Sep 2019 15:35:01 +0200 Subject: RTLpath: use idx for indices in match_states relation --- backend/RTLpath.v | 30 +++++++++++++++--------------- 1 file changed, 15 insertions(+), 15 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index f3b05220..fb61cad7 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -535,34 +535,34 @@ Definition is_inst (s: RTL.state): bool := | _ => false end. -Inductive match_inst_states (n: nat) (s1:RTL.state): state -> Prop := +Inductive match_inst_states (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> - (n <= path)%nat -> - path_step_istate ge (path-n) stack f sp rs m pc = Some s2 -> + (idx <= path)%nat -> + path_step_istate ge (path-idx) stack f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> - match_inst_states n s1 (State stack f sp pc rs m). + match_inst_states idx s1 (State stack f sp pc rs m). -Definition match_states (n: nat) (s1:RTL.state) (s2:state): Prop := +Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := if is_inst s1 - then match_inst_states n s1 s2 + then match_inst_states idx s1 s2 else s1 = state_RTL s2. -Lemma path_step_complete stack f sp rs m pc t s1' i path st: +Lemma path_step_complete stack f sp rs m pc t s1' idx path st: (fn_path f)!pc = Some path -> - (i <= path)%nat -> - path_step_istate ge (path-i) stack f sp rs m pc = Some st -> + (idx <= path)%nat -> + path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - exists i' s2', + exists idx' s2', (path_step ge pge path stack f sp rs m pc t s2' - \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (i' < i)%nat)) /\ match_states i' s1' s2'. + \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Admitted. Lemma step_noninst_complete s1 t s1' s2: is_inst s1 = false -> s1 = state_RTL s2 -> RTL.step ge s1 t s1' -> - exists s2', step ge pge s2 t s2' /\ exists i', match_states i' s1' s2'. + exists s2', step ge pge s2 t s2' /\ exists idx, match_states idx s1' s2'. Proof. intros H0 H1 H2; destruct s2; subst; simpl in * |- *; try congruence; inversion H2; clear H2; subst; try_simplify_someHyps; try congruence. @@ -598,10 +598,10 @@ Proof. Admitted. -Lemma step_complete s1 t s1' i s2: - match_states i s1 s2 -> +Lemma step_complete s1 t s1' idx s2: + match_states idx s1 s2 -> RTL.step ge s1 t s1' -> - exists i' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (i' < i)%nat)) /\ match_states i' s1' s2'. + exists idx' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. unfold match_states at 1. destruct (is_inst s1) eqn: His1. - destruct 1 as [path stack f sp pc rs m s2 X X0 X1 X2]; subst. -- cgit From ee0804e93395a6c25b8a3770bb6f00e51fe2a297 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 6 Sep 2019 19:21:19 +0200 Subject: RTLpath: defines wellformation of stacks. --- backend/RTLpath.v | 80 ++++++++++++++++++++++++++++++++++++++++--------------- 1 file changed, 59 insertions(+), 21 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index fb61cad7..426ec3e3 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -127,7 +127,7 @@ Inductive stackframe : Type := (f: function) (**r calling function *) (sp: val) (**r stack pointer in calling function *) (pc: node) (**r program point in calling function *) - (rs: regset) (**r register state in calling function *) + (rs: regset) (**r register state in calling function *) . Definition stf_RTL (st: stackframe): RTL.stackframe := @@ -529,6 +529,33 @@ Proof. eauto. Qed. +Definition wf_stf (st: stackframe): Prop := + match st with + | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc + end. + +Definition wf_stackframe (stack: list stackframe): Prop := + forall st, List.In st stack -> wf_stf st. + +Lemma wf_stackframe_nil: wf_stackframe nil. +Proof. + unfold wf_stackframe; simpl. tauto. +Qed. +Local Hint Resolve wf_stackframe_nil. + +Lemma wf_stackframe_cons st stack: + wf_stackframe (st::stack) <-> (wf_stf st) /\ wf_stackframe stack. +Proof. + unfold wf_stackframe; simpl; intuition (subst; auto). +Qed. + +Definition stack_of (s: state): list stackframe := + match s with + | State stack f sp pc rs m => stack + | Callstate stack f args m => stack + | Returnstate stack v m => stack + end. + Definition is_inst (s: RTL.state): bool := match s with | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => true | None => false end @@ -544,15 +571,17 @@ Inductive match_inst_states (idx: nat) (s1:RTL.state): state -> Prop := match_inst_states idx s1 (State stack f sp pc rs m). Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := - if is_inst s1 - then match_inst_states idx s1 s2 - else s1 = state_RTL s2. + (if is_inst s1 + then match_inst_states idx s1 s2 + else s1 = state_RTL s2) + /\ wf_stackframe (stack_of s2). Lemma path_step_complete stack f sp rs m pc t s1' idx path st: + path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (idx <= path)%nat -> - path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + wf_stackframe stack -> exists idx' s2', (path_step ge pge path stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. @@ -562,9 +591,10 @@ Lemma step_noninst_complete s1 t s1' s2: is_inst s1 = false -> s1 = state_RTL s2 -> RTL.step ge s1 t s1' -> + wf_stackframe (stack_of s2) -> exists s2', step ge pge s2 t s2' /\ exists idx, match_states idx s1' s2'. Proof. - intros H0 H1 H2; destruct s2; subst; simpl in * |- *; try congruence; + intros H0 H1 H2 WFSTACK; destruct s2; subst; simpl in * |- *; try congruence; inversion H2; clear H2; subst; try_simplify_someHyps; try congruence. + (* exec_function_internal *) destruct f; simpl in H3; inversion H3; subst; clear H3. @@ -574,7 +604,7 @@ Proof. destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X; [|exists O; auto]. exploit initialize_path; eauto. destruct 1 as (path & Hpath). - exists path. + exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. omega. @@ -590,12 +620,17 @@ Proof. eexists; constructor 1. * apply exec_return. * unfold match_states; simpl. - admit. - (* FIXME: il nous faut dans le match_states (et donc en hypothèse de ce lemme) - un invariant additionnel qui dit que tous les points de retour - dans le stack_frame sont des points d'entrée valides ! - *) -Admitted. + rewrite wf_stackframe_cons in WFSTACK. + destruct WFSTACK as (H0 & H1); simpl in H0. + destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; auto]. + exploit initialize_path; eauto. + destruct 1 as (path & Hpath). + exists path. constructor; auto. + econstructor; eauto. + - cutrewrite (path-path=O)%nat; simpl; eauto. + omega. + - simpl; auto. +Qed. Lemma step_complete s1 t s1' idx s2: @@ -603,18 +638,19 @@ Lemma step_complete s1 t s1' idx s2: RTL.step ge s1 t s1' -> exists idx' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. - unfold match_states at 1. destruct (is_inst s1) eqn: His1. - - destruct 1 as [path stack f sp pc rs m s2 X X0 X1 X2]; subst. -(* FIXME: USELESS HERE ? + unfold match_states at 1. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. + - destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; subst. + clear His1. simpl in * |- *. +(* FIXME: His1 USELESS HERE ? simpl in His1. destruct ((fn_code f) ! (the_pc s2)) eqn:Hpc; simpl; try congruence. *) intros; exploit path_step_complete; eauto. destruct 1 as (i' & s2' & H0 & H1). - repeat eexists; eauto. + eexists; eexists; eauto. destruct H0. + exploit exec_path; eauto. - + intuition auto. + + intuition subst; eauto. - intros; exploit step_noninst_complete; eauto. firstorder. Qed. @@ -627,11 +663,13 @@ Proof. - unfold initial_state, match_states. destruct 1; simpl; exists O. destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. - - unfold final_state, match_states. intros i s1 s2 r H1 H2; destruct H2. - destruct s2; simpl in * |- *; inversion H1; subst. + - unfold final_state, match_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. + destruct s2; simpl in * |- *; inversion H0; subst. constructor. - Local Hint Resolve plus_one star_refl. - intros; exploit step_complete; eauto. firstorder (subst; eauto). + intros; exploit step_complete; eauto. + destruct 1 as (idx' & s2' & X). + exists idx'. exists s2'. intuition (subst; eauto). - intros id; destruct (senv_preserved p); simpl in * |-. intuition. Qed. -- cgit From 15730206844eb869b0de5e82222bade7c479f231 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 7 Sep 2019 08:42:03 +0200 Subject: RTLpath: lemmas on wellformed paths --- backend/RTLpath.v | 148 +++++++++++++++++++++++++++++++++++++++++++----------- 1 file changed, 119 insertions(+), 29 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index 426ec3e3..09e0a9bb 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -344,27 +344,6 @@ Proof. destruct i; simpl; try congruence; simplify_SOME x. Qed. -(* FIXME: useful ? -Lemma step_istate_internal ge i stack f sp rs m : - step_istate ge i stack f sp rs m <> None -> - default_succ i <> None. -Proof. - destruct i; simpl; try congruence; simplify_SOME x. -Qed. -*) - -(* FIXME: useful ? -Lemma step_istate_false_continue ge i stack f sp rs m st : - step_istate ge i stack f sp rs m = Some st -> - st.(continue) = false -> - exists cond, exists args, exists ifnot, i = Icond cond args st.(the_pc) ifnot /\ st.(the_rs) = rs /\ st.(the_mem) = m. -Proof. - destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. - destruct b; simpl in * |- *; try congruence. - repeat eexists; eauto. -Qed. -*) - Local Hint Resolve star_refl. Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, @@ -477,6 +456,8 @@ Qed. End CORRECTNESS. + + (** The hard way: Forward simulation of RTL by RTLpath This way can be viewed as a completeness property: all transitions in RTL can be represented as RTLpath transitions ! @@ -496,26 +477,119 @@ Proof. eapply match_globdef_var. destruct v; eauto. Qed. +(* Theory of wellformed paths *) -Lemma step_istate_complete ge t i stack f sp rs m pc s': - RTL.step ge (State stack f sp pc rs m) t s' -> - (fn_code f)!pc = Some i -> - default_succ i <> None -> - t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Fixpoint nth_default_succ (c: code) (path:nat) (pc: node): option node := + match path with + | O => Some pc + | S path' => + SOME i <- c!pc IN + SOME pc' <- default_succ i IN + nth_default_succ c path' pc' + end. + +Lemma wellformed_suffix_path c pm path path': + (path' <= path)%nat -> + forall pc, wellformed_path c pm path pc -> + exists pc', nth_default_succ c (path-path') pc = Some pc' /\ wellformed_path c pm path' pc'. Proof. - intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; - (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. + induction 1 as [|m]. + + intros; cutrewrite (path'-path'=O)%nat; [simpl;eauto|omega]. + + intros pc WF; cutrewrite (S m-path'=S (m-path'))%nat; [simpl;eauto|omega]. + inversion WF; subst; clear WF; intros; simplify_someHyps. + intros; simplify_someHyps; eauto. +Qed. + +Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := + SOME pc <- nth_default_succ c path pc IN + c!pc. + +Lemma final_node_path f path pc: + (fn_path f)!pc = Some path -> + exists i, nth_default_succ_inst (fn_code f) path pc = Some i + /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). +Proof. + intros; exploit fn_path_wf; eauto. + intro WF. + exploit (wellformed_suffix_path (fn_code f) (fn_path f) path O); omega || eauto. + destruct 1 as (pc' & NTH_SUCC & WF'); auto. + cutrewrite (path - 0 = path)%nat in NTH_SUCC; try omega. + unfold nth_default_succ_inst. + inversion WF'; clear WF'; subst. simplify_someHyps; eauto. +Qed. + +Lemma internal_node_path path f path0 pc: + (fn_path f)!pc = (Some path0) -> + path <> O -> + (path <= path0)%nat -> + exists i pc', + nth_default_succ_inst (fn_code f) (path0-path) pc = Some i /\ + default_succ i = Some pc' /\ + (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). +Proof. + intros; exploit fn_path_wf; eauto. + intro WF; exploit wellformed_suffix_path; eauto. + destruct 1 as (pc' & NTH_SUCC & WF'). + unfold nth_default_succ_inst. + destruct path as [|path']; try congruence. + inversion WF'; clear WF'; subst. simplify_someHyps; eauto. Qed. Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. -Local Hint Resolve fn_entry_point_wf fn_path_wf. +Local Hint Resolve fn_entry_point_wf. +Lemma step_istate_internal ge i stack f sp rs m st: + step_istate ge i stack f sp rs m = Some st -> + st.(continue) = true -> + default_succ i = Some st.(the_pc). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. + destruct x; simpl in * |- *; try congruence. +Qed. + +Lemma path_step_istate_internal ge path stack f sp: forall rs m pc st, + st.(continue) = true -> + path_step_istate ge path stack f sp rs m pc = Some st -> + nth_default_succ (fn_code f) path pc = Some st.(the_pc). +Proof. + induction path; simpl. { try_simplify_someHyps. } + intros rs m pc st CONT; try_simplify_someHyps. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:X; try congruence. + try_simplify_someHyps. + intros; erewrite step_istate_internal; eauto. +Qed. + +(* FIXME: useful ? +Lemma step_istate_early_exit ge i stack f sp rs m st : + step_istate ge i stack f sp rs m = Some st -> + st.(continue) = false -> + exists cond, exists args, exists ifnot, i = Icond cond args st.(the_pc) ifnot /\ st.(the_rs) = rs /\ st.(the_mem) = m. +Proof. + destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. + destruct b; simpl in * |- *; try congruence. + repeat eexists; eauto. +Qed. +*) + +Lemma step_istate Section COMPLETENESS. +Lemma step_istate_complete ge t i stack f sp rs m pc s': + RTL.step ge (State stack f sp pc rs m) t s' -> + (fn_code f)!pc = Some i -> + default_succ i <> None -> + t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; + (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. +Qed. + Variable p: program. Let pge := Genv.globalenv p. @@ -576,6 +650,22 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := else s1 = state_RTL s2) /\ wf_stackframe (stack_of s2). +(* +Lemma path_istate_complete_silent prev_path_idx stack f sp: forall rs m pc st t s1' idx, + path_step_istate ge (S prev_path_idx) stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some ((S prev_path_idx)+idx)%nat -> + st.(continue) = true -> (* ?*) + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + t = E0. +Proof. + intros; exploit (internal_node_path (S prev_path_idx)); omega || eauto. + destruct 1 as (i & pc' & Hi & Hpc & DUM). + cutrewrite (S prev_path_idx + idx - S prev_path_idx = idx)%nat in Hi; try omega. +Qed. +*) + +(*** SUITE EN ATTENTE !!! ***) + Lemma path_step_complete stack f sp rs m pc t s1' idx path st: path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> -- cgit From e3c940ee38e02e6dda9469cf7d84e8b0b9d2a28d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 7 Sep 2019 23:07:49 +0200 Subject: RTLpath: proof of the stuttering case --- backend/RTLpath.v | 132 +++++++++++++++++++++++++++++++++--------------------- 1 file changed, 82 insertions(+), 50 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index 09e0a9bb..dfd93808 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -297,7 +297,7 @@ Definition semantics (p: program) := (** ** Preliminaries: simple tactics for option-monad *) Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B): - (forall x, e = Some x -> P (f x)) -> (P None) -> (P (SOME x <- e IN f x)). + (forall x, e = Some x -> P (f x)) -> (e = None -> P None) -> (P (SOME x <- e IN f x)). Proof. intros; destruct e; simpl; auto. Qed. @@ -520,19 +520,19 @@ Qed. Lemma internal_node_path path f path0 pc: (fn_path f)!pc = (Some path0) -> - path <> O -> - (path <= path0)%nat -> + (path < path0)%nat -> exists i pc', - nth_default_succ_inst (fn_code f) (path0-path) pc = Some i /\ + nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. - intro WF; exploit wellformed_suffix_path; eauto. + intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } destruct 1 as (pc' & NTH_SUCC & WF'). - unfold nth_default_succ_inst. - destruct path as [|path']; try congruence. - inversion WF'; clear WF'; subst. simplify_someHyps; eauto. + cutrewrite (path0 - (path0 - path) = path)%nat in NTH_SUCC; try omega. + unfold nth_default_succ_inst. + inversion WF'; clear WF'; subst. { omega. } + simplify_someHyps; eauto. Qed. Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. @@ -564,6 +564,23 @@ Proof. intros; erewrite step_istate_internal; eauto. Qed. +Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = true -> + (fn_code f)!(st.(the_pc)) = Some i -> + step_istate ge i stack f sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) stack f sp rs m pc. +Proof. + induction path. + + simpl; intros; try_simplify_someHyps. simplify_SOME st. + destruct st as [b]; destruct b; simpl; auto. + + intros rs m pc st i H. + simpl in H. + generalize H; clear H; simplify_SOME xx. + destruct (continue xx0) eqn: CONTxx0. + * intros; erewrite IHpath; eauto. + * intros; congruence. +Qed. + (* FIXME: useful ? Lemma step_istate_early_exit ge i stack f sp rs m st : step_istate ge i stack f sp rs m = Some st -> @@ -576,8 +593,6 @@ Proof. Qed. *) -Lemma step_istate - Section COMPLETENESS. Lemma step_istate_complete ge t i stack f sp rs m pc s': @@ -630,41 +645,51 @@ Definition stack_of (s: state): list stackframe := | Returnstate stack v m => stack end. -Definition is_inst (s: RTL.state): bool := +Definition is_inst (s: RTL.state): option bool := match s with - | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => true | None => false end - | _ => false + | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => Some true | None => Some false end + | _ => None end. -Inductive match_inst_states (idx: nat) (s1:RTL.state): state -> Prop := +Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> (idx <= path)%nat -> path_step_istate ge (path-idx) stack f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> - match_inst_states idx s1 (State stack f sp pc rs m). + match_inst_states_goal idx s1 (State stack f sp pc rs m). + +Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop := + match is_inst s1 with + | Some b => b = true -> match_inst_states_goal idx s1 s2 + | None => s1 = state_RTL s2 + end. Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := - (if is_inst s1 - then match_inst_states idx s1 s2 - else s1 = state_RTL s2) + match_inst_states idx s1 s2 /\ wf_stackframe (stack_of s2). -(* -Lemma path_istate_complete_silent prev_path_idx stack f sp: forall rs m pc st t s1' idx, - path_step_istate ge (S prev_path_idx) stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some ((S prev_path_idx)+idx)%nat -> - st.(continue) = true -> (* ?*) - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - t = E0. +Lemma path_istate_complete_internal path stack f sp rs m pc st t s1' idx: + path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + (S idx <= path)%nat -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). Proof. - intros; exploit (internal_node_path (S prev_path_idx)); omega || eauto. + intros H H0 H1 CONT STEP; exploit (internal_node_path (path-(S idx))); omega || eauto. destruct 1 as (i & pc' & Hi & Hpc & DUM). - cutrewrite (S prev_path_idx + idx - S prev_path_idx = idx)%nat in Hi; try omega. + unfold nth_default_succ_inst in Hi. + erewrite path_step_istate_internal in Hi; eauto. + exploit step_istate_complete; congruence || eauto. + destruct 1 as (SILENT & st0 & STEP0 & EQ). + intuition; subst; unfold match_inst_states; simpl. + destruct ((fn_code f) ! (the_pc st0)) eqn: X. 2: { intros; try congruence. } + intros; refine (State_match _ _ path stack f sp pc rs m _ H0 _ _ _); simpl; omega || eauto. + cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. + erewrite <- path_step_istate_step_right; eauto. Qed. -*) -(*** SUITE EN ATTENTE !!! ***) Lemma path_step_complete stack f sp rs m pc t s1' idx path st: path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> @@ -675,10 +700,19 @@ Lemma path_step_complete stack f sp rs m pc t s1' idx path st: exists idx' s2', (path_step ge pge path stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. +Proof. + intros; destruct (st.(continue)) eqn: CONT. + destruct idx as [ | idx]. + + (* path_step on normal_exit *) admit. + + (* stuttering step *) + exploit path_istate_complete_internal; eauto. + unfold match_states; exists idx; exists (State stack f sp pc rs m); + intuition. + + admit. Admitted. Lemma step_noninst_complete s1 t s1' s2: - is_inst s1 = false -> + is_inst s1 = None -> s1 = state_RTL s2 -> RTL.step ge s1 t s1' -> wf_stackframe (stack_of s2) -> @@ -690,8 +724,8 @@ Proof. destruct f; simpl in H3; inversion H3; subst; clear H3. eexists; constructor 1. * eapply exec_function_internal; eauto. - * unfold match_states; simpl. - destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X; [|exists O; auto]. + * unfold match_states, match_inst_states; simpl. + destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X. 2:{ exists O; intuition. } exploit initialize_path; eauto. destruct 1 as (path & Hpath). exists path. constructor; auto. @@ -703,16 +737,16 @@ Proof. destruct f; simpl in H3 |-; inversion H3; subst; clear H3. eexists; constructor 1. * apply exec_function_external; eauto. - * unfold match_states; simpl. exists O; auto. + * unfold match_states, match_inst_states; simpl. exists O; auto. + (* exec_return *) destruct stack eqn: Hstack; simpl in H1; inversion H1; clear H1; subst. destruct s0 eqn: Hs0; simpl in H0; inversion H0; clear H0; subst. eexists; constructor 1. * apply exec_return. - * unfold match_states; simpl. + * unfold match_states, match_inst_states; simpl. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. - destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; auto]. + destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; intuition ]. exploit initialize_path; eauto. destruct 1 as (path & Hpath). exists path. constructor; auto. @@ -728,19 +762,17 @@ Lemma step_complete s1 t s1' idx s2: RTL.step ge s1 t s1' -> exists idx' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. - unfold match_states at 1. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - - destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; subst. - clear His1. simpl in * |- *. -(* FIXME: His1 USELESS HERE ? - simpl in His1. - destruct ((fn_code f) ! (the_pc s2)) eqn:Hpc; simpl; try congruence. -*) - intros; exploit path_step_complete; eauto. - destruct 1 as (i' & s2' & H0 & H1). - eexists; eexists; eauto. - destruct H0. - + exploit exec_path; eauto. - + intuition subst; eauto. + unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. + - destruct b. + * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. + clear His1. simpl in * |- *. + intros STEP; exploit path_step_complete; eauto. + destruct 1 as (idx' & s2' & H0 & H1). + eexists; eexists; eauto. + destruct H0. + + exploit exec_path; eauto. + + intuition subst; eauto. + * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. - intros; exploit step_noninst_complete; eauto. firstorder. Qed. @@ -750,10 +782,10 @@ Proof. eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). constructor 1; simpl. - apply lt_wf. - - unfold initial_state, match_states. destruct 1; simpl; exists O. + - unfold initial_state, match_states, match_inst_states. destruct 1; simpl; exists O. destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. - - unfold final_state, match_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. + - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. destruct s2; simpl in * |- *; inversion H0; subst. constructor. - Local Hint Resolve plus_one star_refl. -- cgit From 4d2df5a5a08066c5b7074e6b4b0fb3d93fd29997 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 8 Sep 2019 09:48:13 +0200 Subject: RTLpath: starting the normal exit case... --- backend/RTLpath.v | 81 +++++++++++++++++++++++++++++++++++++++++++------------ 1 file changed, 64 insertions(+), 17 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index dfd93808..d0e28422 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -541,7 +541,7 @@ Proof. Qed. Local Hint Resolve fn_entry_point_wf. -Lemma step_istate_internal ge i stack f sp rs m st: +Lemma step_istate_normal_exit ge i stack f sp rs m st: step_istate ge i stack f sp rs m = Some st -> st.(continue) = true -> default_succ i = Some st.(the_pc). @@ -550,7 +550,7 @@ Proof. destruct x; simpl in * |- *; try congruence. Qed. -Lemma path_step_istate_internal ge path stack f sp: forall rs m pc st, +Lemma path_step_istate_normal_exit ge path stack f sp: forall rs m pc st, st.(continue) = true -> path_step_istate ge path stack f sp rs m pc = Some st -> nth_default_succ (fn_code f) path pc = Some st.(the_pc). @@ -561,7 +561,7 @@ Proof. inversion_SOME st0; intros Hst0. destruct (continue st0) eqn:X; try congruence. try_simplify_someHyps. - intros; erewrite step_istate_internal; eauto. + intros; erewrite step_istate_normal_exit; eauto. Qed. Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, @@ -669,7 +669,7 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := match_inst_states idx s1 s2 /\ wf_stackframe (stack_of s2). -Lemma path_istate_complete_internal path stack f sp rs m pc st t s1' idx: +Lemma path_istate_complete_normal_exit path stack f sp rs m pc st t s1' idx: path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (S idx <= path)%nat -> @@ -677,19 +677,58 @@ Lemma path_istate_complete_internal path stack f sp rs m pc st t s1' idx: RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). Proof. - intros H H0 H1 CONT STEP; exploit (internal_node_path (path-(S idx))); omega || eauto. - destruct 1 as (i & pc' & Hi & Hpc & DUM). + intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path-(S idx))); omega || eauto. + intros (i & pc' & Hi & Hpc & DUM). unfold nth_default_succ_inst in Hi. - erewrite path_step_istate_internal in Hi; eauto. + erewrite path_step_istate_normal_exit in Hi; eauto. exploit step_istate_complete; congruence || eauto. - destruct 1 as (SILENT & st0 & STEP0 & EQ). + intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. destruct ((fn_code f) ! (the_pc st0)) eqn: X. 2: { intros; try congruence. } - intros; refine (State_match _ _ path stack f sp pc rs m _ H0 _ _ _); simpl; omega || eauto. + intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. erewrite <- path_step_istate_step_right; eauto. Qed. +Lemma path_last_step_complete_part1 path stack f sp rs m pc st t s1': + path_step_istate ge path stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + exists s2', path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s2'. +Proof. + intros PSTEP PATH CONT RSTEP; exploit (final_node_path f path); eauto. + intros (i & Hi & SUCCS). + unfold nth_default_succ_inst in Hi. + erewrite path_step_istate_normal_exit in Hi; eauto. + destruct (default_succ i) eqn:Hn0. + + (* exec_istate *) + exploit step_istate_complete; congruence || eauto. + intros (SILENT & st0 & STEP0 & EQ); subst. + exploit (exec_istate ge pge); eauto. +Admitted. + +(* Can we do a generic proof here for match_states ? *) +Lemma path_last_step_complete_part2 path stack f sp rs m pc st t s1' s2': + path_step_istate ge path stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + wf_stackframe stack -> + path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s2' -> + (exists idx', match_states idx' s1' s2'). +Admitted. + +Lemma path_last_step_complete path stack f sp rs m pc st t s1': + path_step_istate ge path stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + wf_stackframe stack -> + exists s2', + (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' + /\ (exists idx', match_states idx' s1' s2'). +Admitted. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> @@ -701,14 +740,20 @@ Lemma path_step_complete stack f sp rs m pc t s1' idx path st: (path_step ge pge path stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. - intros; destruct (st.(continue)) eqn: CONT. + intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. - + (* path_step on normal_exit *) admit. + + (* path_step on normal_exit *) + Local Hint Resolve exec_normal_exit. + cutrewrite (path-0=path)%nat in PSTEP; [|omega]. + exploit path_last_step_complete; eauto. + intros (s2' & LSTEP & (idx' & MATCH)). + exists idx'; exists s2'; intuition eauto. + (* stuttering step *) - exploit path_istate_complete_internal; eauto. + exploit path_istate_complete_normal_exit; eauto. unfold match_states; exists idx; exists (State stack f sp pc rs m); intuition. - + admit. + + (* one or two path_step on early_exit *) admit. + (* ISSUE: define two "path_step" *) Admitted. Lemma step_noninst_complete s1 t s1' s2: @@ -760,10 +805,11 @@ Qed. Lemma step_complete s1 t s1' idx s2: match_states idx s1 s2 -> RTL.step ge s1 t s1' -> - exists idx' s2', (step ge pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. + exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. + Local Hint Resolve plus_one. unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - - destruct b. + - destruct b. * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. clear His1. simpl in * |- *. intros STEP; exploit path_step_complete; eauto. @@ -774,7 +820,8 @@ Proof. + intuition subst; eauto. * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. - intros; exploit step_noninst_complete; eauto. - firstorder. + intros (s2' & STEP & (idx0 & MATCH)). + exists idx0; exists s2'; intuition auto. Qed. Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). @@ -788,7 +835,7 @@ Proof. - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. destruct s2; simpl in * |- *; inversion H0; subst. constructor. - - Local Hint Resolve plus_one star_refl. + - Local Hint Resolve star_refl. intros; exploit step_complete; eauto. destruct 1 as (idx' & s2' & X). exists idx'. exists s2'. intuition (subst; eauto). -- cgit From db9fb34d28cf52bfcd3aa36c402ed37f8cf007bf Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 8 Sep 2019 21:49:57 +0200 Subject: RTLpath: avancement du "last_step" --- backend/RTLpath.v | 47 ++++++++++++++++++++++++++++++++++++++++------- 1 file changed, 40 insertions(+), 7 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index d0e28422..f1f6feec 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -405,14 +405,12 @@ Lemma find_function_RTL_match ros rs fd: find_function pge ros rs = Some fd -> RTL.find_function ge ros rs = Some (fundef_RTL fd). Proof. destruct ros; simpl. - + destruct (rs#r); simpl; try congruence. - destruct (Ptrofs.eq_dec i Ptrofs.zero); simpl; try congruence. - intros H; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. - destruct 1 as (cuint & tf & H1 & H2 & H3); subst; auto. + + intro; exploit (Genv.find_funct_match match_prog_RTL); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; auto. + rewrite symbols_preserved. destruct (Genv.find_symbol pge i); simpl; try congruence. - intros H; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. - destruct 1 as (cuint & tf & H1 & H2 & H3); subst; auto. + intro; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; auto. Qed. Local Hint Resolve step_istate_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. @@ -618,6 +616,19 @@ Proof. eauto. Qed. +Lemma find_RTL_function_match ros rs fd: + RTL.find_function ge ros rs = Some fd -> exists fd', fd = fundef_RTL fd' /\ find_function pge ros rs = Some fd'. +Proof. + destruct ros; simpl. + + intro; exploit (Genv.find_funct_match (match_RTL_prog p)); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; eauto. + + rewrite (symbols_preserved p); unfold pge. + destruct (Genv.find_symbol (Genv.globalenv p) i); simpl; try congruence. + intro; exploit find_funct_ptr_RTL_preserv; eauto. + intros (tf & H1 & H2); subst; eauto. +Qed. + + Definition wf_stf (st: stackframe): Prop := match st with | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc @@ -690,6 +701,19 @@ Proof. erewrite <- path_step_istate_step_right; eauto. Qed. + +Ltac simplify_someHyp_RH := + match goal with + | H: None = Some _ |- _ => inversion H; subst + | H: Some _ = None |- _ => inversion H; subst + | H: ?t = ?t |- _ => clear H + | H: Some _ = Some _ |- _ => inversion H; subst + | H: Some _ <> None |- _ => clear H + | H: None <> Some _ |- _ => clear H + | H: _ = Some _ |- _ => (rewrite !H in * |- || generalize H); clear H + end. + + Lemma path_last_step_complete_part1 path stack f sp rs m pc st t s1': path_step_istate ge path stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> @@ -706,7 +730,16 @@ Proof. exploit step_istate_complete; congruence || eauto. intros (SILENT & st0 & STEP0 & EQ); subst. exploit (exec_istate ge pge); eauto. -Admitted. + + Local Hint Resolve exec_Ibuiltin exec_Ijumptable exec_Ireturn. + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp_RH; simpl in * |- *)); try congruence; eauto. + - intros; exploit find_RTL_function_match; eauto. + intros (fd' & MATCHfd & Hfd'); subst. + exploit (exec_Icall ge pge); eauto. + - intros; exploit find_RTL_function_match; eauto. + intros (fd' & MATCHfd & Hfd'); subst. + exploit (exec_Itailcall ge pge); eauto. + - intros; exploit exec_Ibuiltin; eauto. +Qed. (* Can we do a generic proof here for match_states ? *) Lemma path_last_step_complete_part2 path stack f sp rs m pc st t s1' s2': -- cgit From fa51b854d856e6c414a38f25917b775ec506c6c7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Sep 2019 13:01:18 +0200 Subject: RTLpath: proof of the normal exit --- backend/RTLpath.v | 121 ++++++++++++++++++++++++++++++++++++------------------ 1 file changed, 82 insertions(+), 39 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index f1f6feec..ab130c33 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -1,8 +1,19 @@ -(** We introduce a data-structure for partitioning the nodes of a RTL CFG into a set of "traces" (in the sense of "trace-scheduling") +(** We introduce a data-structure extending the RTL CFG, in a control-flow graph over "traces" (in the sense of "trace-scheduling") Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert: - a "path" is simply a list of successive nodes in the CFG. + a "path" is simply a list of successive nodes in the CFG (modulo some additional wellformness conditions). - Here, we extend syntactically the notion of RTL programs with a structure of "path_map". + Actually, we extend syntactically the notion of RTL programs with a structure of "path_map": + this gives an alternative view of the CFG -- where "nodes" are paths instead of simple instructions. + Our wellformness condition on paths express that: + - the CFG on paths is wellformed: any successor of a given path points to another path (possibly the same). + - execution of a paths only emit single events. + + We represent each path only by a natural: the number of nodes in the path. These nodes are recovered from a static notion of "default successor". + This notion of path is thus incomplete. For example, if a path contains a whole loop (and for example, unrools it several times), + then this loop must be a suffix of the path. + + However: it is sufficient in order to represent superblocks (each superblock being represented as a path). + A superblock decomposition of the CFG exactly corresponds to the case where each node is in at most one path. Our goal is to provide two bisimulable semantics: - one is simply the RTL semantics @@ -538,6 +549,15 @@ Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. Local Hint Resolve fn_entry_point_wf. +Local Opaque path_entry. + +Lemma step_istate_successors ge i stack f sp rs m st: + step_istate ge i stack f sp rs m = Some st -> + In (the_pc st) (successors_instr i). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. + destruct x; simpl in * |- *; intuition congruence. +Qed. Lemma step_istate_normal_exit ge i stack f sp rs m st: step_istate ge i stack f sp rs m = Some st -> @@ -713,15 +733,18 @@ Ltac simplify_someHyp_RH := | H: _ = Some _ |- _ => (rewrite !H in * |- || generalize H); clear H end. - -Lemma path_last_step_complete_part1 path stack f sp rs m pc st t s1': +Lemma path_last_step_complete path stack f sp rs m pc st t s1': path_step_istate ge path stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - exists s2', path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s2'. + wf_stackframe stack -> + exists s2', + (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' + /\ (exists idx', match_states idx' s1' s2'). Proof. - intros PSTEP PATH CONT RSTEP; exploit (final_node_path f path); eauto. + Local Hint Resolve step_istate_successors list_nth_z_in. (* Hint for path_entry proofs *) + intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto. intros (i & Hi & SUCCS). unfold nth_default_succ_inst in Hi. erewrite path_step_istate_normal_exit in Hi; eauto. @@ -730,39 +753,57 @@ Proof. exploit step_istate_complete; congruence || eauto. intros (SILENT & st0 & STEP0 & EQ); subst. exploit (exec_istate ge pge); eauto. - + Local Hint Resolve exec_Ibuiltin exec_Ijumptable exec_Ireturn. - generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp_RH; simpl in * |- *)); try congruence; eauto. - - intros; exploit find_RTL_function_match; eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! (the_pc st0)) eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp_RH; simpl in * |- * )); try congruence; eauto. + - (* Icall *) + intros; exploit find_RTL_function_match; eauto. intros (fd' & MATCHfd & Hfd'); subst. exploit (exec_Icall ge pge); eauto. - - intros; exploit find_RTL_function_match; eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. + rewrite wf_stackframe_cons; intuition simpl; eauto. + - (* Itailcall *) + intros; exploit find_RTL_function_match; eauto. intros (fd' & MATCHfd & Hfd'); subst. exploit (exec_Itailcall ge pge); eauto. - - intros; exploit exec_Ibuiltin; eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. + - (* Ibuiltin *) + intros; exploit exec_Ibuiltin; eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + - (* Ijumptable *) + intros; exploit exec_Ijumptable; eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + - (* Ireturn *) + intros; exploit exec_Ireturn; eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. Qed. -(* Can we do a generic proof here for match_states ? *) -Lemma path_last_step_complete_part2 path stack f sp rs m pc st t s1' s2': - path_step_istate ge path stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some path -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - wf_stackframe stack -> - path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s2' -> - (exists idx', match_states idx' s1' s2'). -Admitted. - -Lemma path_last_step_complete path stack f sp rs m pc st t s1': - path_step_istate ge path stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some path -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - wf_stackframe stack -> - exists s2', - (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' - /\ (exists idx', match_states idx' s1' s2'). -Admitted. - Lemma path_step_complete stack f sp rs m pc t s1' idx path st: path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> @@ -771,7 +812,11 @@ Lemma path_step_complete stack f sp rs m pc t s1' idx path st: wf_stackframe stack -> exists idx' s2', (path_step ge pge path stack f sp rs m pc t s2' - \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. + \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) + \/ (path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + /\ (fn_path f)!(the_pc st) = Some O /\ path_step ge pge 0 stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') + ) + /\ match_states idx' s1' s2'. Proof. intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. @@ -840,7 +885,7 @@ Lemma step_complete s1 t s1' idx s2: RTL.step ge s1 t s1' -> exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. - Local Hint Resolve plus_one. + Local Hint Resolve plus_one plus_two exec_path. unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - destruct b. * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. @@ -848,9 +893,7 @@ Proof. intros STEP; exploit path_step_complete; eauto. destruct 1 as (idx' & s2' & H0 & H1). eexists; eexists; eauto. - destruct H0. - + exploit exec_path; eauto. - + intuition subst; eauto. + destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. - intros; exploit step_noninst_complete; eauto. intros (s2' & STEP & (idx0 & MATCH)). -- cgit From 76aaaa0d3d26de88e0524ca160c5036a28455cfa Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Sep 2019 19:00:43 +0200 Subject: RTLpath: preliminaryies to early exit case... --- backend/RTLpath.v | 87 ++++++++++++++++++++++++++++++++++++++++--------------- 1 file changed, 64 insertions(+), 23 deletions(-) diff --git a/backend/RTLpath.v b/backend/RTLpath.v index ab130c33..82c787d4 100644 --- a/backend/RTLpath.v +++ b/backend/RTLpath.v @@ -111,7 +111,6 @@ Record function : Type := fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path }. - Definition fundef := AST.fundef function. Definition program := AST.program fundef unit. Definition genv := Genv.t fundef unit. @@ -318,13 +317,13 @@ Local Ltac inversion_SOME x := Ltac simplify_someHyp := match goal with - | H: None = Some _ |- _ => inversion H; subst - | H: Some _ = None |- _ => inversion H; subst + | H: None = Some _ |- _ => inversion H; clear H; subst + | H: Some _ = None |- _ => inversion H; clear H; subst | H: ?t = ?t |- _ => clear H - | H: Some _ = Some _ |- _ => inversion H; subst + | H: Some _ = Some _ |- _ => inversion H; clear H; subst | H: Some _ <> None |- _ => clear H | H: None <> Some _ |- _ => clear H - | H: _ = Some _ |- _ => (rewrite !H in * |- * || generalize H); clear H + | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H end. Ltac simplify_someHyps := @@ -599,17 +598,71 @@ Proof. * intros; congruence. Qed. -(* FIXME: useful ? +(* USEFUL ? +Lemma path_step_istate_inv_normal ge path stack f sp: forall rs m pc st, + path_step_istate ge (S path) stack f sp rs m pc = Some st -> + (continue st)=true -> + exists st0 i, + path_step_istate ge path stack f sp rs m pc = Some st0 /\ + st0.(continue) = true /\ + (fn_code f)!(st0.(the_pc)) = Some i /\ + step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. +Proof. + simpl; intros rs m pc st. inversion_SOME i; inversion_SOME st0. + destruct (continue st0) eqn: CONT; try (intro; simplify_someHyps; congruence). + generalize rs m pc st i st0 CONT; + clear rs m pc st st i st0 CONT. + induction path; simpl; intros rs m pc st i st0. + - try_simplify_someHyps; simpl. + repeat eexists; intuition eauto. + destruct st0; simpl in * |- *; subst; auto. + - simplify_SOME yy. + rewrite CONT. + destruct (continue yy0) eqn: CONT0; try_simplify_someHyps; congruence || auto. +Qed. +*) + +Lemma path_step_istate_inversion_early ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + (continue st)=false -> + exists st0 i path0, + (path > path0)%nat /\ + path_step_istate ge path0 stack f sp rs m pc = Some st0 /\ + st0.(continue) = true /\ + (fn_code f)!(st0.(the_pc)) = Some i /\ + step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. +Proof. + induction path as [|path]; simpl. + - intros; try_simplify_someHyps; try congruence. + - intros rs m pc st; inversion_SOME i; inversion_SOME st0. + cutrewrite (path-0=path)%nat; try omega. + destruct (continue st0) eqn: CONT. + + intros STEP PC STEPS CONT0. exploit IHpath; eauto. + clear STEPS. + intros (st1 & i0 & path0 & BOUND & STEP1 & CONT1 & X1 & X2); auto. + exists st1. exists i0. exists (S path0). intuition. + simpl; try_simplify_someHyps. + rewrite CONT. auto. + + intros; try_simplify_someHyps; try congruence. + eexists. exists i. exists O; simpl. intuition eauto. + omega. +Qed. + + +Inductive is_early_exit pc: instruction -> Prop := + | Icond_early_exit cond args ifnot: + is_early_exit pc (Icond cond args pc ifnot). + Lemma step_istate_early_exit ge i stack f sp rs m st : step_istate ge i stack f sp rs m = Some st -> st.(continue) = false -> - exists cond, exists args, exists ifnot, i = Icond cond args st.(the_pc) ifnot /\ st.(the_rs) = rs /\ st.(the_mem) = m. + st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. Proof. + Local Hint Resolve Icond_early_exit. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. - destruct b; simpl in * |- *; try congruence. - repeat eexists; eauto. + destruct b; simpl in * |- *; intuition eauto. + congruence. Qed. -*) Section COMPLETENESS. @@ -721,18 +774,6 @@ Proof. erewrite <- path_step_istate_step_right; eauto. Qed. - -Ltac simplify_someHyp_RH := - match goal with - | H: None = Some _ |- _ => inversion H; subst - | H: Some _ = None |- _ => inversion H; subst - | H: ?t = ?t |- _ => clear H - | H: Some _ = Some _ |- _ => inversion H; subst - | H: Some _ <> None |- _ => clear H - | H: None <> Some _ |- _ => clear H - | H: _ = Some _ |- _ => (rewrite !H in * |- || generalize H); clear H - end. - Lemma path_last_step_complete path stack f sp rs m pc st t s1': path_step_istate ge path stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> @@ -762,7 +803,7 @@ Proof. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. * simpl; eauto. - + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp_RH; simpl in * |- * )); try congruence; eauto. + + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. - (* Icall *) intros; exploit find_RTL_function_match; eauto. intros (fd' & MATCHfd & Hfd'); subst. -- cgit From 2b6efc56c8a3e6b46689e53d8baba444155c79ef Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Sep 2019 22:52:48 +0200 Subject: RTLpath: end of the proof --- backend/RTLpath.v | 962 --------------------------------------------- mppa_k1c/lib/RTLpath.v | 1009 ++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 1009 insertions(+), 962 deletions(-) delete mode 100644 backend/RTLpath.v create mode 100644 mppa_k1c/lib/RTLpath.v diff --git a/backend/RTLpath.v b/backend/RTLpath.v deleted file mode 100644 index 82c787d4..00000000 --- a/backend/RTLpath.v +++ /dev/null @@ -1,962 +0,0 @@ -(** We introduce a data-structure extending the RTL CFG, in a control-flow graph over "traces" (in the sense of "trace-scheduling") - Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert: - a "path" is simply a list of successive nodes in the CFG (modulo some additional wellformness conditions). - - Actually, we extend syntactically the notion of RTL programs with a structure of "path_map": - this gives an alternative view of the CFG -- where "nodes" are paths instead of simple instructions. - Our wellformness condition on paths express that: - - the CFG on paths is wellformed: any successor of a given path points to another path (possibly the same). - - execution of a paths only emit single events. - - We represent each path only by a natural: the number of nodes in the path. These nodes are recovered from a static notion of "default successor". - This notion of path is thus incomplete. For example, if a path contains a whole loop (and for example, unrools it several times), - then this loop must be a suffix of the path. - - However: it is sufficient in order to represent superblocks (each superblock being represented as a path). - A superblock decomposition of the CFG exactly corresponds to the case where each node is in at most one path. - - Our goal is to provide two bisimulable semantics: - - one is simply the RTL semantics - - the other is based on a notion of "path-step": each path is executed in a single step. - - Remark that all analyses on RTL programs should thus be appliable for "free" also for RTLpath programs ! -*) - -Require Import Coqlib Maps. -Require Import AST Integers Values Events Memory Globalenvs Smallstep. -Require Import Op Registers. -Require Import RTL Linking. - -Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end) - (at level 200, X ident, A at level 100, B at level 200) - : option_monad_scope. - -Local Open Scope option_monad_scope. - -(** * Syntax of RTLpath programs *) - -(** Internal instruction = instruction with a default successor in a path. *) - -Definition default_succ (i: instruction): option node := - match i with - | Inop s => Some s - | Iop op args res s => Some s - | Iload chunk addr args dst s => Some s - | Istore chunk addr args src s => Some s - | Icond cond args ifso ifnot => Some ifnot - | _ => None (* TODO: we could choose a successor for jumptable ? *) - end. - -Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *) - match i with - | Icond cond args ifso ifnot => Some ifso - | _ => None - end. - -(** Our notion of path. - - We do not formally require that the set of path is a partition of the CFG. - path may have intersections ! - - Moreover, we do not formally require that path have a single entry-point (a superblock structure) - - But, in practice, these properties are probably necessary in order to ensure the success of dynamic verification of scheduling. - - Here: we only require that each exit-point of a path is the entry-point of a path - (and that internal node of a path are internal instructions) -*) - - -(* By convention, we say that node [n] of the CFG is an entry-point of a path - if it is an entry of the path_map. - - Such a path of entry [n] is defined from a natural [path] representing the [path] default-successors of [n]. - - Remark: a path can loop several times in the CFG. - -*) - -Definition path_map: Type := PTree.t nat. - -Definition path_entry (c:code) (pm: path_map) (n: node): Prop - := pm!n = None -> c!n=None. - -Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := - | wf_last_node i pc: - c!pc = Some i -> - (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> - wellformed_path c pm 0 pc - | wf_internal_node path i pc pc': - c!pc = Some i -> - default_succ i = Some pc' -> - (forall n, early_exit i = Some n -> path_entry c pm n) -> - wellformed_path c pm path pc' -> - wellformed_path c pm (S path) pc. - -(* all paths defined from the path_map are wellformed *) -Definition wellformed_path_map (c:code) (pm: path_map): Prop := - forall n path, pm!n = Some path -> wellformed_path c pm path n. - -(** We "extend" the notion of RTL program with the additional structure for path. - - There is thus a trivial "forgetful functor" from RTLpath programs to RTL ones. -*) - -Record function : Type := - { fn_RTL:> RTL.function; - fn_path: path_map; - (* condition 1 below: the entry-point of the code is an entry-point of a path *) - fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); - (* condition 2 below: the path_map is well-formed *) - fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path - }. - -Definition fundef := AST.fundef function. -Definition program := AST.program fundef unit. -Definition genv := Genv.t fundef unit. - -Definition fundef_RTL (fu: fundef) : RTL.fundef := - match fu with - | Internal f => Internal f.(fn_RTL) - | External ef => External ef - end. -Coercion fundef_RTL: fundef >-> RTL.fundef. - -Definition program_RTL (p: program) : RTL.program := transform_program fundef_RTL p. -Coercion program_RTL: program >-> RTL.program. - -(** * Path-step semantics of RTLpath programs *) - -(* Semantics of internal instructions (mimicking RTL semantics) *) - -Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. - -Inductive stackframe : Type := - | Stackframe - (res: reg) (**r where to store the result *) - (f: function) (**r calling function *) - (sp: val) (**r stack pointer in calling function *) - (pc: node) (**r program point in calling function *) - (rs: regset) (**r register state in calling function *) - . - -Definition stf_RTL (st: stackframe): RTL.stackframe := - match st with - | Stackframe res f sp pc rs => RTL.Stackframe res f sp pc rs - end. - -Fixpoint stack_RTL (stack: list stackframe): list RTL.stackframe := - match stack with - | nil => nil - | cons stf stack' => cons (stf_RTL stf) (stack_RTL stack') - end. - -Inductive state : Type := - | State - (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 *) - | Callstate - (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 *) - | Returnstate - (stack: list stackframe) (**r call stack *) - (v: val) (**r return value for the call *) - (m: mem) (**r memory state *) - . - -Definition state_RTL (s: state): RTL.state := - match s with - | State stack f sp pc rs m => RTL.State (stack_RTL stack) f sp pc rs m - | Callstate stack f args m => RTL.Callstate (stack_RTL stack) f args m - | Returnstate stack v m => RTL.Returnstate (stack_RTL stack) v m - end. -Coercion state_RTL: state >-> RTL.state. - -Definition step_istate (ge: RTL.genv) (i: instruction) (stack: list stackframe) (f: function) (sp: val) (rs: regset) (m: mem): option istate := - match i with - | Inop pc' => Some (ist true pc' rs m) - | Iop op args res pc' => - SOME v <- eval_operation ge sp op rs##args m IN - Some (ist true pc' (rs#res <- v) m) - | Iload chunk addr args dst pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME v <- Mem.loadv chunk m a IN - Some (ist true pc' (rs#dst <- v) m) - | Istore chunk addr args src pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME m' <- Mem.storev chunk m a rs#src IN - Some (ist true pc' rs m') - | Icond cond args ifso ifnot => - SOME b <- eval_condition cond rs##args m IN - Some (ist (negb b) (if b then ifso else ifnot) rs m) - | _ => None (* TODO jumptable ? *) - end. - -(** Execution of a path in a single step *) - -Fixpoint path_step_istate ge (path:nat) stack (f: function) sp rs m pc: option istate := - match path with - | O => Some (ist true pc rs m) - | S p => - SOME i <- (fn_code f)!pc IN - SOME st <- step_istate ge i stack f sp rs m IN - if st.(continue) then - path_step_istate ge p stack f sp st.(the_rs) st.(the_mem) st.(the_pc) - else - Some st - end. - -Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef := - match ros with - | inl r => Genv.find_funct pge rs#r - | inr symb => - match Genv.find_symbol pge symb with - | None => None - | Some b => Genv.find_funct_ptr pge b - end - end. - -Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := - | exec_istate i sp pc rs m st: - (fn_code f)!pc = Some i -> - step_istate ge i stack f sp rs m = Some st -> - path_last_step ge pge stack f sp pc rs m - E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - | exec_Icall sp pc rs m sig ros args res pc' fd: - (fn_code f)!pc = Some(Icall sig ros args res pc') -> - find_function pge ros rs = Some fd -> - funsig fd = sig -> - path_last_step ge pge stack f sp pc rs m - E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) - | exec_Itailcall stk pc rs m sig ros args fd m': - (fn_code f)!pc = Some(Itailcall sig ros args) -> - find_function pge ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m - E0 (Callstate stack fd rs##args m') - | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': - (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - path_last_step ge pge stack f sp pc rs m - t (State stack f sp pc' (regmap_setres res vres rs) m') - | exec_Ijumptable sp pc rs m arg tbl n pc': - (fn_code f)!pc = Some(Ijumptable arg tbl) -> - rs#arg = Vint n -> - list_nth_z tbl (Int.unsigned n) = Some pc' -> - path_last_step ge pge stack f sp pc rs m - E0 (State stack f sp pc' rs m) - | exec_Ireturn stk pc rs m or m': - (fn_code f)!pc = Some(Ireturn or) -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m - E0 (Returnstate stack (regmap_optget or Vundef rs) m'). - -Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := - | exec_early_exit st: - path_step_istate ge path stack f sp rs m pc = Some st -> - st.(continue) = false -> - path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - | exec_normal_exit st t s: - path_step_istate ge path stack f sp rs m pc = Some st -> - st.(continue) = true -> - path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> - path_step ge pge path stack f sp rs m pc t s. - -Inductive step ge pge: state -> trace -> state -> Prop := - | exec_path path stack f sp rs m pc t s: - (fn_path f)!pc = Some path -> - path_step ge pge path stack f sp rs m pc t s -> - step ge pge (State stack f sp pc rs m) t s - | exec_function_internal s f args m m' stk: - Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) -> - step ge pge (Callstate s (Internal f) args m) - E0 (State s - f - (Vptr stk Ptrofs.zero) - f.(fn_entrypoint) - (init_regs args f.(fn_params)) - m') - | exec_function_external s ef args res t m m': - external_call ef ge args m t res m' -> - step ge pge (Callstate s (External ef) args m) - t (Returnstate s res m') - | exec_return res f sp pc rs s vres m: - step ge pge (Returnstate (Stackframe res f sp pc rs :: s) vres m) - E0 (State s f sp pc (rs#res <- vres) m). - -Definition initial_state (p: program) (st:state): Prop - := RTL.initial_state p st. - -Definition final_state (st: state) (i:int): Prop - := RTL.final_state st i. - -Definition semantics (p: program) := - Semantics (step (Genv.globalenv (program_RTL p))) (initial_state p) final_state (Genv.globalenv p). - - -(** * Proving the bisimulation between (semantics p) and (RTL.semantics p). *) - -(** ** Preliminaries: simple tactics for option-monad *) - -Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B): - (forall x, e = Some x -> P (f x)) -> (e = None -> P None) -> (P (SOME x <- e IN f x)). -Proof. - intros; destruct e; simpl; auto. -Qed. - -Local Ltac inversion_SOME x := - try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]). - -Ltac simplify_someHyp := - match goal with - | H: None = Some _ |- _ => inversion H; clear H; subst - | H: Some _ = None |- _ => inversion H; clear H; subst - | H: ?t = ?t |- _ => clear H - | H: Some _ = Some _ |- _ => inversion H; clear H; subst - | H: Some _ <> None |- _ => clear H - | H: None <> Some _ |- _ => clear H - | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H - end. - -Ltac simplify_someHyps := - repeat (simplify_someHyp; simpl in * |- *). - -Ltac try_simplify_someHyps := - try (intros; simplify_someHyps; eauto). - -Ltac simplify_SOME x := - (repeat inversion_SOME x); try_simplify_someHyps. - - -Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. - -(** ** The easy way: Forward simulation of RTLpath by RTL - -This way can be viewed as a correctness property: all transitions in RTLpath are valid RTL transitions ! - -*) - -Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. - -Lemma step_istate_correct ge i stack f sp rs m st : - step_istate ge i stack f sp rs m = Some st -> - forall pc, (fn_code f)!pc = Some i -> - RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). -Proof. - destruct i; simpl; try congruence; simplify_SOME x. -Qed. - -Local Hint Resolve star_refl. - -Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> - star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). -Proof. - induction path; simpl; try_simplify_someHyps. - inversion_SOME i; intros Hi. - inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:cont. - + intros; eapply star_step. - - eapply step_istate_correct; eauto. - - simpl; eauto. - - auto. - + intros; simplify_someHyp; eapply star_step. - - eapply step_istate_correct; eauto. - - simpl; eauto. - - auto. -Qed. - -Lemma path_step_istate_correct_early_exit ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> - st.(continue) = false -> - plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). -Proof. - destruct path; simpl; try_simplify_someHyps; try congruence. - inversion_SOME i; intros Hi. - inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:cont. - + intros; eapply plus_left. - - eapply step_istate_correct; eauto. - - eapply path_step_istate_correct; eauto. - - auto. - + intros X; inversion X; subst. - eapply plus_one. - eapply step_istate_correct; eauto. -Qed. - - -Section CORRECTNESS. - -Variable p: program. - -Lemma match_prog_RTL: match_program (fun _ f tf => tf = fundef_RTL f) eq p (program_RTL p). -Proof. - eapply match_transform_program; eauto. -Qed. - -Let pge := Genv.globalenv p. -Let ge := Genv.globalenv (program_RTL p). - -Lemma senv_preserved: Senv.equiv pge ge. -Proof (Genv.senv_match match_prog_RTL). - -Lemma symbols_preserved s: Genv.find_symbol ge s = Genv.find_symbol pge s. -Proof (Genv.find_symbol_match match_prog_RTL s). - -Lemma find_function_RTL_match ros rs fd: - find_function pge ros rs = Some fd -> RTL.find_function ge ros rs = Some (fundef_RTL fd). -Proof. - destruct ros; simpl. - + intro; exploit (Genv.find_funct_match match_prog_RTL); eauto. - intros (cuint & tf & H1 & H2 & H3); subst; auto. - + rewrite symbols_preserved. - destruct (Genv.find_symbol pge i); simpl; try congruence. - intro; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. - intros (cuint & tf & H1 & H2 & H3); subst; auto. -Qed. - -Local Hint Resolve step_istate_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. - -Lemma path_last_step_correct stack f sp pc rs m t s: - path_last_step ge pge stack f sp pc rs m t s -> - RTL.step ge (State stack f sp pc rs m) t s. -Proof. - destruct 1; try (eapply step_istate_correct); simpl; eauto. -Qed. - -Lemma path_step_correct path stack f sp pc rs m t s: - path_step ge pge path stack f sp rs m pc t s -> - plus RTL.step ge (State stack f sp pc rs m) t s. -Proof. - destruct 1. - + eapply path_step_istate_correct_early_exit; eauto. - + eapply plus_right. - eapply path_step_istate_correct; eauto. - eapply path_last_step_correct; eauto. - auto. -Qed. - -Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return. - -Lemma step_correct s t s': step ge pge s t s' -> plus RTL.step ge s t s'. -Proof. - destruct 1; try (eapply path_step_correct); simpl; eauto. -Qed. - -Theorem RTLpath_correct: forward_simulation (semantics p) (RTL.semantics p). -Proof. - eapply forward_simulation_plus with (match_states := fun s1 s2 => s2 = state_RTL s1); simpl; auto. - - apply senv_preserved. - - unfold initial_state; eauto. - - unfold final_state; intros; subst; eauto. - - intros; subst. eexists; intuition. - eapply step_correct; eauto. -Qed. - -End CORRECTNESS. - - - - -(** The hard way: Forward simulation of RTL by RTLpath - -This way can be viewed as a completeness property: all transitions in RTL can be represented as RTLpath transitions ! - -*) - -(* This lemma is probably needed to compose a pass from RTL -> RTLpath with other passes.*) -Lemma match_RTL_prog {LA: Linker fundef} {LV: Linker unit} p: match_program (fun _ f tf => f = fundef_RTL tf) eq (program_RTL p) p. -Proof. - unfold match_program, match_program_gen; intuition. - unfold program_RTL at 2; simpl. - generalize (prog_defs p). - induction l as [|a l]; simpl; eauto. - destruct a; simpl. - intros; eapply list_forall2_cons; eauto. - unfold match_ident_globdef; simpl; intuition; destruct g as [f|v]; simpl; eauto. - eapply match_globdef_var. destruct v; eauto. -Qed. - -(* Theory of wellformed paths *) - -Fixpoint nth_default_succ (c: code) (path:nat) (pc: node): option node := - match path with - | O => Some pc - | S path' => - SOME i <- c!pc IN - SOME pc' <- default_succ i IN - nth_default_succ c path' pc' - end. - -Lemma wellformed_suffix_path c pm path path': - (path' <= path)%nat -> - forall pc, wellformed_path c pm path pc -> - exists pc', nth_default_succ c (path-path') pc = Some pc' /\ wellformed_path c pm path' pc'. -Proof. - induction 1 as [|m]. - + intros; cutrewrite (path'-path'=O)%nat; [simpl;eauto|omega]. - + intros pc WF; cutrewrite (S m-path'=S (m-path'))%nat; [simpl;eauto|omega]. - inversion WF; subst; clear WF; intros; simplify_someHyps. - intros; simplify_someHyps; eauto. -Qed. - -Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := - SOME pc <- nth_default_succ c path pc IN - c!pc. - -Lemma final_node_path f path pc: - (fn_path f)!pc = Some path -> - exists i, nth_default_succ_inst (fn_code f) path pc = Some i - /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). -Proof. - intros; exploit fn_path_wf; eauto. - intro WF. - exploit (wellformed_suffix_path (fn_code f) (fn_path f) path O); omega || eauto. - destruct 1 as (pc' & NTH_SUCC & WF'); auto. - cutrewrite (path - 0 = path)%nat in NTH_SUCC; try omega. - unfold nth_default_succ_inst. - inversion WF'; clear WF'; subst. simplify_someHyps; eauto. -Qed. - -Lemma internal_node_path path f path0 pc: - (fn_path f)!pc = (Some path0) -> - (path < path0)%nat -> - exists i pc', - nth_default_succ_inst (fn_code f) path pc = Some i /\ - default_succ i = Some pc' /\ - (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). -Proof. - intros; exploit fn_path_wf; eauto. - intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } - destruct 1 as (pc' & NTH_SUCC & WF'). - cutrewrite (path0 - (path0 - path) = path)%nat in NTH_SUCC; try omega. - unfold nth_default_succ_inst. - inversion WF'; clear WF'; subst. { omega. } - simplify_someHyps; eauto. -Qed. - -Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. -Proof. - unfold path_entry; destruct pm!n; eauto. intuition congruence. -Qed. -Local Hint Resolve fn_entry_point_wf. -Local Opaque path_entry. - -Lemma step_istate_successors ge i stack f sp rs m st: - step_istate ge i stack f sp rs m = Some st -> - In (the_pc st) (successors_instr i). -Proof. - destruct i; simpl; try congruence; simplify_SOME x. - destruct x; simpl in * |- *; intuition congruence. -Qed. - -Lemma step_istate_normal_exit ge i stack f sp rs m st: - step_istate ge i stack f sp rs m = Some st -> - st.(continue) = true -> - default_succ i = Some st.(the_pc). -Proof. - destruct i; simpl; try congruence; simplify_SOME x. - destruct x; simpl in * |- *; try congruence. -Qed. - -Lemma path_step_istate_normal_exit ge path stack f sp: forall rs m pc st, - st.(continue) = true -> - path_step_istate ge path stack f sp rs m pc = Some st -> - nth_default_succ (fn_code f) path pc = Some st.(the_pc). -Proof. - induction path; simpl. { try_simplify_someHyps. } - intros rs m pc st CONT; try_simplify_someHyps. - inversion_SOME i; intros Hi. - inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:X; try congruence. - try_simplify_someHyps. - intros; erewrite step_istate_normal_exit; eauto. -Qed. - -Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, - path_step_istate ge path stack f sp rs m pc = Some st -> - st.(continue) = true -> - (fn_code f)!(st.(the_pc)) = Some i -> - step_istate ge i stack f sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) stack f sp rs m pc. -Proof. - induction path. - + simpl; intros; try_simplify_someHyps. simplify_SOME st. - destruct st as [b]; destruct b; simpl; auto. - + intros rs m pc st i H. - simpl in H. - generalize H; clear H; simplify_SOME xx. - destruct (continue xx0) eqn: CONTxx0. - * intros; erewrite IHpath; eauto. - * intros; congruence. -Qed. - -(* USEFUL ? -Lemma path_step_istate_inv_normal ge path stack f sp: forall rs m pc st, - path_step_istate ge (S path) stack f sp rs m pc = Some st -> - (continue st)=true -> - exists st0 i, - path_step_istate ge path stack f sp rs m pc = Some st0 /\ - st0.(continue) = true /\ - (fn_code f)!(st0.(the_pc)) = Some i /\ - step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. -Proof. - simpl; intros rs m pc st. inversion_SOME i; inversion_SOME st0. - destruct (continue st0) eqn: CONT; try (intro; simplify_someHyps; congruence). - generalize rs m pc st i st0 CONT; - clear rs m pc st st i st0 CONT. - induction path; simpl; intros rs m pc st i st0. - - try_simplify_someHyps; simpl. - repeat eexists; intuition eauto. - destruct st0; simpl in * |- *; subst; auto. - - simplify_SOME yy. - rewrite CONT. - destruct (continue yy0) eqn: CONT0; try_simplify_someHyps; congruence || auto. -Qed. -*) - -Lemma path_step_istate_inversion_early ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> - (continue st)=false -> - exists st0 i path0, - (path > path0)%nat /\ - path_step_istate ge path0 stack f sp rs m pc = Some st0 /\ - st0.(continue) = true /\ - (fn_code f)!(st0.(the_pc)) = Some i /\ - step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. -Proof. - induction path as [|path]; simpl. - - intros; try_simplify_someHyps; try congruence. - - intros rs m pc st; inversion_SOME i; inversion_SOME st0. - cutrewrite (path-0=path)%nat; try omega. - destruct (continue st0) eqn: CONT. - + intros STEP PC STEPS CONT0. exploit IHpath; eauto. - clear STEPS. - intros (st1 & i0 & path0 & BOUND & STEP1 & CONT1 & X1 & X2); auto. - exists st1. exists i0. exists (S path0). intuition. - simpl; try_simplify_someHyps. - rewrite CONT. auto. - + intros; try_simplify_someHyps; try congruence. - eexists. exists i. exists O; simpl. intuition eauto. - omega. -Qed. - - -Inductive is_early_exit pc: instruction -> Prop := - | Icond_early_exit cond args ifnot: - is_early_exit pc (Icond cond args pc ifnot). - -Lemma step_istate_early_exit ge i stack f sp rs m st : - step_istate ge i stack f sp rs m = Some st -> - st.(continue) = false -> - st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. -Proof. - Local Hint Resolve Icond_early_exit. - destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. - destruct b; simpl in * |- *; intuition eauto. - congruence. -Qed. - -Section COMPLETENESS. - -Lemma step_istate_complete ge t i stack f sp rs m pc s': - RTL.step ge (State stack f sp pc rs m) t s' -> - (fn_code f)!pc = Some i -> - default_succ i <> None -> - t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). -Proof. - intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; - (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. -Qed. - -Variable p: program. - -Let pge := Genv.globalenv p. -Let ge := Genv.globalenv (program_RTL p). - -Lemma find_funct_ptr_RTL_preserv b f: - Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0). -Proof. - intros; exploit (Genv.find_funct_ptr_match (match_RTL_prog p)); eauto. - destruct 1 as (cunit & tf & X & Y & Z); subst. - eauto. -Qed. - -Lemma find_RTL_function_match ros rs fd: - RTL.find_function ge ros rs = Some fd -> exists fd', fd = fundef_RTL fd' /\ find_function pge ros rs = Some fd'. -Proof. - destruct ros; simpl. - + intro; exploit (Genv.find_funct_match (match_RTL_prog p)); eauto. - intros (cuint & tf & H1 & H2 & H3); subst; eauto. - + rewrite (symbols_preserved p); unfold pge. - destruct (Genv.find_symbol (Genv.globalenv p) i); simpl; try congruence. - intro; exploit find_funct_ptr_RTL_preserv; eauto. - intros (tf & H1 & H2); subst; eauto. -Qed. - - -Definition wf_stf (st: stackframe): Prop := - match st with - | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc - end. - -Definition wf_stackframe (stack: list stackframe): Prop := - forall st, List.In st stack -> wf_stf st. - -Lemma wf_stackframe_nil: wf_stackframe nil. -Proof. - unfold wf_stackframe; simpl. tauto. -Qed. -Local Hint Resolve wf_stackframe_nil. - -Lemma wf_stackframe_cons st stack: - wf_stackframe (st::stack) <-> (wf_stf st) /\ wf_stackframe stack. -Proof. - unfold wf_stackframe; simpl; intuition (subst; auto). -Qed. - -Definition stack_of (s: state): list stackframe := - match s with - | State stack f sp pc rs m => stack - | Callstate stack f args m => stack - | Returnstate stack v m => stack - end. - -Definition is_inst (s: RTL.state): option bool := - match s with - | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => Some true | None => Some false end - | _ => None - end. - -Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := - | State_match path stack f sp pc rs m s2: - (fn_path f)!pc = Some path -> - (idx <= path)%nat -> - path_step_istate ge (path-idx) stack f sp rs m pc = Some s2 -> - s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> - match_inst_states_goal idx s1 (State stack f sp pc rs m). - -Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop := - match is_inst s1 with - | Some b => b = true -> match_inst_states_goal idx s1 s2 - | None => s1 = state_RTL s2 - end. - -Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := - match_inst_states idx s1 s2 - /\ wf_stackframe (stack_of s2). - -Lemma path_istate_complete_normal_exit path stack f sp rs m pc st t s1' idx: - path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some path -> - (S idx <= path)%nat -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). -Proof. - intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path-(S idx))); omega || eauto. - intros (i & pc' & Hi & Hpc & DUM). - unfold nth_default_succ_inst in Hi. - erewrite path_step_istate_normal_exit in Hi; eauto. - exploit step_istate_complete; congruence || eauto. - intros (SILENT & st0 & STEP0 & EQ). - intuition; subst; unfold match_inst_states; simpl. - destruct ((fn_code f) ! (the_pc st0)) eqn: X. 2: { intros; try congruence. } - intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. - cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. - erewrite <- path_step_istate_step_right; eauto. -Qed. - -Lemma path_last_step_complete path stack f sp rs m pc st t s1': - path_step_istate ge path stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some path -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - wf_stackframe stack -> - exists s2', - (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' - /\ (exists idx', match_states idx' s1' s2'). -Proof. - Local Hint Resolve step_istate_successors list_nth_z_in. (* Hint for path_entry proofs *) - intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto. - intros (i & Hi & SUCCS). - unfold nth_default_succ_inst in Hi. - erewrite path_step_istate_normal_exit in Hi; eauto. - destruct (default_succ i) eqn:Hn0. - + (* exec_istate *) - exploit step_istate_complete; congruence || eauto. - intros (SILENT & st0 & STEP0 & EQ); subst. - exploit (exec_istate ge pge); eauto. - eexists; intuition eauto. - unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! (the_pc st0)) eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). - exists path0; intuition eauto. - econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. - * simpl; eauto. - + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. - - (* Icall *) - intros; exploit find_RTL_function_match; eauto. - intros (fd' & MATCHfd & Hfd'); subst. - exploit (exec_Icall ge pge); eauto. - eexists; intuition eauto. - eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. - rewrite wf_stackframe_cons; intuition simpl; eauto. - - (* Itailcall *) - intros; exploit find_RTL_function_match; eauto. - intros (fd' & MATCHfd & Hfd'); subst. - exploit (exec_Itailcall ge pge); eauto. - eexists; intuition eauto. - eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. - - (* Ibuiltin *) - intros; exploit exec_Ibuiltin; eauto. - eexists; intuition eauto. - unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). - exists path0; intuition eauto. - econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. - * simpl; eauto. - - (* Ijumptable *) - intros; exploit exec_Ijumptable; eauto. - eexists; intuition eauto. - unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). - exists path0; intuition eauto. - econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. - * simpl; eauto. - - (* Ireturn *) - intros; exploit exec_Ireturn; eauto. - eexists; intuition eauto. - eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. -Qed. - -Lemma path_step_complete stack f sp rs m pc t s1' idx path st: - path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> - (fn_path f)!pc = Some path -> - (idx <= path)%nat -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> - wf_stackframe stack -> - exists idx' s2', - (path_step ge pge path stack f sp rs m pc t s2' - \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) - \/ (path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - /\ (fn_path f)!(the_pc st) = Some O /\ path_step ge pge 0 stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') - ) - /\ match_states idx' s1' s2'. -Proof. - intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. - destruct idx as [ | idx]. - + (* path_step on normal_exit *) - Local Hint Resolve exec_normal_exit. - cutrewrite (path-0=path)%nat in PSTEP; [|omega]. - exploit path_last_step_complete; eauto. - intros (s2' & LSTEP & (idx' & MATCH)). - exists idx'; exists s2'; intuition eauto. - + (* stuttering step *) - exploit path_istate_complete_normal_exit; eauto. - unfold match_states; exists idx; exists (State stack f sp pc rs m); - intuition. - + (* one or two path_step on early_exit *) admit. - (* ISSUE: define two "path_step" *) -Admitted. - -Lemma step_noninst_complete s1 t s1' s2: - is_inst s1 = None -> - s1 = state_RTL s2 -> - RTL.step ge s1 t s1' -> - wf_stackframe (stack_of s2) -> - exists s2', step ge pge s2 t s2' /\ exists idx, match_states idx s1' s2'. -Proof. - intros H0 H1 H2 WFSTACK; destruct s2; subst; simpl in * |- *; try congruence; - inversion H2; clear H2; subst; try_simplify_someHyps; try congruence. - + (* exec_function_internal *) - destruct f; simpl in H3; inversion H3; subst; clear H3. - eexists; constructor 1. - * eapply exec_function_internal; eauto. - * unfold match_states, match_inst_states; simpl. - destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X. 2:{ exists O; intuition. } - exploit initialize_path; eauto. - destruct 1 as (path & Hpath). - exists path. constructor; auto. - econstructor; eauto. - - cutrewrite (path-path=O)%nat; simpl; eauto. - omega. - - simpl; auto. - + (* exec_function_external *) - destruct f; simpl in H3 |-; inversion H3; subst; clear H3. - eexists; constructor 1. - * apply exec_function_external; eauto. - * unfold match_states, match_inst_states; simpl. exists O; auto. - + (* exec_return *) - destruct stack eqn: Hstack; simpl in H1; inversion H1; clear H1; subst. - destruct s0 eqn: Hs0; simpl in H0; inversion H0; clear H0; subst. - eexists; constructor 1. - * apply exec_return. - * unfold match_states, match_inst_states; simpl. - rewrite wf_stackframe_cons in WFSTACK. - destruct WFSTACK as (H0 & H1); simpl in H0. - destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; intuition ]. - exploit initialize_path; eauto. - destruct 1 as (path & Hpath). - exists path. constructor; auto. - econstructor; eauto. - - cutrewrite (path-path=O)%nat; simpl; eauto. - omega. - - simpl; auto. -Qed. - - -Lemma step_complete s1 t s1' idx s2: - match_states idx s1 s2 -> - RTL.step ge s1 t s1' -> - exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. -Proof. - Local Hint Resolve plus_one plus_two exec_path. - unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - - destruct b. - * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. - clear His1. simpl in * |- *. - intros STEP; exploit path_step_complete; eauto. - destruct 1 as (idx' & s2' & H0 & H1). - eexists; eexists; eauto. - destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. - * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. - - intros; exploit step_noninst_complete; eauto. - intros (s2' & STEP & (idx0 & MATCH)). - exists idx0; exists s2'; intuition auto. -Qed. - -Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). -Proof. - eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). - constructor 1; simpl. - - apply lt_wf. - - unfold initial_state, match_states, match_inst_states. destruct 1; simpl; exists O. - destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. - exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. - - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. - destruct s2; simpl in * |- *; inversion H0; subst. - constructor. - - Local Hint Resolve star_refl. - intros; exploit step_complete; eauto. - destruct 1 as (idx' & s2' & X). - exists idx'. exists s2'. intuition (subst; eauto). - - intros id; destruct (senv_preserved p); simpl in * |-. intuition. -Qed. - -End COMPLETENESS. \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v new file mode 100644 index 00000000..93766f33 --- /dev/null +++ b/mppa_k1c/lib/RTLpath.v @@ -0,0 +1,1009 @@ +(** We introduce a data-structure extending the RTL CFG, in a control-flow graph over "traces" (in the sense of "trace-scheduling") + Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert: + a "path" is simply a list of successive nodes in the CFG (modulo some additional wellformness conditions). + + Actually, we extend syntactically the notion of RTL programs with a structure of "path_map": + this gives an alternative view of the CFG -- where "nodes" are paths instead of simple instructions. + Our wellformness condition on paths express that: + - the CFG on paths is wellformed: any successor of a given path points to another path (possibly the same). + - execution of a paths only emit single events. + + We represent each path only by a natural: the number of nodes in the path. These nodes are recovered from a static notion of "default successor". + This notion of path is thus incomplete. For example, if a path contains a whole loop (and for example, unrools it several times), + then this loop must be a suffix of the path. + + However: it is sufficient in order to represent superblocks (each superblock being represented as a path). + A superblock decomposition of the CFG exactly corresponds to the case where each node is in at most one path. + + Our goal is to provide two bisimulable semantics: + - one is simply the RTL semantics + - the other is based on a notion of "path-step": each path is executed in a single step. + + Remark that all analyses on RTL programs should thus be appliable for "free" also for RTLpath programs ! +*) + +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL Linking. + +Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end) + (at level 200, X ident, A at level 100, B at level 200) + : option_monad_scope. + +Local Open Scope option_monad_scope. + +(** * Syntax of RTLpath programs *) + +(** Internal instruction = instruction with a default successor in a path. *) + +Definition default_succ (i: instruction): option node := + match i with + | Inop s => Some s + | Iop op args res s => Some s + | Iload chunk addr args dst s => Some s + | Istore chunk addr args src s => Some s + | Icond cond args ifso ifnot => Some ifnot + | _ => None (* TODO: we could choose a successor for jumptable ? *) + end. + +Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *) + match i with + | Icond cond args ifso ifnot => Some ifso + | _ => None + end. + +(** Our notion of path. + + We do not formally require that the set of path is a partition of the CFG. + path may have intersections ! + + Moreover, we do not formally require that path have a single entry-point (a superblock structure) + + But, in practice, these properties are probably necessary in order to ensure the success of dynamic verification of scheduling. + + Here: we only require that each exit-point of a path is the entry-point of a path + (and that internal node of a path are internal instructions) +*) + + +(* By convention, we say that node [n] of the CFG is an entry-point of a path + if it is an entry of the path_map. + + Such a path of entry [n] is defined from a natural [path] representing the [path] default-successors of [n]. + + Remark: a path can loop several times in the CFG. + +*) + +Definition path_map: Type := PTree.t nat. + +Definition path_entry (c:code) (pm: path_map) (n: node): Prop + := pm!n = None -> c!n=None. + +Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := + | wf_last_node i pc: + c!pc = Some i -> + (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> + wellformed_path c pm 0 pc + | wf_internal_node path i pc pc': + c!pc = Some i -> + default_succ i = Some pc' -> + (forall n, early_exit i = Some n -> path_entry c pm n) -> + wellformed_path c pm path pc' -> + wellformed_path c pm (S path) pc. + +(* all paths defined from the path_map are wellformed *) +Definition wellformed_path_map (c:code) (pm: path_map): Prop := + forall n path, pm!n = Some path -> wellformed_path c pm path n. + +(** We "extend" the notion of RTL program with the additional structure for path. + + There is thus a trivial "forgetful functor" from RTLpath programs to RTL ones. +*) + +Record function : Type := + { fn_RTL:> RTL.function; + fn_path: path_map; + (* condition 1 below: the entry-point of the code is an entry-point of a path *) + fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); + (* condition 2 below: the path_map is well-formed *) + fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path + }. + +Definition fundef := AST.fundef function. +Definition program := AST.program fundef unit. +Definition genv := Genv.t fundef unit. + +Definition fundef_RTL (fu: fundef) : RTL.fundef := + match fu with + | Internal f => Internal f.(fn_RTL) + | External ef => External ef + end. +Coercion fundef_RTL: fundef >-> RTL.fundef. + +Definition program_RTL (p: program) : RTL.program := transform_program fundef_RTL p. +Coercion program_RTL: program >-> RTL.program. + +(** * Path-step semantics of RTLpath programs *) + +(* Semantics of internal instructions (mimicking RTL semantics) *) + +Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. + +Inductive stackframe : Type := + | Stackframe + (res: reg) (**r where to store the result *) + (f: function) (**r calling function *) + (sp: val) (**r stack pointer in calling function *) + (pc: node) (**r program point in calling function *) + (rs: regset) (**r register state in calling function *) + . + +Definition stf_RTL (st: stackframe): RTL.stackframe := + match st with + | Stackframe res f sp pc rs => RTL.Stackframe res f sp pc rs + end. + +Fixpoint stack_RTL (stack: list stackframe): list RTL.stackframe := + match stack with + | nil => nil + | cons stf stack' => cons (stf_RTL stf) (stack_RTL stack') + end. + +Inductive state : Type := + | State + (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 *) + | Callstate + (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 *) + | Returnstate + (stack: list stackframe) (**r call stack *) + (v: val) (**r return value for the call *) + (m: mem) (**r memory state *) + . + +Definition state_RTL (s: state): RTL.state := + match s with + | State stack f sp pc rs m => RTL.State (stack_RTL stack) f sp pc rs m + | Callstate stack f args m => RTL.Callstate (stack_RTL stack) f args m + | Returnstate stack v m => RTL.Returnstate (stack_RTL stack) v m + end. +Coercion state_RTL: state >-> RTL.state. + +Definition step_istate (ge: RTL.genv) (i: instruction) (stack: list stackframe) (f: function) (sp: val) (rs: regset) (m: mem): option istate := + match i with + | Inop pc' => Some (ist true pc' rs m) + | Iop op args res pc' => + SOME v <- eval_operation ge sp op rs##args m IN + Some (ist true pc' (rs#res <- v) m) + | Iload chunk addr args dst pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME v <- Mem.loadv chunk m a IN + Some (ist true pc' (rs#dst <- v) m) + | Istore chunk addr args src pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME m' <- Mem.storev chunk m a rs#src IN + Some (ist true pc' rs m') + | Icond cond args ifso ifnot => + SOME b <- eval_condition cond rs##args m IN + Some (ist (negb b) (if b then ifso else ifnot) rs m) + | _ => None (* TODO jumptable ? *) + end. + +(** Execution of a path in a single step *) + +Fixpoint path_step_istate ge (path:nat) stack (f: function) sp rs m pc: option istate := + match path with + | O => Some (ist true pc rs m) + | S p => + SOME i <- (fn_code f)!pc IN + SOME st <- step_istate ge i stack f sp rs m IN + if st.(continue) then + path_step_istate ge p stack f sp st.(the_rs) st.(the_mem) st.(the_pc) + else + Some st + end. + +Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef := + match ros with + | inl r => Genv.find_funct pge rs#r + | inr symb => + match Genv.find_symbol pge symb with + | None => None + | Some b => Genv.find_funct_ptr pge b + end + end. + +Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := + | exec_istate i sp pc rs m st: + (fn_code f)!pc = Some i -> + step_istate ge i stack f sp rs m = Some st -> + path_last_step ge pge stack f sp pc rs m + E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + | exec_Icall sp pc rs m sig ros args res pc' fd: + (fn_code f)!pc = Some(Icall sig ros args res pc') -> + find_function pge ros rs = Some fd -> + funsig fd = sig -> + path_last_step ge pge stack f sp pc rs m + E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m) + | exec_Itailcall stk pc rs m sig ros args fd m': + (fn_code f)!pc = Some(Itailcall sig ros args) -> + find_function pge ros rs = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Callstate stack fd rs##args m') + | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': + (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> + eval_builtin_args ge (fun r => rs#r) sp m args vargs -> + external_call ef ge vargs m t vres m' -> + path_last_step ge pge stack f sp pc rs m + t (State stack f sp pc' (regmap_setres res vres rs) m') + | exec_Ijumptable sp pc rs m arg tbl n pc': + (fn_code f)!pc = Some(Ijumptable arg tbl) -> + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + path_last_step ge pge stack f sp pc rs m + E0 (State stack f sp pc' rs m) + | exec_Ireturn stk pc rs m or m': + (fn_code f)!pc = Some(Ireturn or) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Returnstate stack (regmap_optget or Vundef rs) m'). + +Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := + | exec_early_exit st: + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = false -> + path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + | exec_normal_exit st t s: + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = true -> + path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> + path_step ge pge path stack f sp rs m pc t s. + +Inductive step ge pge: state -> trace -> state -> Prop := + | exec_path path stack f sp rs m pc t s: + (fn_path f)!pc = Some path -> + path_step ge pge path stack f sp rs m pc t s -> + step ge pge (State stack f sp pc rs m) t s + | exec_function_internal s f args m m' stk: + Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) -> + step ge pge (Callstate s (Internal f) args m) + E0 (State s + f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + m') + | exec_function_external s ef args res t m m': + external_call ef ge args m t res m' -> + step ge pge (Callstate s (External ef) args m) + t (Returnstate s res m') + | exec_return res f sp pc rs s vres m: + step ge pge (Returnstate (Stackframe res f sp pc rs :: s) vres m) + E0 (State s f sp pc (rs#res <- vres) m). + +Definition initial_state (p: program) (st:state): Prop + := RTL.initial_state p st. + +Definition final_state (st: state) (i:int): Prop + := RTL.final_state st i. + +Definition semantics (p: program) := + Semantics (step (Genv.globalenv (program_RTL p))) (initial_state p) final_state (Genv.globalenv p). + + +(** * Proving the bisimulation between (semantics p) and (RTL.semantics p). *) + +(** ** Preliminaries: simple tactics for option-monad *) + +Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B): + (forall x, e = Some x -> P (f x)) -> (e = None -> P None) -> (P (SOME x <- e IN f x)). +Proof. + intros; destruct e; simpl; auto. +Qed. + +Local Ltac inversion_SOME x := + try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]). + +Ltac simplify_someHyp := + match goal with + | H: None = Some _ |- _ => inversion H; clear H; subst + | H: Some _ = None |- _ => inversion H; clear H; subst + | H: ?t = ?t |- _ => clear H + | H: Some _ = Some _ |- _ => inversion H; clear H; subst + | H: Some _ <> None |- _ => clear H + | H: None <> Some _ |- _ => clear H + | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H + end. + +Ltac simplify_someHyps := + repeat (simplify_someHyp; simpl in * |- *). + +Ltac try_simplify_someHyps := + try (intros; simplify_someHyps; eauto). + +Ltac simplify_SOME x := + (repeat inversion_SOME x); try_simplify_someHyps. + + +Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. + +(** ** The easy way: Forward simulation of RTLpath by RTL + +This way can be viewed as a correctness property: all transitions in RTLpath are valid RTL transitions ! + +*) + +Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. + +Lemma step_istate_correct ge i stack f sp rs m st : + step_istate ge i stack f sp rs m = Some st -> + forall pc, (fn_code f)!pc = Some i -> + RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. +Qed. + +Local Hint Resolve star_refl. + +Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + induction path; simpl; try_simplify_someHyps. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:cont. + + intros; eapply star_step. + - eapply step_istate_correct; eauto. + - simpl; eauto. + - auto. + + intros; simplify_someHyp; eapply star_step. + - eapply step_istate_correct; eauto. + - simpl; eauto. + - auto. +Qed. + +Lemma path_step_istate_correct_early_exit ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = false -> + plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + destruct path; simpl; try_simplify_someHyps; try congruence. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:cont. + + intros; eapply plus_left. + - eapply step_istate_correct; eauto. + - eapply path_step_istate_correct; eauto. + - auto. + + intros X; inversion X; subst. + eapply plus_one. + eapply step_istate_correct; eauto. +Qed. + + +Section CORRECTNESS. + +Variable p: program. + +Lemma match_prog_RTL: match_program (fun _ f tf => tf = fundef_RTL f) eq p (program_RTL p). +Proof. + eapply match_transform_program; eauto. +Qed. + +Let pge := Genv.globalenv p. +Let ge := Genv.globalenv (program_RTL p). + +Lemma senv_preserved: Senv.equiv pge ge. +Proof (Genv.senv_match match_prog_RTL). + +Lemma symbols_preserved s: Genv.find_symbol ge s = Genv.find_symbol pge s. +Proof (Genv.find_symbol_match match_prog_RTL s). + +Lemma find_function_RTL_match ros rs fd: + find_function pge ros rs = Some fd -> RTL.find_function ge ros rs = Some (fundef_RTL fd). +Proof. + destruct ros; simpl. + + intro; exploit (Genv.find_funct_match match_prog_RTL); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; auto. + + rewrite symbols_preserved. + destruct (Genv.find_symbol pge i); simpl; try congruence. + intro; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; auto. +Qed. + +Local Hint Resolve step_istate_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. + +Lemma path_last_step_correct stack f sp pc rs m t s: + path_last_step ge pge stack f sp pc rs m t s -> + RTL.step ge (State stack f sp pc rs m) t s. +Proof. + destruct 1; try (eapply step_istate_correct); simpl; eauto. +Qed. + +Lemma path_step_correct path stack f sp pc rs m t s: + path_step ge pge path stack f sp rs m pc t s -> + plus RTL.step ge (State stack f sp pc rs m) t s. +Proof. + destruct 1. + + eapply path_step_istate_correct_early_exit; eauto. + + eapply plus_right. + eapply path_step_istate_correct; eauto. + eapply path_last_step_correct; eauto. + auto. +Qed. + +Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return. + +Lemma step_correct s t s': step ge pge s t s' -> plus RTL.step ge s t s'. +Proof. + destruct 1; try (eapply path_step_correct); simpl; eauto. +Qed. + +Theorem RTLpath_correct: forward_simulation (semantics p) (RTL.semantics p). +Proof. + eapply forward_simulation_plus with (match_states := fun s1 s2 => s2 = state_RTL s1); simpl; auto. + - apply senv_preserved. + - unfold initial_state; eauto. + - unfold final_state; intros; subst; eauto. + - intros; subst. eexists; intuition. + eapply step_correct; eauto. +Qed. + +End CORRECTNESS. + + + + +(** The hard way: Forward simulation of RTL by RTLpath + +This way can be viewed as a completeness property: all transitions in RTL can be represented as RTLpath transitions ! + +*) + +(* This lemma is probably needed to compose a pass from RTL -> RTLpath with other passes.*) +Lemma match_RTL_prog {LA: Linker fundef} {LV: Linker unit} p: match_program (fun _ f tf => f = fundef_RTL tf) eq (program_RTL p) p. +Proof. + unfold match_program, match_program_gen; intuition. + unfold program_RTL at 2; simpl. + generalize (prog_defs p). + induction l as [|a l]; simpl; eauto. + destruct a; simpl. + intros; eapply list_forall2_cons; eauto. + unfold match_ident_globdef; simpl; intuition; destruct g as [f|v]; simpl; eauto. + eapply match_globdef_var. destruct v; eauto. +Qed. + +(* Theory of wellformed paths *) + +Fixpoint nth_default_succ (c: code) (path:nat) (pc: node): option node := + match path with + | O => Some pc + | S path' => + SOME i <- c!pc IN + SOME pc' <- default_succ i IN + nth_default_succ c path' pc' + end. + +Lemma wellformed_suffix_path c pm path path': + (path' <= path)%nat -> + forall pc, wellformed_path c pm path pc -> + exists pc', nth_default_succ c (path-path') pc = Some pc' /\ wellformed_path c pm path' pc'. +Proof. + induction 1 as [|m]. + + intros; cutrewrite (path'-path'=O)%nat; [simpl;eauto|omega]. + + intros pc WF; cutrewrite (S m-path'=S (m-path'))%nat; [simpl;eauto|omega]. + inversion WF; subst; clear WF; intros; simplify_someHyps. + intros; simplify_someHyps; eauto. +Qed. + +Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := + SOME pc <- nth_default_succ c path pc IN + c!pc. + +Lemma final_node_path f path pc: + (fn_path f)!pc = Some path -> + exists i, nth_default_succ_inst (fn_code f) path pc = Some i + /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). +Proof. + intros; exploit fn_path_wf; eauto. + intro WF. + exploit (wellformed_suffix_path (fn_code f) (fn_path f) path O); omega || eauto. + destruct 1 as (pc' & NTH_SUCC & WF'); auto. + cutrewrite (path - 0 = path)%nat in NTH_SUCC; try omega. + unfold nth_default_succ_inst. + inversion WF'; clear WF'; subst. simplify_someHyps; eauto. +Qed. + +Lemma internal_node_path path f path0 pc: + (fn_path f)!pc = (Some path0) -> + (path < path0)%nat -> + exists i pc', + nth_default_succ_inst (fn_code f) path pc = Some i /\ + default_succ i = Some pc' /\ + (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). +Proof. + intros; exploit fn_path_wf; eauto. + intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } + destruct 1 as (pc' & NTH_SUCC & WF'). + cutrewrite (path0 - (path0 - path) = path)%nat in NTH_SUCC; try omega. + unfold nth_default_succ_inst. + inversion WF'; clear WF'; subst. { omega. } + simplify_someHyps; eauto. +Qed. + +Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. +Proof. + unfold path_entry; destruct pm!n; eauto. intuition congruence. +Qed. +Local Hint Resolve fn_entry_point_wf. +Local Opaque path_entry. + +Lemma step_istate_successors ge i stack f sp rs m st: + step_istate ge i stack f sp rs m = Some st -> + In (the_pc st) (successors_instr i). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. + destruct x; simpl in * |- *; intuition congruence. +Qed. + +Lemma step_istate_normal_exit ge i stack f sp rs m st: + step_istate ge i stack f sp rs m = Some st -> + st.(continue) = true -> + default_succ i = Some st.(the_pc). +Proof. + destruct i; simpl; try congruence; simplify_SOME x. + destruct x; simpl in * |- *; try congruence. +Qed. + +Lemma path_step_istate_normal_exit ge path stack f sp: forall rs m pc st, + st.(continue) = true -> + path_step_istate ge path stack f sp rs m pc = Some st -> + nth_default_succ (fn_code f) path pc = Some st.(the_pc). +Proof. + induction path; simpl. { try_simplify_someHyps. } + intros rs m pc st CONT; try_simplify_someHyps. + inversion_SOME i; intros Hi. + inversion_SOME st0; intros Hst0. + destruct (continue st0) eqn:X; try congruence. + try_simplify_someHyps. + intros; erewrite step_istate_normal_exit; eauto. +Qed. + +Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, + path_step_istate ge path stack f sp rs m pc = Some st -> + st.(continue) = true -> + (fn_code f)!(st.(the_pc)) = Some i -> + step_istate ge i stack f sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) stack f sp rs m pc. +Proof. + induction path. + + simpl; intros; try_simplify_someHyps. simplify_SOME st. + destruct st as [b]; destruct b; simpl; auto. + + intros rs m pc st i H. + simpl in H. + generalize H; clear H; simplify_SOME xx. + destruct (continue xx0) eqn: CONTxx0. + * intros; erewrite IHpath; eauto. + * intros; congruence. +Qed. + +(* USEFUL ? +Lemma path_step_istate_inv_normal ge path stack f sp: forall rs m pc st, + path_step_istate ge (S path) stack f sp rs m pc = Some st -> + (continue st)=true -> + exists st0 i, + path_step_istate ge path stack f sp rs m pc = Some st0 /\ + st0.(continue) = true /\ + (fn_code f)!(st0.(the_pc)) = Some i /\ + step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. +Proof. + simpl; intros rs m pc st. inversion_SOME i; inversion_SOME st0. + destruct (continue st0) eqn: CONT; try (intro; simplify_someHyps; congruence). + generalize rs m pc st i st0 CONT; + clear rs m pc st st i st0 CONT. + induction path; simpl; intros rs m pc st i st0. + - try_simplify_someHyps; simpl. + repeat eexists; intuition eauto. + destruct st0; simpl in * |- *; subst; auto. + - simplify_SOME yy. + rewrite CONT. + destruct (continue yy0) eqn: CONT0; try_simplify_someHyps; congruence || auto. +Qed. +*) + +Lemma path_step_istate_inversion_early ge path stack f sp: forall rs m pc st, + path_step_istate ge path stack f sp rs m pc = Some st -> + (continue st)=false -> + exists st0 i path0, + (path > path0)%nat /\ + path_step_istate ge path0 stack f sp rs m pc = Some st0 /\ + st0.(continue) = true /\ + (fn_code f)!(st0.(the_pc)) = Some i /\ + step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. +Proof. + induction path as [|path]; simpl. + - intros; try_simplify_someHyps; try congruence. + - intros rs m pc st; inversion_SOME i; inversion_SOME st0. + cutrewrite (path-0=path)%nat; try omega. + destruct (continue st0) eqn: CONT. + + intros STEP PC STEPS CONT0. exploit IHpath; eauto. + clear STEPS. + intros (st1 & i0 & path0 & BOUND & STEP1 & CONT1 & X1 & X2); auto. + exists st1. exists i0. exists (S path0). intuition. + simpl; try_simplify_someHyps. + rewrite CONT. auto. + + intros; try_simplify_someHyps; try congruence. + eexists. exists i. exists O; simpl. intuition eauto. + omega. +Qed. + +Lemma path_step_istate_resize ge path0 path1 stack f sp rs m pc st: + (path0 <= path1)%nat -> + path_step_istate ge path0 stack f sp rs m pc = Some st -> + (continue st)=false -> + path_step_istate ge path1 stack f sp rs m pc = Some st. +Proof. + induction 1 as [|path1]; simpl; auto. + intros PSTEP CONT. exploit IHle; auto. clear PSTEP IHle H path0. + generalize rs m pc st CONT; clear rs m pc st CONT. + induction path1 as [|path]; simpl; auto. + - intros; try_simplify_someHyps; try congruence. + - intros rs m pc st; inversion_SOME i; inversion_SOME st0; intros; try_simplify_someHyps. + destruct (continue st0) eqn: CONT0; eauto. +Qed. + +Inductive is_early_exit pc: instruction -> Prop := + | Icond_early_exit cond args ifnot: + is_early_exit pc (Icond cond args pc ifnot). + +Lemma step_istate_early_exit ge i stack f sp rs m st : + step_istate ge i stack f sp rs m = Some st -> + st.(continue) = false -> + st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. +Proof. + Local Hint Resolve Icond_early_exit. + destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. + destruct b; simpl in * |- *; intuition eauto. + congruence. +Qed. + +Section COMPLETENESS. + +Lemma step_istate_complete ge t i stack f sp rs m pc s': + RTL.step ge (State stack f sp pc rs m) t s' -> + (fn_code f)!pc = Some i -> + default_succ i <> None -> + t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; + (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. +Qed. + +Variable p: program. + +Let pge := Genv.globalenv p. +Let ge := Genv.globalenv (program_RTL p). + +Lemma find_funct_ptr_RTL_preserv b f: + Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0). +Proof. + intros; exploit (Genv.find_funct_ptr_match (match_RTL_prog p)); eauto. + destruct 1 as (cunit & tf & X & Y & Z); subst. + eauto. +Qed. + +Lemma find_RTL_function_match ros rs fd: + RTL.find_function ge ros rs = Some fd -> exists fd', fd = fundef_RTL fd' /\ find_function pge ros rs = Some fd'. +Proof. + destruct ros; simpl. + + intro; exploit (Genv.find_funct_match (match_RTL_prog p)); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; eauto. + + rewrite (symbols_preserved p); unfold pge. + destruct (Genv.find_symbol (Genv.globalenv p) i); simpl; try congruence. + intro; exploit find_funct_ptr_RTL_preserv; eauto. + intros (tf & H1 & H2); subst; eauto. +Qed. + + +Definition wf_stf (st: stackframe): Prop := + match st with + | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc + end. + +Definition wf_stackframe (stack: list stackframe): Prop := + forall st, List.In st stack -> wf_stf st. + +Lemma wf_stackframe_nil: wf_stackframe nil. +Proof. + unfold wf_stackframe; simpl. tauto. +Qed. +Local Hint Resolve wf_stackframe_nil. + +Lemma wf_stackframe_cons st stack: + wf_stackframe (st::stack) <-> (wf_stf st) /\ wf_stackframe stack. +Proof. + unfold wf_stackframe; simpl; intuition (subst; auto). +Qed. + +Definition stack_of (s: state): list stackframe := + match s with + | State stack f sp pc rs m => stack + | Callstate stack f args m => stack + | Returnstate stack v m => stack + end. + +Definition is_inst (s: RTL.state): option bool := + match s with + | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => Some true | None => Some false end + | _ => None + end. + +Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := + | State_match path stack f sp pc rs m s2: + (fn_path f)!pc = Some path -> + (idx <= path)%nat -> + path_step_istate ge (path-idx) stack f sp rs m pc = Some s2 -> + s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> + match_inst_states_goal idx s1 (State stack f sp pc rs m). + +Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop := + match is_inst s1 with + | Some b => b = true -> match_inst_states_goal idx s1 s2 + | None => s1 = state_RTL s2 + end. + +Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := + match_inst_states idx s1 s2 + /\ wf_stackframe (stack_of s2). + +Lemma path_istate_complete_normal_exit path idx stack f sp rs m pc st t s1': + path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + (S idx <= path)%nat -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). +Proof. + intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path-(S idx))); omega || eauto. + intros (i & pc' & Hi & Hpc & DUM). + unfold nth_default_succ_inst in Hi. + erewrite path_step_istate_normal_exit in Hi; eauto. + exploit step_istate_complete; congruence || eauto. + intros (SILENT & st0 & STEP0 & EQ). + intuition; subst; unfold match_inst_states; simpl. + destruct ((fn_code f) ! (the_pc st0)) eqn: X. 2: { intros; try congruence. } + intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. + cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. + erewrite <- path_step_istate_step_right; eauto. +Qed. + +Lemma path_last_step_complete path stack f sp rs m pc st t s1': + path_step_istate ge path stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + st.(continue) = true -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + wf_stackframe stack -> + exists s2', + (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' + /\ (exists idx', match_states idx' s1' s2'). +Proof. + Local Hint Resolve step_istate_successors list_nth_z_in. (* Hint for path_entry proofs *) + intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto. + intros (i & Hi & SUCCS). + unfold nth_default_succ_inst in Hi. + erewrite path_step_istate_normal_exit in Hi; eauto. + destruct (default_succ i) eqn:Hn0. + + (* exec_istate *) + exploit step_istate_complete; congruence || eauto. + intros (SILENT & st0 & STEP0 & EQ); subst. + exploit (exec_istate ge pge); eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! (the_pc st0)) eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. + - (* Icall *) + intros; exploit find_RTL_function_match; eauto. + intros (fd' & MATCHfd & Hfd'); subst. + exploit (exec_Icall ge pge); eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. + rewrite wf_stackframe_cons; intuition simpl; eauto. + - (* Itailcall *) + intros; exploit find_RTL_function_match; eauto. + intros (fd' & MATCHfd & Hfd'); subst. + exploit (exec_Itailcall ge pge); eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. + - (* Ibuiltin *) + intros; exploit exec_Ibuiltin; eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + - (* Ijumptable *) + intros; exploit exec_Ijumptable; eauto. + eexists; intuition eauto. + unfold match_states, match_inst_states; simpl. + destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. + exploit initialize_path; eauto. + intros (path0 & Hpath0). + exists path0; intuition eauto. + econstructor; eauto. + * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * simpl; eauto. + - (* Ireturn *) + intros; exploit exec_Ireturn; eauto. + eexists; intuition eauto. + eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. +Qed. + +Lemma path_step_complete stack f sp rs m pc t s1' idx path st: + path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> + (fn_path f)!pc = Some path -> + (idx <= path)%nat -> + RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + wf_stackframe stack -> + exists idx' s2', + (path_step ge pge path stack f sp rs m pc t s2' + \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) + \/ (path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + /\ (fn_path f)!(the_pc st) = Some O /\ path_step ge pge 0 stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') + ) + /\ match_states idx' s1' s2'. +Proof. + Local Hint Resolve exec_early_exit exec_normal_exit. + intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. + destruct idx as [ | idx]. + + (* path_step on normal_exit *) + cutrewrite (path-0=path)%nat in PSTEP; [|omega]. + exploit path_last_step_complete; eauto. + intros (s2' & LSTEP & (idx' & MATCH)). + exists idx'; exists s2'; intuition eauto. + + (* stuttering step *) + exploit path_istate_complete_normal_exit; eauto. + unfold match_states; exists idx; exists (State stack f sp pc rs m); + intuition. + + (* one or two path_step on early_exit *) + exploit path_step_istate_inversion_early; eauto. + intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). + exploit step_istate_early_exit; eauto. + intros (X1 & X2 & EARLY_EXIT). + destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. + exploit (internal_node_path path0); omega || eauto. + intros (i' & pc' & Hi' & Hpc' & ENTRY). + unfold nth_default_succ_inst in Hi'. + erewrite path_step_istate_normal_exit in Hi'; eauto. + clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. + exploit (path_step_istate_resize ge (path - idx)%nat path); eauto; try omega. + clear PSTEP; intros PSTEP. + destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY. + destruct ((fn_code f) ! pc0) as [i0|] eqn: Hi0. 2: { inversion RSTEP; intros; try_simplify_someHyps; try congruence. } + exploit initialize_path; eauto. + intros ([|path1] & Hpath1). + * (* two step case *) + exploit (path_last_step_complete 0 stack f sp (the_rs st0) (the_mem st0) pc0); simpl; eauto. + simpl; intros (s2' & LSTEP & (idx' & MATCH)). + exists idx'. exists s2'. constructor; auto. + right. right. intuition eauto. + { exploit (exec_early_exit ge pge path); eauto. } + { exploit exec_normal_exit. + 4:{eauto. } + - simpl; eauto. + - simpl; eauto. + - simpl; eauto. } + * (* single step case *) + exploit (path_istate_complete_normal_exit (S path1) path1 stack f sp (the_rs st0) (the_mem st0) pc0); simpl; auto. + - { cutrewrite (path1-path1=0)%nat; try omega. simpl; eauto. } + - simpl; auto. + - simpl; eauto. + - intuition subst. + repeat eexists; intuition eauto. +Qed. + +Lemma step_noninst_complete s1 t s1' s2: + is_inst s1 = None -> + s1 = state_RTL s2 -> + RTL.step ge s1 t s1' -> + wf_stackframe (stack_of s2) -> + exists s2', step ge pge s2 t s2' /\ exists idx, match_states idx s1' s2'. +Proof. + intros H0 H1 H2 WFSTACK; destruct s2; subst; simpl in * |- *; try congruence; + inversion H2; clear H2; subst; try_simplify_someHyps; try congruence. + + (* exec_function_internal *) + destruct f; simpl in H3; inversion H3; subst; clear H3. + eexists; constructor 1. + * eapply exec_function_internal; eauto. + * unfold match_states, match_inst_states; simpl. + destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X. 2:{ exists O; intuition. } + exploit initialize_path; eauto. + destruct 1 as (path & Hpath). + exists path. constructor; auto. + econstructor; eauto. + - cutrewrite (path-path=O)%nat; simpl; eauto. + omega. + - simpl; auto. + + (* exec_function_external *) + destruct f; simpl in H3 |-; inversion H3; subst; clear H3. + eexists; constructor 1. + * apply exec_function_external; eauto. + * unfold match_states, match_inst_states; simpl. exists O; auto. + + (* exec_return *) + destruct stack eqn: Hstack; simpl in H1; inversion H1; clear H1; subst. + destruct s0 eqn: Hs0; simpl in H0; inversion H0; clear H0; subst. + eexists; constructor 1. + * apply exec_return. + * unfold match_states, match_inst_states; simpl. + rewrite wf_stackframe_cons in WFSTACK. + destruct WFSTACK as (H0 & H1); simpl in H0. + destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; intuition ]. + exploit initialize_path; eauto. + destruct 1 as (path & Hpath). + exists path. constructor; auto. + econstructor; eauto. + - cutrewrite (path-path=O)%nat; simpl; eauto. + omega. + - simpl; auto. +Qed. + + +Lemma step_complete s1 t s1' idx s2: + match_states idx s1 s2 -> + RTL.step ge s1 t s1' -> + exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. +Proof. + Local Hint Resolve plus_one plus_two exec_path. + unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. + - destruct b. + * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. + clear His1. simpl in * |- *. + intros STEP; exploit path_step_complete; eauto. + destruct 1 as (idx' & s2' & H0 & H1). + eexists; eexists; eauto. + destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. + * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. + - intros; exploit step_noninst_complete; eauto. + intros (s2' & STEP & (idx0 & MATCH)). + exists idx0; exists s2'; intuition auto. +Qed. + +Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p). +Proof. + eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). + constructor 1; simpl. + - apply lt_wf. + - unfold initial_state, match_states, match_inst_states. destruct 1; simpl; exists O. + destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. + exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. + - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. + destruct s2; simpl in * |- *; inversion H0; subst. + constructor. + - Local Hint Resolve star_refl. + intros; exploit step_complete; eauto. + destruct 1 as (idx' & s2' & X). + exists idx'. exists s2'. intuition (subst; eauto). + - intros id; destruct (senv_preserved p); simpl in * |-. intuition. +Qed. + +End COMPLETENESS. \ No newline at end of file -- cgit From 9a2ea947c19256079a53d4b612c2cc0c5c2917bb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Sep 2019 23:00:58 +0200 Subject: RTLpath: add it into the configure script --- configure | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure b/configure index 3bdb5d2b..be1a8d4c 100755 --- a/configure +++ b/configure @@ -818,7 +818,7 @@ cat >> Makefile.config < Date: Tue, 10 Sep 2019 11:15:52 +0200 Subject: RTLpath: more comments and a bit of cleaning --- mppa_k1c/lib/RTLpath.v | 129 ++++++++++++++++++++++--------------------------- 1 file changed, 59 insertions(+), 70 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 93766f33..597e299d 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -1,4 +1,4 @@ -(** We introduce a data-structure extending the RTL CFG, in a control-flow graph over "traces" (in the sense of "trace-scheduling") +(** We introduce a data-structure extending the RTL CFG into a control-flow graph over "traces" (in the sense of "trace-scheduling") Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert: a "path" is simply a list of successive nodes in the CFG (modulo some additional wellformness conditions). @@ -247,7 +247,7 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me external_call ef ge vargs m t vres m' -> path_last_step ge pge stack f sp pc rs m t (State stack f sp pc' (regmap_setres res vres rs) m') - | exec_Ijumptable sp pc rs m arg tbl n pc': + | exec_Ijumptable sp pc rs m arg tbl n pc': (* TODO remove jumptable from here ? *) (fn_code f)!pc = Some(Ijumptable arg tbl) -> rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> @@ -332,12 +332,10 @@ Ltac simplify_someHyps := Ltac try_simplify_someHyps := try (intros; simplify_someHyps; eauto). +(* TODO: try to improve this tactic with a better control over names and inversion *) Ltac simplify_SOME x := (repeat inversion_SOME x); try_simplify_someHyps. - -Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. - (** ** The easy way: Forward simulation of RTLpath by RTL This way can be viewed as a correctness property: all transitions in RTLpath are valid RTL transitions ! @@ -392,6 +390,7 @@ Proof. eapply step_istate_correct; eauto. Qed. +Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. Section CORRECTNESS. @@ -581,6 +580,10 @@ Proof. intros; erewrite step_istate_normal_exit; eauto. Qed. + +(* TODO: the three following lemmas could maybe simplified by introducing an auxiliary + left-recursive definition equivalent to path_step_istate ? +*) Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, path_step_istate ge path stack f sp rs m pc = Some st -> st.(continue) = true -> @@ -598,30 +601,6 @@ Proof. * intros; congruence. Qed. -(* USEFUL ? -Lemma path_step_istate_inv_normal ge path stack f sp: forall rs m pc st, - path_step_istate ge (S path) stack f sp rs m pc = Some st -> - (continue st)=true -> - exists st0 i, - path_step_istate ge path stack f sp rs m pc = Some st0 /\ - st0.(continue) = true /\ - (fn_code f)!(st0.(the_pc)) = Some i /\ - step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. -Proof. - simpl; intros rs m pc st. inversion_SOME i; inversion_SOME st0. - destruct (continue st0) eqn: CONT; try (intro; simplify_someHyps; congruence). - generalize rs m pc st i st0 CONT; - clear rs m pc st st i st0 CONT. - induction path; simpl; intros rs m pc st i st0. - - try_simplify_someHyps; simpl. - repeat eexists; intuition eauto. - destruct st0; simpl in * |- *; subst; auto. - - simplify_SOME yy. - rewrite CONT. - destruct (continue yy0) eqn: CONT0; try_simplify_someHyps; congruence || auto. -Qed. -*) - Lemma path_step_istate_inversion_early ge path stack f sp: forall rs m pc st, path_step_istate ge path stack f sp rs m pc = Some st -> (continue st)=false -> @@ -665,7 +644,8 @@ Qed. Inductive is_early_exit pc: instruction -> Prop := | Icond_early_exit cond args ifnot: - is_early_exit pc (Icond cond args pc ifnot). + is_early_exit pc (Icond cond args pc ifnot) + . (* TODO add jumptable here ? *) Lemma step_istate_early_exit ge i stack f sp rs m st : step_istate ge i stack f sp rs m = Some st -> @@ -680,16 +660,6 @@ Qed. Section COMPLETENESS. -Lemma step_istate_complete ge t i stack f sp rs m pc s': - RTL.step ge (State stack f sp pc rs m) t s' -> - (fn_code f)!pc = Some i -> - default_succ i <> None -> - t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). -Proof. - intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; - (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. -Qed. - Variable p: program. Let pge := Genv.globalenv p. @@ -716,6 +686,7 @@ Proof. Qed. +(** *** Definition of well-formed stacks and of match_states *) Definition wf_stf (st: stackframe): Prop := match st with | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc @@ -743,6 +714,10 @@ Definition stack_of (s: state): list stackframe := | Returnstate stack v m => stack end. +(* Distinguish RTL.state into three cases: + - None: this is not a RTL.State. + - Some b: this is a RTL.State and b indicates whether there is an associated instruction in the CFG +*) Definition is_inst (s: RTL.state): option bool := match s with | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => Some true | None => Some false end @@ -767,7 +742,18 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := match_inst_states idx s1 s2 /\ wf_stackframe (stack_of s2). -Lemma path_istate_complete_normal_exit path idx stack f sp rs m pc st t s1': +(** *** Auxiliary lemmas of completeness *) +Lemma step_istate_complete t i stack f sp rs m pc s': + RTL.step ge (State stack f sp pc rs m) t s' -> + (fn_code f)!pc = Some i -> + default_succ i <> None -> + t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). +Proof. + intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; + (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. +Qed. + +Lemma stuttering path idx stack f sp rs m pc st t s1': path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (S idx <= path)%nat -> @@ -788,7 +774,7 @@ Proof. erewrite <- path_step_istate_step_right; eauto. Qed. -Lemma path_last_step_complete path stack f sp rs m pc st t s1': +Lemma normal_exit path stack f sp rs m pc st t s1': path_step_istate ge path stack f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> @@ -877,44 +863,47 @@ Proof. intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. + (* path_step on normal_exit *) - cutrewrite (path-0=path)%nat in PSTEP; [|omega]. - exploit path_last_step_complete; eauto. + cutrewrite (path-0=path)%nat in PSTEP; [|omega]. + exploit normal_exit; eauto. intros (s2' & LSTEP & (idx' & MATCH)). exists idx'; exists s2'; intuition eauto. + (* stuttering step *) - exploit path_istate_complete_normal_exit; eauto. + exploit stuttering; eauto. unfold match_states; exists idx; exists (State stack f sp pc rs m); intuition. - + (* one or two path_step on early_exit *) - exploit path_step_istate_inversion_early; eauto. - intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). - exploit step_istate_early_exit; eauto. - intros (X1 & X2 & EARLY_EXIT). - destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. - exploit (internal_node_path path0); omega || eauto. - intros (i' & pc' & Hi' & Hpc' & ENTRY). - unfold nth_default_succ_inst in Hi'. - erewrite path_step_istate_normal_exit in Hi'; eauto. - clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. + + (* one or two path_step on early_exit *) exploit (path_step_istate_resize ge (path - idx)%nat path); eauto; try omega. clear PSTEP; intros PSTEP. - destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY. - destruct ((fn_code f) ! pc0) as [i0|] eqn: Hi0. 2: { inversion RSTEP; intros; try_simplify_someHyps; try congruence. } - exploit initialize_path; eauto. - intros ([|path1] & Hpath1). + (* TODO for clarification: move the assert below into a separate lemma *) + assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). + { exploit path_step_istate_inversion_early; eauto. + intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). + exploit step_istate_early_exit; eauto. + intros (X1 & X2 & EARLY_EXIT). + destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. + exploit (internal_node_path path0); omega || eauto. + intros (i' & pc' & Hi' & Hpc' & ENTRY). + unfold nth_default_succ_inst in Hi'. + erewrite path_step_istate_normal_exit in Hi'; eauto. + clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. + destruct ((fn_code f) ! pc0) as [i0|] eqn: Hi0. + 2: { inversion RSTEP; intros; try_simplify_someHyps; try congruence. } + destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; + exploit initialize_path; eauto. + } + destruct HPATH0 as ([|path1] & Hpath1). * (* two step case *) - exploit (path_last_step_complete 0 stack f sp (the_rs st0) (the_mem st0) pc0); simpl; eauto. + exploit (normal_exit 0); simpl; eauto. simpl; intros (s2' & LSTEP & (idx' & MATCH)). exists idx'. exists s2'. constructor; auto. right. right. intuition eauto. - { exploit (exec_early_exit ge pge path); eauto. } - { exploit exec_normal_exit. - 4:{eauto. } - - simpl; eauto. - - simpl; eauto. - - simpl; eauto. } + (* now, prove the last step *) + exploit exec_normal_exit. 4:{ eauto. } + - simpl; eauto. + - simpl; eauto. + - simpl; eauto. * (* single step case *) - exploit (path_istate_complete_normal_exit (S path1) path1 stack f sp (the_rs st0) (the_mem st0) pc0); simpl; auto. + exploit (stuttering (S path1) path1 stack f sp (the_rs st) (the_mem st) (the_pc st)); simpl; auto. - { cutrewrite (path1-path1=0)%nat; try omega. simpl; eauto. } - simpl; auto. - simpl; eauto. @@ -967,7 +956,7 @@ Proof. - simpl; auto. Qed. - +(** *** The main completeness lemma and the simulation theorem...*) Lemma step_complete s1 t s1' idx s2: match_states idx s1 s2 -> RTL.step ge s1 t s1' -> @@ -1006,4 +995,4 @@ Proof. - intros id; destruct (senv_preserved p); simpl in * |-. intuition. Qed. -End COMPLETENESS. \ No newline at end of file +End COMPLETENESS. -- cgit From de208f6e67fb5fbc94ecd9f9e98804bcdeb14fa5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 12 Sep 2019 11:13:30 +0200 Subject: RTLpath: simplifying (a bit) the proof by simplifying a bit a pedantic condition on wellformed path. --- mppa_k1c/lib/RTLpath.v | 82 +++++++++++++++++++------------------------------- 1 file changed, 31 insertions(+), 51 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 597e299d..9c8abbe2 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -67,8 +67,7 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, *) -(* By convention, we say that node [n] of the CFG is an entry-point of a path - if it is an entry of the path_map. +(* By convention, we say that node [n] is the entry-point of a path if it is a key of the path_map. Such a path of entry [n] is defined from a natural [path] representing the [path] default-successors of [n]. @@ -78,18 +77,21 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, Definition path_map: Type := PTree.t nat. -Definition path_entry (c:code) (pm: path_map) (n: node): Prop - := pm!n = None -> c!n=None. +Definition path_entry (*c:code*) (pm: path_map) (n: node): Prop + := pm!n = None -> False. +(* NB: In theory, we could relax the above [False] into [c!n=None], but this seems PEDANTIC in practice ! + Indeed, in practice, the CFG will not have dandling node ! +*) Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := | wf_last_node i pc: c!pc = Some i -> - (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> + (forall n, List.In n (successors_instr i) -> path_entry (*c*) pm n) -> wellformed_path c pm 0 pc | wf_internal_node path i pc pc': c!pc = Some i -> default_succ i = Some pc' -> - (forall n, early_exit i = Some n -> path_entry c pm n) -> + (forall n, early_exit i = Some n -> path_entry (*c*) pm n) -> wellformed_path c pm path pc' -> wellformed_path c pm (S path) pc. @@ -106,7 +108,7 @@ Record function : Type := { fn_RTL:> RTL.function; fn_path: path_map; (* condition 1 below: the entry-point of the code is an entry-point of a path *) - fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); + fn_entry_point_wf: path_entry (*fn_RTL.(fn_code)*) fn_path fn_RTL.(fn_entrypoint); (* condition 2 below: the path_map is well-formed *) fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path }. @@ -514,7 +516,7 @@ Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := Lemma final_node_path f path pc: (fn_path f)!pc = Some path -> exists i, nth_default_succ_inst (fn_code f) path pc = Some i - /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). + /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF. @@ -531,7 +533,7 @@ Lemma internal_node_path path f path0 pc: exists i pc', nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ - (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). + (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } @@ -542,7 +544,7 @@ Proof. simplify_someHyps; eauto. Qed. -Lemma initialize_path c pm n i: c!n = Some i -> path_entry c pm n -> exists path, pm!n = Some path. +Lemma initialize_path (*c*) pm n (*i*): (* PEDANTIC: c!n = Some i -> *) path_entry (*c*) pm n -> exists path, pm!n = Some path. Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. @@ -689,7 +691,7 @@ Qed. (** *** Definition of well-formed stacks and of match_states *) Definition wf_stf (st: stackframe): Prop := match st with - | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc + | Stackframe res f sp pc rs => path_entry (*f.(fn_code)*) f.(fn_path) pc end. Definition wf_stackframe (stack: list stackframe): Prop := @@ -714,14 +716,10 @@ Definition stack_of (s: state): list stackframe := | Returnstate stack v m => stack end. -(* Distinguish RTL.state into three cases: - - None: this is not a RTL.State. - - Some b: this is a RTL.State and b indicates whether there is an associated instruction in the CFG -*) -Definition is_inst (s: RTL.state): option bool := +Definition is_inst (s: RTL.state): bool := match s with - | RTL.State stack f sp pc rs m => match (fn_code f)!pc with Some _ => Some true | None => Some false end - | _ => None + | RTL.State stack f sp pc rs m => true + | _ => false end. Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := @@ -733,10 +731,7 @@ Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := match_inst_states_goal idx s1 (State stack f sp pc rs m). Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop := - match is_inst s1 with - | Some b => b = true -> match_inst_states_goal idx s1 s2 - | None => s1 = state_RTL s2 - end. + if is_inst s1 then match_inst_states_goal idx s1 s2 else s1 = state_RTL s2. Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := match_inst_states idx s1 s2 @@ -768,7 +763,6 @@ Proof. exploit step_istate_complete; congruence || eauto. intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. - destruct ((fn_code f) ! (the_pc st0)) eqn: X. 2: { intros; try congruence. } intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. erewrite <- path_step_istate_step_right; eauto. @@ -796,9 +790,7 @@ Proof. exploit (exec_istate ge pge); eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! (the_pc st0)) eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). + destruct (initialize_path (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -821,9 +813,7 @@ Proof. intros; exploit exec_Ibuiltin; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). + destruct (initialize_path (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -832,9 +822,7 @@ Proof. intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct ((fn_code f) ! pc') eqn: X; simpl; [|exists O; intuition ]. - exploit initialize_path; eauto. - intros (path0 & Hpath0). + destruct (initialize_path (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -876,7 +864,8 @@ Proof. clear PSTEP; intros PSTEP. (* TODO for clarification: move the assert below into a separate lemma *) assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). - { exploit path_step_istate_inversion_early; eauto. + { clear RSTEP. + exploit path_step_istate_inversion_early; eauto. intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). exploit step_istate_early_exit; eauto. intros (X1 & X2 & EARLY_EXIT). @@ -886,10 +875,8 @@ Proof. unfold nth_default_succ_inst in Hi'. erewrite path_step_istate_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. - destruct ((fn_code f) ! pc0) as [i0|] eqn: Hi0. - 2: { inversion RSTEP; intros; try_simplify_someHyps; try congruence. } destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; - exploit initialize_path; eauto. + destruct (initialize_path (fn_path f) pc0); eauto. } destruct HPATH0 as ([|path1] & Hpath1). * (* two step case *) @@ -912,7 +899,7 @@ Proof. Qed. Lemma step_noninst_complete s1 t s1' s2: - is_inst s1 = None -> + is_inst s1 = false -> s1 = state_RTL s2 -> RTL.step ge s1 t s1' -> wf_stackframe (stack_of s2) -> @@ -925,9 +912,7 @@ Proof. eexists; constructor 1. * eapply exec_function_internal; eauto. * unfold match_states, match_inst_states; simpl. - destruct ((fn_code (fn_RTL f)) ! (fn_entrypoint (fn_RTL f))) eqn: X. 2:{ exists O; intuition. } - exploit initialize_path; eauto. - destruct 1 as (path & Hpath). + destruct (initialize_path (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. @@ -946,9 +931,7 @@ Proof. * unfold match_states, match_inst_states; simpl. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. - destruct ((fn_code (fn_RTL f0)) ! pc0) eqn: X; [|exists O; intuition ]. - exploit initialize_path; eauto. - destruct 1 as (path & Hpath). + destruct (initialize_path (fn_path f0) pc0) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. @@ -964,14 +947,11 @@ Lemma step_complete s1 t s1' idx s2: Proof. Local Hint Resolve plus_one plus_two exec_path. unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - - destruct b. - * destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst. - clear His1. simpl in * |- *. - intros STEP; exploit path_step_complete; eauto. - destruct 1 as (idx' & s2' & H0 & H1). - eexists; eexists; eauto. - destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. - * intros X; generalize His1; clear His1. destruct X; simpl; try_simplify_someHyps; try congruence. + - clear His1; destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst; simpl in * |- *. + intros STEP; exploit path_step_complete; eauto. + intros (idx' & s2' & H0 & H1). + eexists; eexists; eauto. + destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. - intros; exploit step_noninst_complete; eauto. intros (s2' & STEP & (idx0 & MATCH)). exists idx0; exists s2'; intuition auto. -- cgit From 667ea52dc75c9e8422a8dbe99848d22a75bba1d4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 16 Oct 2019 15:12:25 +0200 Subject: git checkout 64a80f81297fb20c4f952d4b36cd0ae5d5da8f1e -- cparser/Machine.mli --- cparser/Machine.mli | 1 - 1 file changed, 1 deletion(-) diff --git a/cparser/Machine.mli b/cparser/Machine.mli index ea25c4f6..24d36e6c 100644 --- a/cparser/Machine.mli +++ b/cparser/Machine.mli @@ -88,7 +88,6 @@ val arm_bigendian : t val rv32 : t val rv64 : t val mppa_k1c : t -val aarch64 : t val gcc_extensions : t -> t val compcert_interpreter : t -> t -- cgit From 8b8f27f6d507a6d48da5b45c937ef3c66e59d67e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 16 Oct 2019 15:16:06 +0200 Subject: fix incorrect merge of RTLpath with mppa-work ? --- driver/Frontend.ml | 1 - 1 file changed, 1 deletion(-) diff --git a/driver/Frontend.ml b/driver/Frontend.ml index b9db0d23..2584db90 100644 --- a/driver/Frontend.ml +++ b/driver/Frontend.ml @@ -117,7 +117,6 @@ let init () = then Machine.rv64 else Machine.rv32 | "mppa_k1c" -> Machine.mppa_k1c - | "aarch64" -> Machine.aarch64 | _ -> assert false end; Env.set_builtins C2C.builtins; -- cgit From b985c9f2309d942c4617d91cae5b36969b5173f3 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 16 Oct 2019 19:49:12 +0200 Subject: experimenting symbolic execution for RTL --- mppa_k1c/lib/RTLpathSE_theory.v | 111 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 111 insertions(+) create mode 100644 mppa_k1c/lib/RTLpathSE_theory.v diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v new file mode 100644 index 00000000..3bd5bb1b --- /dev/null +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -0,0 +1,111 @@ +(* A theory of symbolic execution on RTLpath + +NB: an efficient implementation with hash-consing will be defined in ... + +*) + +Require Import Coqlib Maps. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL RTLpath. + +Record sgenv := { + the_ge: RTL.genv; + the_sp: val +}. + +(* symbolic value *) +Inductive sval := + | Sinput (r: reg) + | Sop (op:operation) (lsv: list_sval) (* TODO: pb des fausses dependances sur la memoire, via comparaisons de pointeurs ? *) + | Sload (sm: smem) (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. + +Local Open Scope option_monad_scope. + +Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs: regset) (m0: mem): option val := + match sv with + | Sinput r => Some (rs#r) + | Sop op l => + SOME args <- list_sval_eval sge l rs m0 IN + eval_operation sge.(the_ge) sge.(the_sp) op args m0 + | Sload sm chunk addr lsv => + SOME args <- list_sval_eval sge lsv rs m0 IN + SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN + SOME m <- smem_eval sge sm rs m0 IN + Mem.loadv chunk m a + end +with list_sval_eval (sge: sgenv) (lsv: list_sval) (rs: regset) (m0: mem): option (list val) := + match lsv with + | Snil => Some nil + | Scons sv lsv' => + SOME v <- sval_eval sge sv rs m0 IN + SOME lv <- list_sval_eval sge lsv' rs m0 IN + Some (v::lv) + end +with smem_eval (sge: sgenv) (sm: smem) (rs: regset) (m0: mem): option mem := + match sm with + | Sinit => Some m0 + | Sstore sm chunk addr lsv srce => + SOME args <- list_sval_eval sge lsv rs m0 IN + SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN + SOME m <- smem_eval sge sm rs m0 IN + SOME sv <- sval_eval sge srce rs m0 IN + Mem.storev chunk m a sv + end. + +Record local_istate := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. + +(* +Definition assign (rs: reg -> sval) (x:reg) (sv: sval): reg -> sval + := fun y => if Pos.eq_dec x y then sv else rs y. +*) + +Record cond_istate := { the_cond: condition; cond_args: list_sval; cond_ist: local_istate; cond_ifso: node }. + +(* symbolic internal state *) +Record s_istate := Sist { the_pc: node; early_exit: list cond_istate; curr: local_istate }. + +Fixpoint list_val_inj (l: list sval): list_sval := + match l with + | nil => Snil + | v::l => Scons v (list_val_inj l) + end. + +Definition sreg_set (st:local_istate) r (sv:sval) := + {| pre:=(fun _ _ _ => False) ; (* TODO *) + the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; + the_smem:= st.(the_smem)|}. + +Definition symb_istep (i: instruction) (st: s_istate) : option s_istate := + match i with + | Inop pc' => Some (Sist pc' st.(early_exit) st.(curr)) + | Iop op args res pc' => + let rs := st.(curr).(the_sreg) in + let vargs := list_val_inj (List.map rs args) in + let sv := Sop op vargs in + Some (Sist pc' st.(early_exit) (sreg_set st.(curr) res sv)) +(* | Iload chunk addr args dst pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME v <- Mem.loadv chunk m a IN + Some (ist true pc' (rs#dst <- v) m) + | Istore chunk addr args src pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME m' <- Mem.storev chunk m a rs#src IN + Some (ist true pc' rs m') + | Icond cond args ifso ifnot => + SOME b <- eval_condition cond rs##args m IN + Some (ist (negb b) (if b then ifso else ifnot) rs m) +*) + | _ => None (* TODO jumptable ? *) + end. -- cgit From 933aeeda897da5076f5a53fe6752a860c30a29bb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 08:10:01 +0200 Subject: fix configure broken by merge --- configure | 3 --- 1 file changed, 3 deletions(-) diff --git a/configure b/configure index ede8020f..be6f2180 100755 --- a/configure +++ b/configure @@ -470,10 +470,7 @@ if test "$arch" = "mppa_k1c"; then system="linux" fi -<<<<<<< HEAD -======= # ->>>>>>> mppa-work # AArch64 (ARMv8 64 bits) Target Configuration # if test "$arch" = "aarch64"; then -- cgit From b20a23a5b6bef28bc0415055a5e6248d4f7d7b27 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 08:51:15 +0200 Subject: internal step of symbolic execution: first draft --- mppa_k1c/lib/RTLpathSE_theory.v | 72 +++++++++++++++++++++++++---------------- 1 file changed, 44 insertions(+), 28 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 3bd5bb1b..ffaff8ac 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -31,6 +31,12 @@ 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_val_inj (l: list sval): list_sval := + match l with + | nil => Snil + | v::l => Scons v (list_val_inj l) + end. + Local Open Scope option_monad_scope. Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs: regset) (m0: mem): option val := @@ -66,46 +72,56 @@ with smem_eval (sge: sgenv) (sm: smem) (rs: regset) (m0: mem): option mem := Record local_istate := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. -(* -Definition assign (rs: reg -> sval) (x:reg) (sv: sval): reg -> sval - := fun y => if Pos.eq_dec x y then sv else rs y. -*) - Record cond_istate := { the_cond: condition; cond_args: list_sval; cond_ist: local_istate; cond_ifso: node }. (* symbolic internal state *) -Record s_istate := Sist { the_pc: node; early_exit: list cond_istate; curr: local_istate }. +Record s_istate := { the_pc: node; early_exit: list cond_istate; curr: local_istate + cond_istate }. -Fixpoint list_val_inj (l: list sval): list_sval := - match l with - | nil => Snil - | v::l => Scons v (list_val_inj l) +Definition get_curr (st:s_istate): local_istate := + match st.(curr) with + | inl x => x + | inr x => x.(cond_ist) end. +Definition sist (st: s_istate) (pc: node) (nxt: local_istate + cond_istate) := + match st.(curr) with + | inl _ => {| the_pc := pc; early_exit := st.(early_exit); curr:= nxt |} + | inr cnd => {| the_pc := pc; early_exit := cnd::st.(early_exit); curr:=nxt |} + end. + Definition sreg_set (st:local_istate) r (sv:sval) := - {| pre:=(fun _ _ _ => False) ; (* TODO *) + {| pre:=(fun _ _ _ => True ) ; (* TODO *) the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; the_smem:= st.(the_smem)|}. +Definition smem_set (st:local_istate) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval) := + {| pre:=(fun _ _ _ => True ) ; (* TODO *) + the_sreg:= st.(the_sreg); + the_smem:= Sstore st.(the_smem) chunk addr lsv srce |}. + +(* one internal step of symbolic execution *) Definition symb_istep (i: instruction) (st: s_istate) : option s_istate := match i with - | Inop pc' => Some (Sist pc' st.(early_exit) st.(curr)) - | Iop op args res pc' => - let rs := st.(curr).(the_sreg) in - let vargs := list_val_inj (List.map rs args) in - let sv := Sop op vargs in - Some (Sist pc' st.(early_exit) (sreg_set st.(curr) res sv)) -(* | Iload chunk addr args dst pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME v <- Mem.loadv chunk m a IN - Some (ist true pc' (rs#dst <- v) m) + | Inop pc' => Some (sist st pc' (inl (get_curr st))) + | Iop op args dst pc' => + let prev := get_curr st in + let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let next := sreg_set prev dst (Sop op vargs) in + Some (sist st pc' (inl next)) + | Iload chunk addr args dst pc' => + let prev := get_curr st in + let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in + Some (sist st pc' (inl next)) | Istore chunk addr args src pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME m' <- Mem.storev chunk m a rs#src IN - Some (ist true pc' rs m') - | Icond cond args ifso ifnot => - SOME b <- eval_condition cond rs##args m IN - Some (ist (negb b) (if b then ifso else ifnot) rs m) -*) + let prev := get_curr st in + let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let next := smem_set prev chunk addr vargs (prev.(the_sreg) src) in + Some (sist st pc' (inl next)) + | Icond cond args ifso ifnot => + let prev := get_curr st in + let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let next := {| the_cond:=cond; cond_args:=vargs; cond_ist := prev; cond_ifso := ifso |} in + Some (sist st ifnot (inr next)) | _ => None (* TODO jumptable ? *) end. -- cgit From 1c15144b104b954695c85a077a6ab99bedb11c98 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 10:17:38 +0200 Subject: remove useless dependencies --- mppa_k1c/lib/RTLpath.v | 152 ++++++++++++++++++++++++------------------------- 1 file changed, 76 insertions(+), 76 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 9c8abbe2..a0365344 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -133,6 +133,50 @@ Coercion program_RTL: program >-> RTL.program. Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. +Definition step_istate (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := + match i with + | Inop pc' => Some (ist true pc' rs m) + | Iop op args res pc' => + SOME v <- eval_operation ge sp op rs##args m IN + Some (ist true pc' (rs#res <- v) m) + | Iload chunk addr args dst pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME v <- Mem.loadv chunk m a IN + Some (ist true pc' (rs#dst <- v) m) + | Istore chunk addr args src pc' => + SOME a <- eval_addressing ge sp addr rs##args IN + SOME m' <- Mem.storev chunk m a rs#src IN + Some (ist true pc' rs m') + | Icond cond args ifso ifnot => + SOME b <- eval_condition cond rs##args m IN + Some (ist (negb b) (if b then ifso else ifnot) rs m) + | _ => None (* TODO jumptable ? *) + end. + +(** Execution of a path in a single step *) + +Fixpoint path_step_istate ge (path:nat) (f: function) sp rs m pc: option istate := + match path with + | O => Some (ist true pc rs m) + | S p => + SOME i <- (fn_code f)!pc IN + SOME st <- step_istate ge i sp rs m IN + if st.(continue) then + path_step_istate ge p f sp st.(the_rs) st.(the_mem) st.(the_pc) + else + Some st + end. + +Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef := + match ros with + | inl r => Genv.find_funct pge rs#r + | inr symb => + match Genv.find_symbol pge symb with + | None => None + | Some b => Genv.find_funct_ptr pge b + end + end. + Inductive stackframe : Type := | Stackframe (res: reg) (**r where to store the result *) @@ -180,54 +224,10 @@ Definition state_RTL (s: state): RTL.state := end. Coercion state_RTL: state >-> RTL.state. -Definition step_istate (ge: RTL.genv) (i: instruction) (stack: list stackframe) (f: function) (sp: val) (rs: regset) (m: mem): option istate := - match i with - | Inop pc' => Some (ist true pc' rs m) - | Iop op args res pc' => - SOME v <- eval_operation ge sp op rs##args m IN - Some (ist true pc' (rs#res <- v) m) - | Iload chunk addr args dst pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME v <- Mem.loadv chunk m a IN - Some (ist true pc' (rs#dst <- v) m) - | Istore chunk addr args src pc' => - SOME a <- eval_addressing ge sp addr rs##args IN - SOME m' <- Mem.storev chunk m a rs#src IN - Some (ist true pc' rs m') - | Icond cond args ifso ifnot => - SOME b <- eval_condition cond rs##args m IN - Some (ist (negb b) (if b then ifso else ifnot) rs m) - | _ => None (* TODO jumptable ? *) - end. - -(** Execution of a path in a single step *) - -Fixpoint path_step_istate ge (path:nat) stack (f: function) sp rs m pc: option istate := - match path with - | O => Some (ist true pc rs m) - | S p => - SOME i <- (fn_code f)!pc IN - SOME st <- step_istate ge i stack f sp rs m IN - if st.(continue) then - path_step_istate ge p stack f sp st.(the_rs) st.(the_mem) st.(the_pc) - else - Some st - end. - -Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef := - match ros with - | inl r => Genv.find_funct pge rs#r - | inr symb => - match Genv.find_symbol pge symb with - | None => None - | Some b => Genv.find_funct_ptr pge b - end - end. - Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := | exec_istate i sp pc rs m st: (fn_code f)!pc = Some i -> - step_istate ge i stack f sp rs m = Some st -> + step_istate ge i sp rs m = Some st -> path_last_step ge pge stack f sp pc rs m E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) | exec_Icall sp pc rs m sig ros args res pc' fd: @@ -263,11 +263,11 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := | exec_early_exit st: - path_step_istate ge path stack f sp rs m pc = Some st -> + path_step_istate ge path f sp rs m pc = Some st -> st.(continue) = false -> path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) | exec_normal_exit st t s: - path_step_istate ge path stack f sp rs m pc = Some st -> + path_step_istate ge path f sp rs m pc = Some st -> st.(continue) = true -> path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> path_step ge pge path stack f sp rs m pc t s. @@ -346,9 +346,9 @@ This way can be viewed as a correctness property: all transitions in RTLpath are Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. -Lemma step_istate_correct ge i stack f sp rs m st : - step_istate ge i stack f sp rs m = Some st -> - forall pc, (fn_code f)!pc = Some i -> +Lemma step_istate_correct ge i stack (f:function) sp rs m st : + step_istate ge i sp rs m = Some st -> + forall pc, (fn_code f)!pc = Some i -> RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. destruct i; simpl; try congruence; simplify_SOME x. @@ -357,7 +357,7 @@ Qed. Local Hint Resolve star_refl. Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> + path_step_istate ge path f sp rs m pc = Some st -> star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. induction path; simpl; try_simplify_someHyps. @@ -375,8 +375,8 @@ Proof. Qed. Lemma path_step_istate_correct_early_exit ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> - st.(continue) = false -> + path_step_istate ge path f sp rs m pc = Some st -> + st.(continue) = false -> plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. destruct path; simpl; try_simplify_someHyps; try congruence. @@ -551,16 +551,16 @@ Qed. Local Hint Resolve fn_entry_point_wf. Local Opaque path_entry. -Lemma step_istate_successors ge i stack f sp rs m st: - step_istate ge i stack f sp rs m = Some st -> +Lemma step_istate_successors ge i sp rs m st: + step_istate ge i sp rs m = Some st -> In (the_pc st) (successors_instr i). Proof. destruct i; simpl; try congruence; simplify_SOME x. destruct x; simpl in * |- *; intuition congruence. Qed. -Lemma step_istate_normal_exit ge i stack f sp rs m st: - step_istate ge i stack f sp rs m = Some st -> +Lemma step_istate_normal_exit ge i sp rs m st: + step_istate ge i sp rs m = Some st -> st.(continue) = true -> default_succ i = Some st.(the_pc). Proof. @@ -568,9 +568,9 @@ Proof. destruct x; simpl in * |- *; try congruence. Qed. -Lemma path_step_istate_normal_exit ge path stack f sp: forall rs m pc st, +Lemma path_step_istate_normal_exit ge path f sp: forall rs m pc st, st.(continue) = true -> - path_step_istate ge path stack f sp rs m pc = Some st -> + path_step_istate ge path f sp rs m pc = Some st -> nth_default_succ (fn_code f) path pc = Some st.(the_pc). Proof. induction path; simpl. { try_simplify_someHyps. } @@ -586,11 +586,11 @@ Qed. (* TODO: the three following lemmas could maybe simplified by introducing an auxiliary left-recursive definition equivalent to path_step_istate ? *) -Lemma path_step_istate_step_right ge path stack f sp: forall rs m pc st i, - path_step_istate ge path stack f sp rs m pc = Some st -> +Lemma path_step_istate_step_right ge path f sp: forall rs m pc st i, + path_step_istate ge path f sp rs m pc = Some st -> st.(continue) = true -> (fn_code f)!(st.(the_pc)) = Some i -> - step_istate ge i stack f sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) stack f sp rs m pc. + step_istate ge i sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) f sp rs m pc. Proof. induction path. + simpl; intros; try_simplify_someHyps. simplify_SOME st. @@ -603,15 +603,15 @@ Proof. * intros; congruence. Qed. -Lemma path_step_istate_inversion_early ge path stack f sp: forall rs m pc st, - path_step_istate ge path stack f sp rs m pc = Some st -> +Lemma path_step_istate_inversion_early ge path f sp: forall rs m pc st, + path_step_istate ge path f sp rs m pc = Some st -> (continue st)=false -> exists st0 i path0, (path > path0)%nat /\ - path_step_istate ge path0 stack f sp rs m pc = Some st0 /\ + path_step_istate ge path0 f sp rs m pc = Some st0 /\ st0.(continue) = true /\ (fn_code f)!(st0.(the_pc)) = Some i /\ - step_istate ge i stack f sp st0.(the_rs) st0.(the_mem) = Some st. + step_istate ge i sp st0.(the_rs) st0.(the_mem) = Some st. Proof. induction path as [|path]; simpl. - intros; try_simplify_someHyps; try congruence. @@ -629,11 +629,11 @@ Proof. omega. Qed. -Lemma path_step_istate_resize ge path0 path1 stack f sp rs m pc st: +Lemma path_step_istate_resize ge path0 path1 f sp rs m pc st: (path0 <= path1)%nat -> - path_step_istate ge path0 stack f sp rs m pc = Some st -> + path_step_istate ge path0 f sp rs m pc = Some st -> (continue st)=false -> - path_step_istate ge path1 stack f sp rs m pc = Some st. + path_step_istate ge path1 f sp rs m pc = Some st. Proof. induction 1 as [|path1]; simpl; auto. intros PSTEP CONT. exploit IHle; auto. clear PSTEP IHle H path0. @@ -649,8 +649,8 @@ Inductive is_early_exit pc: instruction -> Prop := is_early_exit pc (Icond cond args pc ifnot) . (* TODO add jumptable here ? *) -Lemma step_istate_early_exit ge i stack f sp rs m st : - step_istate ge i stack f sp rs m = Some st -> +Lemma step_istate_early_exit ge i sp rs m st : + step_istate ge i sp rs m = Some st -> st.(continue) = false -> st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. Proof. @@ -726,7 +726,7 @@ Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> (idx <= path)%nat -> - path_step_istate ge (path-idx) stack f sp rs m pc = Some s2 -> + path_step_istate ge (path-idx) f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> match_inst_states_goal idx s1 (State stack f sp pc rs m). @@ -742,14 +742,14 @@ Lemma step_istate_complete t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> (fn_code f)!pc = Some i -> default_succ i <> None -> - t = E0 /\ exists st, step_istate ge i stack f sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + t = E0 /\ exists st, step_istate ge i sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. Qed. Lemma stuttering path idx stack f sp rs m pc st t s1': - path_step_istate ge (path-(S idx)) stack f sp rs m pc = Some st -> + path_step_istate ge (path-(S idx)) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (S idx <= path)%nat -> st.(continue) = true -> @@ -769,7 +769,7 @@ Proof. Qed. Lemma normal_exit path stack f sp rs m pc st t s1': - path_step_istate ge path stack f sp rs m pc = Some st -> + path_step_istate ge path f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> @@ -834,7 +834,7 @@ Proof. Qed. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: - path_step_istate ge (path-idx) stack f sp rs m pc = Some st -> + path_step_istate ge (path-idx) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (idx <= path)%nat -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> -- cgit From 6bc9223f2f6efc6e5927715d456eaea317cfeeb0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 14:23:01 +0200 Subject: correctness proof of symbolic execution of one internal step --- mppa_k1c/lib/RTLpathSE_theory.v | 107 ++++++++++++++++++++++++++++++++-------- 1 file changed, 86 insertions(+), 21 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index ffaff8ac..10ab915e 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -70,58 +70,123 @@ with smem_eval (sge: sgenv) (sm: smem) (rs: regset) (m0: mem): option mem := Mem.storev chunk m a sv end. +(* Local symbolic internal states *) Record local_istate := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. -Record cond_istate := { the_cond: condition; cond_args: list_sval; cond_ist: local_istate; cond_ifso: node }. +Definition sem_local_istate (sge: sgenv) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := + st.(pre) sge rs0 m0 + /\ smem_eval sge st.(the_smem) rs0 m0 = Some m + /\ forall (r:reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = Some (rs#r). + +(* symbolic exit states *) +Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. (* symbolic internal state *) -Record s_istate := { the_pc: node; early_exit: list cond_istate; curr: local_istate + cond_istate }. +Record s_istate := { the_pc: node; early_exit: list exit_istate; curr: local_istate + exit_istate }. Definition get_curr (st:s_istate): local_istate := match st.(curr) with | inl x => x - | inr x => x.(cond_ist) + | inr x => x.(exit_ist) + end. + +Definition get_exit (st:s_istate): list exit_istate := + match st.(curr) with + | inl _ => st.(early_exit) + | inr cnd => cnd::(st.(early_exit)) end. -Definition sist (st: s_istate) (pc: node) (nxt: local_istate + cond_istate) := - match st.(curr) with - | inl _ => {| the_pc := pc; early_exit := st.(early_exit); curr:= nxt |} - | inr cnd => {| the_pc := pc; early_exit := cnd::st.(early_exit); curr:=nxt |} - end. +Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := + if (is.(continue)) + then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) + else exists cnd, List.In cnd (get_exit st) /\ sem_local_istate sge cnd.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem). -Definition sreg_set (st:local_istate) r (sv:sval) := - {| pre:=(fun _ _ _ => True ) ; (* TODO *) +Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := + forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. + + +Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := + {| pre:=(fun sge rs m => sval_eval sge (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge rs m)); the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; the_smem:= st.(the_smem)|}. -Definition smem_set (st:local_istate) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval) := - {| pre:=(fun _ _ _ => True ) ; (* TODO *) +Definition smem_set (st:local_istate) (sm:smem) := + {| pre:=(fun sge rs m => smem_eval sge st.(the_smem) rs m <> None /\ (st.(pre) sge rs m)); the_sreg:= st.(the_sreg); - the_smem:= Sstore st.(the_smem) chunk addr lsv srce |}. + the_smem:= sm |}. + +Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): option s_istate := + Some {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. (* one internal step of symbolic execution *) -Definition symb_istep (i: instruction) (st: s_istate) : option s_istate := +Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with - | Inop pc' => Some (sist st pc' (inl (get_curr st))) + | Inop pc' => + sist st pc' (inl (get_curr st)) | Iop op args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs) in - Some (sist st pc' (inl next)) + sist st pc' (inl next) | Iload chunk addr args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in - Some (sist st pc' (inl next)) + sist st pc' (inl next) | Istore chunk addr args src pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in - let next := smem_set prev chunk addr vargs (prev.(the_sreg) src) in - Some (sist st pc' (inl next)) + let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in + sist st pc' (inl next) | Icond cond args ifso ifnot => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in - let next := {| the_cond:=cond; cond_args:=vargs; cond_ist := prev; cond_ifso := ifso |} in - Some (sist st ifnot (inr next)) + let next := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in + sist st ifnot (inr next) | _ => None (* TODO jumptable ? *) end. + +Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m m0, eval_operation ge sp o args m = eval_operation ge sp o args m0. + +Lemma list_sval_eval_inj sge l rs0 m0 (sreg: reg -> sval) rs: + (forall r : reg, sval_eval sge (sreg r) rs0 m0 = Some (rs # r)) -> + list_sval_eval sge (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Proof. + intros H; induction l as [|r l]; simpl; auto. + inversion_SOME v. + inversion_SOME lv. + generalize (H r). + try_simplify_someHyps. +Qed. + +Lemma symb_istep_correct sge i st rs0 m0 rs m st' is: + sem_local_istate sge (get_curr st) rs0 m0 rs m -> + symb_istep i st = Some st' -> + step_istate sge.(the_ge) i sge.(the_sp) rs m = Some is -> + sem_istate sge st' rs0 m0 is. +Proof. + intros (PRE & MEM & REG). + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + + inversion_SOME v. try_simplify_someHyps. + intros EVALOP EVALMEM; intuition. + * exploit REG. try_simplify_someHyps. + * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl; auto. + erewrite list_sval_eval_inj; simpl; auto. + erewrite eval_operation_mem_irrel, EVALOP; auto. + + inversion_SOME a0. inversion_SOME v. try_simplify_someHyps. + intros LOAD EVALADD EVALMEM; intuition. + * exploit REG. try_simplify_someHyps. + * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + + inversion_SOME a0. inversion_SOME m'. try_simplify_someHyps. + intros STORE EVALADD EVALMEM; intuition. + * discriminate. + * erewrite list_sval_eval_inj; simpl; auto. + erewrite REG. + try_simplify_someHyps. + + inversion_SOME b. destruct b; try_simplify_someHyps. + intros EVALCOND EVALMEM; eexists; intuition eauto. +Qed. \ No newline at end of file -- cgit From 04a2ce441e639ddb817e138bbec76a66acb64297 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 14:25:18 +0200 Subject: m --- mppa_k1c/lib/RTLpath.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index a0365344..4e25b2b4 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -314,7 +314,7 @@ Proof. intros; destruct e; simpl; auto. Qed. -Local Ltac inversion_SOME x := +Ltac inversion_SOME x := try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]). Ltac simplify_someHyp := -- cgit From f6b26b154ca85303ab954381cacbfdef625d7dd8 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 15:59:39 +0200 Subject: fix dependency of Op on memory in the theory --- mppa_k1c/lib/RTLpathSE_theory.v | 31 +++++++++++++++++++------------ 1 file changed, 19 insertions(+), 12 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 10ab915e..300a2b0f 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -9,6 +9,15 @@ Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. +(* NOTE for the implementation of the symbolic execution. + It is important to remove dependency of Op on memory. + Here is an way to formalize this + +Parameter mem_irrel: operation -> bool. +Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_irrel o = true -> eval_operation ge sp o args m1 = eval_operation ge sp o args m2. + +*) + Record sgenv := { the_ge: RTL.genv; the_sp: val @@ -17,7 +26,7 @@ Record sgenv := { (* symbolic value *) Inductive sval := | Sinput (r: reg) - | Sop (op:operation) (lsv: list_sval) (* TODO: pb des fausses dependances sur la memoire, via comparaisons de pointeurs ? *) + | Sop (op:operation) (lsv: list_sval) (sm: smem) (* TODO for the implementation: add an additionnal case without dependency on sm*) | Sload (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) with list_sval := | Snil @@ -42,9 +51,10 @@ Local Open Scope option_monad_scope. Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs: regset) (m0: mem): option val := match sv with | Sinput r => Some (rs#r) - | Sop op l => + | Sop op l sm => SOME args <- list_sval_eval sge l rs m0 IN - eval_operation sge.(the_ge) sge.(the_sp) op args m0 + SOME m <- smem_eval sge sm rs m0 IN + eval_operation sge.(the_ge) sge.(the_sp) op args m | Sload sm chunk addr lsv => SOME args <- list_sval_eval sge lsv rs m0 IN SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN @@ -97,14 +107,13 @@ Definition get_exit (st:s_istate): list exit_istate := end. Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := - if (is.(continue)) - then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) - else exists cnd, List.In cnd (get_exit st) /\ sem_local_istate sge cnd.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem). + (if (is.(continue)) + then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) + else exists cnd, List.In cnd (get_exit st) /\ sem_local_istate sge cnd.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem)). Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. - Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := {| pre:=(fun sge rs m => sval_eval sge (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge rs m)); the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; @@ -126,7 +135,7 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := | Iop op args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in - let next := sreg_set prev dst (Sop op vargs) in + let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in sist st pc' (inl next) | Iload chunk addr args dst pc' => let prev := get_curr st in @@ -146,8 +155,6 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := | _ => None (* TODO jumptable ? *) end. -Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m m0, eval_operation ge sp o args m = eval_operation ge sp o args m0. - Lemma list_sval_eval_inj sge l rs0 m0 (sreg: reg -> sval) rs: (forall r : reg, sval_eval sge (sreg r) rs0 m0 = Some (rs # r)) -> list_sval_eval sge (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). @@ -173,7 +180,7 @@ Proof. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl; auto. erewrite list_sval_eval_inj; simpl; auto. - erewrite eval_operation_mem_irrel, EVALOP; auto. + rewrite EVALMEM, EVALOP; auto. + inversion_SOME a0. inversion_SOME v. try_simplify_someHyps. intros LOAD EVALADD EVALMEM; intuition. * exploit REG. try_simplify_someHyps. @@ -189,4 +196,4 @@ Proof. try_simplify_someHyps. + inversion_SOME b. destruct b; try_simplify_someHyps. intros EVALCOND EVALMEM; eexists; intuition eauto. -Qed. \ No newline at end of file +Qed. -- cgit From 375b11452d7babb5a7f8278394c6580c4051df35 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 16:05:56 +0200 Subject: invariant on nodes in symbolic execution --- mppa_k1c/lib/RTLpathSE_theory.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 300a2b0f..3b679e28 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -107,9 +107,9 @@ Definition get_exit (st:s_istate): list exit_istate := end. Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := - (if (is.(continue)) - then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) - else exists cnd, List.In cnd (get_exit st) /\ sem_local_istate sge cnd.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem)). + if (is.(continue)) + then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) + else exists ext, List.In ext (get_exit st) /\ sem_local_istate sge ext.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem) /\ ext.(cond_ifso) = is.(RTLpath.the_pc). Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. -- cgit From 45dc89b7dc37a36f032092e2b2816972d7e53d85 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 17 Oct 2019 17:21:25 +0200 Subject: start completness proof of symb exec of one internal step --- mppa_k1c/lib/RTLpathSE_theory.v | 81 +++++++++++++++++++++++++++++++++++++++-- 1 file changed, 77 insertions(+), 4 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 3b679e28..fc78d9dd 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -8,6 +8,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. +Require Coq.Logic.FunctionalExtensionality. (* NOTE for the implementation of the symbolic execution. It is important to remove dependency of Op on memory. @@ -111,8 +112,10 @@ Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: i then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) else exists ext, List.In ext (get_exit st) /\ sem_local_istate sge ext.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem) /\ ext.(cond_ifso) = is.(RTLpath.the_pc). +(* Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. +*) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := {| pre:=(fun sge rs m => sval_eval sge (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge rs m)); @@ -174,26 +177,96 @@ Lemma symb_istep_correct sge i st rs0 m0 rs m st' is: Proof. intros (PRE & MEM & REG). destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. - + inversion_SOME v. try_simplify_someHyps. + + (* Op *) + inversion_SOME v. try_simplify_someHyps. intros EVALOP EVALMEM; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl; auto. erewrite list_sval_eval_inj; simpl; auto. rewrite EVALMEM, EVALOP; auto. - + inversion_SOME a0. inversion_SOME v. try_simplify_someHyps. + + (* LOAD *) + inversion_SOME a0. inversion_SOME v. try_simplify_someHyps. intros LOAD EVALADD EVALMEM; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. - + inversion_SOME a0. inversion_SOME m'. try_simplify_someHyps. + + (* STORE *) + inversion_SOME a0. inversion_SOME m'. try_simplify_someHyps. intros STORE EVALADD EVALMEM; intuition. * discriminate. * erewrite list_sval_eval_inj; simpl; auto. erewrite REG. try_simplify_someHyps. - + inversion_SOME b. destruct b; try_simplify_someHyps. + + (* COND *) + inversion_SOME b. destruct b; try_simplify_someHyps. intros EVALCOND EVALMEM; eexists; intuition eauto. Qed. + +Definition istate_eq ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + +Definition istate_eq_option ist1 oist := + exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. + +Lemma symb_istep_complete sge i st rs0 m0 rs m st' is: + is.(continue) = true -> + sem_local_istate sge (get_curr st) rs0 m0 rs m -> + sem_istate sge st' rs0 m0 is -> + symb_istep i st = Some st' -> + istate_eq_option is (step_istate sge.(the_ge) i sge.(the_sp) rs m). +Proof. + unfold sem_istate, istate_eq_option, istate_eq; destruct is as [c pc rs1 m1]; simpl. + intro; subst. + intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC); subst. + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. + + (* Nop *) + intros MEM MEM1; eexists; intuition eauto; simpl. + - generalize (REG1 r). unfold get_curr at 1; simpl. + rewrite REG; try congruence. + - unfold get_curr at 1 in MEM1; simpl in MEM1; congruence. + + (* Op *) + intros MEM. + generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. + erewrite list_sval_eval_inj; simpl; auto. + rewrite MEM. + try_simplify_someHyps. + eexists; intuition eauto; simpl. + generalize (REG1 r0). + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. + subst; rewrite Regmap.gss. auto. + + (* LOAD *) + intros MEM. + generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. + erewrite list_sval_eval_inj; simpl; auto. + rewrite MEM. + inversion_SOME addr. + intros EVAL LOAD. rewrite EVAL, LOAD; simpl. + eexists; intuition eauto; simpl. + generalize (REG1 r0). + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. + subst; rewrite Regmap.gss. auto. + + (* STORE *) + intros MEM. + generalize (REG r); destruct (Pos.eq_dec r r); try congruence; simpl. + intros Hr; rewrite Hr; simpl. + erewrite list_sval_eval_inj; simpl; auto. + inversion_SOME addr. + intros EVAL STORE. rewrite EVAL, STORE; simpl. + eexists; intuition eauto; simpl. + generalize (REG r0); rewrite REG1. congruence. + + (* COND *) + intros MEM MEM1. + unfold get_curr at 1 in REG1; + unfold get_curr at 1 in PRE1; + unfold get_curr at 1 in MEM1; + simpl in * |-. + (* TODO: A REVOIR. Il faut exprimer le fait que toutes les conditions sur les sorties intermediaires ont repondu "false" + depuis rs0, m0 + *) +Admitted. -- cgit From 950b63bc896fba666fc7106fbedfc2346dd42696 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 18 Oct 2019 08:46:21 +0200 Subject: completness proof of symb exec of one internal step --- mppa_k1c/lib/RTLpathSE_theory.v | 107 ++++++++++++++++++++++++++-------------- 1 file changed, 70 insertions(+), 37 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index fc78d9dd..8063672c 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -49,39 +49,39 @@ Fixpoint list_val_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. -Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs: regset) (m0: mem): option val := +Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs0: regset) (m0: mem): option val := match sv with - | Sinput r => Some (rs#r) + | Sinput r => Some (rs0#r) | Sop op l sm => - SOME args <- list_sval_eval sge l rs m0 IN - SOME m <- smem_eval sge sm rs m0 IN + SOME args <- list_sval_eval sge l rs0 m0 IN + SOME m <- smem_eval sge sm rs0 m0 IN eval_operation sge.(the_ge) sge.(the_sp) op args m | Sload sm chunk addr lsv => - SOME args <- list_sval_eval sge lsv rs m0 IN + SOME args <- list_sval_eval sge lsv rs0 m0 IN SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN - SOME m <- smem_eval sge sm rs m0 IN + SOME m <- smem_eval sge sm rs0 m0 IN Mem.loadv chunk m a end -with list_sval_eval (sge: sgenv) (lsv: list_sval) (rs: regset) (m0: mem): option (list val) := +with list_sval_eval (sge: sgenv) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := match lsv with | Snil => Some nil | Scons sv lsv' => - SOME v <- sval_eval sge sv rs m0 IN - SOME lv <- list_sval_eval sge lsv' rs m0 IN + SOME v <- sval_eval sge sv rs0 m0 IN + SOME lv <- list_sval_eval sge lsv' rs0 m0 IN Some (v::lv) end -with smem_eval (sge: sgenv) (sm: smem) (rs: regset) (m0: mem): option mem := +with smem_eval (sge: sgenv) (sm: smem) (rs0: regset) (m0: mem): option mem := match sm with | Sinit => Some m0 | Sstore sm chunk addr lsv srce => - SOME args <- list_sval_eval sge lsv rs m0 IN + SOME args <- list_sval_eval sge lsv rs0 m0 IN SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN - SOME m <- smem_eval sge sm rs m0 IN - SOME sv <- sval_eval sge srce rs m0 IN + SOME m <- smem_eval sge sm rs0 m0 IN + SOME sv <- sval_eval sge srce rs0 m0 IN Mem.storev chunk m a sv end. -(* Local symbolic internal states *) +(* Syntax and Semantics of local symbolic internal states *) Record local_istate := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. Definition sem_local_istate (sge: sgenv) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := @@ -89,30 +89,58 @@ Definition sem_local_istate (sge: sgenv) (st: local_istate) (rs0: regset) (m0: m /\ smem_eval sge st.(the_smem) rs0 m0 = Some m /\ forall (r:reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = Some (rs#r). -(* symbolic exit states *) +(* Syntax and semantics of symbolic exit states *) Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. -(* symbolic internal state *) +Definition seval_condition sge (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := + SOME args <- list_sval_eval sge lsv rs0 m0 IN + SOME m <- smem_eval sge sm rs0 m0 IN + eval_condition cond args m. + +Fixpoint have_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := + match lx with + | nil => True + | ext::lx => + seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false + /\ have_not_yet_exit sge lx rs0 m0 + end. + +Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := + SEE_intro ext lx': + is_tail (ext::lx') lx -> + have_not_yet_exit sge lx' rs0 m0 -> + seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> + sem_local_istate sge ext.(exit_ist) rs0 m0 rs m -> + ext.(cond_ifso) = pc -> + sem_early_exit sge lx rs0 m0 rs m pc. + + +(* Syntax and Semantics of symbolic internal state *) Record s_istate := { the_pc: node; early_exit: list exit_istate; curr: local_istate + exit_istate }. Definition get_curr (st:s_istate): local_istate := - match st.(curr) with + let (_,_,c) := st in + match c with | inl x => x | inr x => x.(exit_ist) end. Definition get_exit (st:s_istate): list exit_istate := - match st.(curr) with + let (_,_,c) := st in + match c with | inl _ => st.(early_exit) | inr cnd => cnd::(st.(early_exit)) end. Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(continue)) - then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) - else exists ext, List.In ext (get_exit st) /\ sem_local_istate sge ext.(exit_ist) rs0 m0 is.(the_rs) is.(the_mem) /\ ext.(cond_ifso) = is.(RTLpath.the_pc). + then + sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) + /\ st.(the_pc) = is.(RTLpath.the_pc) + /\ have_not_yet_exit sge (get_exit st) rs0 m0 + else sem_early_exit sge (get_exit st) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). -(* +(* The equality on istate below is not general enough... See [istate_eq] below... Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. *) @@ -171,11 +199,12 @@ Qed. Lemma symb_istep_correct sge i st rs0 m0 rs m st' is: sem_local_istate sge (get_curr st) rs0 m0 rs m -> + have_not_yet_exit sge (get_exit st) rs0 m0 -> symb_istep i st = Some st' -> step_istate sge.(the_ge) i sge.(the_sp) rs m = Some is -> sem_istate sge st' rs0 m0 is. Proof. - intros (PRE & MEM & REG). + intros (PRE & MEM & REG) NYE. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + (* Op *) inversion_SOME v. try_simplify_someHyps. @@ -201,8 +230,16 @@ Proof. erewrite REG. try_simplify_someHyps. + (* COND *) + Local Hint Resolve is_tail_refl. + Local Hint Unfold sem_local_istate. inversion_SOME b. destruct b; try_simplify_someHyps. - intros EVALCOND EVALMEM; eexists; intuition eauto. + intros EVALCOND EVALMEM; econstructor; eauto; simpl. + * unfold seval_condition. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + * intuition. unfold seval_condition. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. Qed. Definition istate_eq ist1 ist2 := @@ -223,13 +260,11 @@ Lemma symb_istep_complete sge i st rs0 m0 rs m st' is: Proof. unfold sem_istate, istate_eq_option, istate_eq; destruct is as [c pc rs1 m1]; simpl. intro; subst. - intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC); subst. + intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC & NYE); subst. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. + (* Nop *) - intros MEM MEM1; eexists; intuition eauto; simpl. - - generalize (REG1 r). unfold get_curr at 1; simpl. - rewrite REG; try congruence. - - unfold get_curr at 1 in MEM1; simpl in MEM1; congruence. + intros MEM; eexists; intuition eauto; simpl. + generalize (REG1 r); rewrite REG; try congruence. + (* Op *) intros MEM. generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. @@ -261,12 +296,10 @@ Proof. eexists; intuition eauto; simpl. generalize (REG r0); rewrite REG1. congruence. + (* COND *) - intros MEM MEM1. - unfold get_curr at 1 in REG1; - unfold get_curr at 1 in PRE1; - unfold get_curr at 1 in MEM1; - simpl in * |-. - (* TODO: A REVOIR. Il faut exprimer le fait que toutes les conditions sur les sorties intermediaires ont repondu "false" - depuis rs0, m0 - *) -Admitted. + intros MEM. destruct NYE as (COND & NYE). + unfold seval_condition in COND. + erewrite list_sval_eval_inj in COND; simpl; auto. + rewrite MEM in COND. + rewrite COND. eexists; intuition eauto. simpl. + generalize (REG r); rewrite REG1. congruence. +Qed. -- cgit From 2cb0bbe386e52d017e26eb4a871bd40821f5e9ef Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 18 Oct 2019 16:53:14 +0200 Subject: Correctness of symbolic execution of internal steps --- mppa_k1c/lib/RTLpathSE_theory.v | 81 ++++++++++++++++++++++++++++++++++++++++- 1 file changed, 79 insertions(+), 2 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 8063672c..e1ab2982 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -19,6 +19,9 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ *) +(** * Syntax and semantics of symbolic values *) + + Record sgenv := { the_ge: RTL.genv; the_sp: val @@ -115,7 +118,7 @@ Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: m sem_early_exit sge lx rs0 m0 rs m pc. -(* Syntax and Semantics of symbolic internal state *) +(** * Syntax and Semantics of symbolic internal state *) Record s_istate := { the_pc: node; early_exit: list exit_istate; curr: local_istate + exit_istate }. Definition get_curr (st:s_istate): local_istate := @@ -145,6 +148,8 @@ Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. *) +(** * Symbolic execution of one internal step *) + Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := {| pre:=(fun sge rs m => sval_eval sge (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge rs m)); the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; @@ -158,7 +163,6 @@ Definition smem_set (st:local_istate) (sm:smem) := Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): option s_istate := Some {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. -(* one internal step of symbolic execution *) Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with | Inop pc' => @@ -242,6 +246,26 @@ Proof. try_simplify_someHyps. Qed. +Lemma symb_istep_correct_None sge i st rs0 m0 rs m: + sem_local_istate sge (get_curr st) rs0 m0 rs m -> + symb_istep i st = None -> + step_istate sge.(the_ge) i sge.(the_sp) rs m = None. +Proof. + intros (PRE & MEM & REG). + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. +Qed. + +Lemma symb_istep_preserv_early_exits i sge st rs0 m0 rs m pc st': + sem_early_exit sge (get_exit st) rs0 m0 rs m pc -> + symb_istep i st = Some st' -> + sem_early_exit sge (get_exit st') rs0 m0 rs m pc. +Proof. + intros H. + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + destruct H as [ext lx TAIL NYE COND INV PC]. + econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. +Qed. + Definition istate_eq ist1 ist2 := ist1.(continue) = ist2.(continue) /\ ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ @@ -303,3 +327,56 @@ Proof. rewrite COND. eexists; intuition eauto. simpl. generalize (REG r); rewrite REG1. congruence. Qed. + + + +(** * Symbolic execution of the internal steps of a path *) +Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := + match path with + | O => sist st st.(the_pc) (inl (get_curr st)) + | S p => + SOME i <- (fn_code f)!(st.(the_pc)) IN + match symb_istep i st with + | Some st1 => symb_isteps p f st1 + | None => sist st st.(the_pc) (inl (get_curr st)) + end + end. + +Lemma symb_isteps_correct_false sge path f rs0 m0 st' is: + is.(continue)=false -> + forall st, sem_istate sge st rs0 m0 is -> + symb_isteps path f st = Some st' -> + sem_istate sge st' rs0 m0 is. +Proof. + Local Hint Resolve symb_istep_preserv_early_exits. + intros CONT; unfold sem_istate; rewrite CONT. + induction path; simpl. + + unfold sist; try_simplify_someHyps. + + intros st; inversion_SOME i. + destruct (symb_istep i st) as [st0|] eqn:STEP; eauto. + unfold sist; try_simplify_someHyps. +Qed. + +Lemma symb_isteps_correct_true sge path f rs0 m0 st' is': forall st is, + is.(continue)=true -> + sem_istate sge st rs0 m0 is -> + symb_isteps path f st = Some st' -> + path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc) = Some is' -> + sem_istate sge st' rs0 m0 is'. +Proof. + Local Hint Resolve symb_istep_correct symb_isteps_correct_false. + induction path; simpl. + + intros st is; + unfold sem_istate, sist in * |- *; + try_simplify_someHyps. + destruct is; simpl in * |- *; subst; intuition auto. + + intros st is CONT; unfold sem_istate at 1; rewrite CONT. + intros (LOCAL & PC & NYE). + rewrite <- PC. + inversion_SOME i; try_simplify_someHyps; intro Hi. + inversion_SOME is0. + destruct (symb_istep i st) eqn:STEP. 2: { erewrite symb_istep_correct_None; eauto. congruence. } + destruct (continue is0) eqn:CONT0; eauto. + (* continue is0 = false -> EARLY EXIT *) + try_simplify_someHyps; eauto. +Qed. -- cgit From 04b1bf870a3aacf803dfdcd8e5aed8a2e6ba39bd Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 18 Oct 2019 17:04:23 +0200 Subject: improved SE of internal steps ? --- mppa_k1c/lib/RTLpathSE_theory.v | 30 +++++++++++++----------------- 1 file changed, 13 insertions(+), 17 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index e1ab2982..20cf3689 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -160,33 +160,33 @@ Definition smem_set (st:local_istate) (sm:smem) := the_sreg:= st.(the_sreg); the_smem:= sm |}. -Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): option s_istate := - Some {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. +Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): s_istate := + {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with | Inop pc' => - sist st pc' (inl (get_curr st)) + Some (sist st pc' (inl (get_curr st))) | Iop op args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in - sist st pc' (inl next) + Some (sist st pc' (inl next)) | Iload chunk addr args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in - sist st pc' (inl next) + Some (sist st pc' (inl next)) | Istore chunk addr args src pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in - sist st pc' (inl next) + Some (sist st pc' (inl next)) | Icond cond args ifso ifnot => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in - sist st ifnot (inr next) + Some (sist st ifnot (inr next)) | _ => None (* TODO jumptable ? *) end. @@ -331,22 +331,20 @@ Qed. (** * Symbolic execution of the internal steps of a path *) -Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := +Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): s_istate := match path with | O => sist st st.(the_pc) (inl (get_curr st)) | S p => - SOME i <- (fn_code f)!(st.(the_pc)) IN - match symb_istep i st with + match (SOME i <- (fn_code f)!(st.(the_pc)) IN symb_istep i st) with | Some st1 => symb_isteps p f st1 | None => sist st st.(the_pc) (inl (get_curr st)) end end. -Lemma symb_isteps_correct_false sge path f rs0 m0 st' is: +Lemma symb_isteps_correct_false sge path f rs0 m0 is: is.(continue)=false -> forall st, sem_istate sge st rs0 m0 is -> - symb_isteps path f st = Some st' -> - sem_istate sge st' rs0 m0 is. + sem_istate sge (symb_isteps path f st) rs0 m0 is. Proof. Local Hint Resolve symb_istep_preserv_early_exits. intros CONT; unfold sem_istate; rewrite CONT. @@ -354,15 +352,13 @@ Proof. + unfold sist; try_simplify_someHyps. + intros st; inversion_SOME i. destruct (symb_istep i st) as [st0|] eqn:STEP; eauto. - unfold sist; try_simplify_someHyps. Qed. -Lemma symb_isteps_correct_true sge path f rs0 m0 st' is': forall st is, +Lemma symb_isteps_correct_true sge path f rs0 m0 is': forall st is, is.(continue)=true -> sem_istate sge st rs0 m0 is -> - symb_isteps path f st = Some st' -> path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc) = Some is' -> - sem_istate sge st' rs0 m0 is'. + sem_istate sge (symb_isteps path f st) rs0 m0 is'. Proof. Local Hint Resolve symb_istep_correct symb_isteps_correct_false. induction path; simpl. -- cgit From 4ea28e089a0000d24725bfaac6e1c12955a1aa14 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 18 Oct 2019 17:51:21 +0200 Subject: Revert "improved SE of internal steps ?" This reverts commit 04b1bf870a3aacf803dfdcd8e5aed8a2e6ba39bd. --- mppa_k1c/lib/RTLpathSE_theory.v | 30 +++++++++++++++++------------- 1 file changed, 17 insertions(+), 13 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 20cf3689..e1ab2982 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -160,33 +160,33 @@ Definition smem_set (st:local_istate) (sm:smem) := the_sreg:= st.(the_sreg); the_smem:= sm |}. -Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): s_istate := - {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. +Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): option s_istate := + Some {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with | Inop pc' => - Some (sist st pc' (inl (get_curr st))) + sist st pc' (inl (get_curr st)) | Iop op args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in - Some (sist st pc' (inl next)) + sist st pc' (inl next) | Iload chunk addr args dst pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in - Some (sist st pc' (inl next)) + sist st pc' (inl next) | Istore chunk addr args src pc' => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in - Some (sist st pc' (inl next)) + sist st pc' (inl next) | Icond cond args ifso ifnot => let prev := get_curr st in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in - Some (sist st ifnot (inr next)) + sist st ifnot (inr next) | _ => None (* TODO jumptable ? *) end. @@ -331,20 +331,22 @@ Qed. (** * Symbolic execution of the internal steps of a path *) -Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): s_istate := +Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := match path with | O => sist st st.(the_pc) (inl (get_curr st)) | S p => - match (SOME i <- (fn_code f)!(st.(the_pc)) IN symb_istep i st) with + SOME i <- (fn_code f)!(st.(the_pc)) IN + match symb_istep i st with | Some st1 => symb_isteps p f st1 | None => sist st st.(the_pc) (inl (get_curr st)) end end. -Lemma symb_isteps_correct_false sge path f rs0 m0 is: +Lemma symb_isteps_correct_false sge path f rs0 m0 st' is: is.(continue)=false -> forall st, sem_istate sge st rs0 m0 is -> - sem_istate sge (symb_isteps path f st) rs0 m0 is. + symb_isteps path f st = Some st' -> + sem_istate sge st' rs0 m0 is. Proof. Local Hint Resolve symb_istep_preserv_early_exits. intros CONT; unfold sem_istate; rewrite CONT. @@ -352,13 +354,15 @@ Proof. + unfold sist; try_simplify_someHyps. + intros st; inversion_SOME i. destruct (symb_istep i st) as [st0|] eqn:STEP; eauto. + unfold sist; try_simplify_someHyps. Qed. -Lemma symb_isteps_correct_true sge path f rs0 m0 is': forall st is, +Lemma symb_isteps_correct_true sge path f rs0 m0 st' is': forall st is, is.(continue)=true -> sem_istate sge st rs0 m0 is -> + symb_isteps path f st = Some st' -> path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc) = Some is' -> - sem_istate sge (symb_isteps path f st) rs0 m0 is'. + sem_istate sge st' rs0 m0 is'. Proof. Local Hint Resolve symb_istep_correct symb_isteps_correct_false. induction path; simpl. -- cgit From 64dd4fb97572d6114f8a0a86db7b8e8013987dda Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 18 Oct 2019 18:48:45 +0200 Subject: Attempt on completeness of internal steps (but wrong goal). --- mppa_k1c/lib/RTLpathSE_theory.v | 61 +++++++++++++++++++++++++++++++---------- 1 file changed, 47 insertions(+), 14 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index e1ab2982..c04da5c9 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -100,13 +100,8 @@ Definition seval_condition sge (cond: condition) (lsv: list_sval) (sm: smem) rs0 SOME m <- smem_eval sge sm rs0 m0 IN eval_condition cond args m. -Fixpoint have_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := - match lx with - | nil => True - | ext::lx => - seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false - /\ have_not_yet_exit sge lx rs0 m0 - end. +Definition have_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := + forall ext, List.In ext lx -> seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := SEE_intro ext lx': @@ -241,7 +236,9 @@ Proof. * unfold seval_condition. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. - * intuition. unfold seval_condition. + * intuition. unfold have_not_yet_exit in * |- *. simpl. + intuition. subst. simpl. + unfold seval_condition. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. Qed. @@ -287,7 +284,7 @@ Proof. intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC & NYE); subst. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. + (* Nop *) - intros MEM; eexists; intuition eauto; simpl. + intros; eexists; intuition eauto; simpl. generalize (REG1 r); rewrite REG; try congruence. + (* Op *) intros MEM. @@ -320,11 +317,11 @@ Proof. eexists; intuition eauto; simpl. generalize (REG r0); rewrite REG1. congruence. + (* COND *) - intros MEM. destruct NYE as (COND & NYE). - unfold seval_condition in COND. - erewrite list_sval_eval_inj in COND; simpl; auto. - rewrite MEM in COND. - rewrite COND. eexists; intuition eauto. simpl. + intros MEM. + exploit NYE. { simpl; eauto. } + simpl; unfold seval_condition; erewrite list_sval_eval_inj; simpl; auto. + rewrite MEM. + intro COND; rewrite COND. eexists; intuition eauto. simpl. generalize (REG r); rewrite REG1. congruence. Qed. @@ -380,3 +377,39 @@ Proof. (* continue is0 = false -> EARLY EXIT *) try_simplify_someHyps; eauto. Qed. + +Lemma symb_isteps_complete sge path f rs0 m0 st' is': forall st is, + is.(continue)=true -> + sem_istate sge st rs0 m0 is -> + symb_isteps path f st = Some st' -> + sem_istate sge st' rs0 m0 is' -> + istate_eq_option is' (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)). +Proof. + induction path. + + unfold sem_istate, istate_eq_option, istate_eq; simpl. + intros st is CONT; rewrite CONT. + intros ((PRE & EVAL & REG) & H & NYE). + unfold sist; intros Hst'; inversion_clear Hst'. + destruct (continue is') eqn: CONT'. + * simpl; intros ((PRE1 & EVAL1 & REG1) & H1). + eexists; intuition eauto; simpl; try congruence. + generalize (REG1 r); rewrite REG. congruence. + * intros [ext lx TAIL NYE1 COND INV PC]. + exploit NYE. { eapply is_tail_in; eauto. } + congruence. + + simpl. unfold sem_istate at 1. + intros st is CONT; rewrite CONT. + intros (LOCAL & PC & NYE). + rewrite <- PC. + inversion_SOME i; try_simplify_someHyps; intro Hi. + destruct (symb_istep i st) eqn:STEP. + * inversion_SOME is0. + - destruct (continue is0) eqn:CONT0; eauto. + (* continue is0 = false -> EARLY EXIT *) + intros ISTEP Hst'. + admit. + - (* absurd case ?*) admit. + * unfold sist; intros Hst'; inversion_clear Hst'. + destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. + (* TODO: unprovable. an assumption is missing on default_succ ? *) +Abort. -- cgit From 576049871a3403e27529f0beeee10a677b47cce8 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 19 Oct 2019 07:52:09 +0200 Subject: m --- mppa_k1c/lib/RTLpathSE_theory.v | 28 ++++++++++++---------------- 1 file changed, 12 insertions(+), 16 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index c04da5c9..07ca9a0f 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -333,10 +333,8 @@ Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := | O => sist st st.(the_pc) (inl (get_curr st)) | S p => SOME i <- (fn_code f)!(st.(the_pc)) IN - match symb_istep i st with - | Some st1 => symb_isteps p f st1 - | None => sist st st.(the_pc) (inl (get_curr st)) - end + SOME st1 <- symb_istep i st IN + symb_isteps p f st1 end. Lemma symb_isteps_correct_false sge path f rs0 m0 st' is: @@ -350,8 +348,7 @@ Proof. induction path; simpl. + unfold sist; try_simplify_someHyps. + intros st; inversion_SOME i. - destruct (symb_istep i st) as [st0|] eqn:STEP; eauto. - unfold sist; try_simplify_someHyps. + inversion_SOME st1; eauto. Qed. Lemma symb_isteps_correct_true sge path f rs0 m0 st' is': forall st is, @@ -371,8 +368,8 @@ Proof. intros (LOCAL & PC & NYE). rewrite <- PC. inversion_SOME i; try_simplify_someHyps; intro Hi. + inversion_SOME st1. inversion_SOME is0. - destruct (symb_istep i st) eqn:STEP. 2: { erewrite symb_istep_correct_None; eauto. congruence. } destruct (continue is0) eqn:CONT0; eauto. (* continue is0 = false -> EARLY EXIT *) try_simplify_someHyps; eauto. @@ -402,14 +399,13 @@ Proof. intros (LOCAL & PC & NYE). rewrite <- PC. inversion_SOME i; try_simplify_someHyps; intro Hi. - destruct (symb_istep i st) eqn:STEP. - * inversion_SOME is0. - - destruct (continue is0) eqn:CONT0; eauto. - (* continue is0 = false -> EARLY EXIT *) - intros ISTEP Hst'. - admit. - - (* absurd case ?*) admit. - * unfold sist; intros Hst'; inversion_clear Hst'. - destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. + inversion_SOME st1. + inversion_SOME is0. + * destruct (continue is0) eqn:CONT0; eauto. + (* continue is0 = false -> EARLY EXIT *) + intros ISTEP Hst1 Hst'. + admit. + * destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. (* TODO: unprovable. an assumption is missing on default_succ ? *) Abort. + -- cgit From ae5122f8d0c1362e690422e85a00a9dad2d18971 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 20 Oct 2019 09:34:10 +0200 Subject: m --- mppa_k1c/lib/RTLpathSE_theory.v | 52 ++++++++++++++++++++++++++++++++++------- 1 file changed, 44 insertions(+), 8 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 07ca9a0f..26fd8cae 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -330,7 +330,7 @@ Qed. (** * Symbolic execution of the internal steps of a path *) Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := match path with - | O => sist st st.(the_pc) (inl (get_curr st)) + | O => sist st st.(the_pc) (inl (get_curr st)) (* NOT CLEAR: returning [Some st] allows better symmetry => right_assoc below *) | S p => SOME i <- (fn_code f)!(st.(the_pc)) IN SOME st1 <- symb_istep i st IN @@ -375,18 +375,51 @@ Proof. try_simplify_someHyps; eauto. Qed. +(* ISSUE with path=0 + +Lemma symb_isteps_right_assoc_decompose f path: forall st st', + symb_isteps (S path) f st = Some st' -> + exists st0, symb_isteps path f st = Some st0 /\ symb_isteps 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 symb_isteps_right_assoc_compose f path: forall st st0 st', + symb_isteps path f st = Some st0 -> + symb_isteps 1%nat f st0 = Some st' -> + symb_isteps (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. +*) + Lemma symb_isteps_complete sge path f rs0 m0 st' is': forall st is, is.(continue)=true -> sem_istate sge st rs0 m0 is -> symb_isteps path f st = Some st' -> sem_istate sge st' rs0 m0 is' -> + (* nth_default_succ (fn_code f) path st.(the_pc) <> None -> *) istate_eq_option is' (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)). Proof. induction path. + unfold sem_istate, istate_eq_option, istate_eq; simpl. intros st is CONT; rewrite CONT. intros ((PRE & EVAL & REG) & H & NYE). - unfold sist; intros Hst'; inversion_clear Hst'. + (* unfold sist; *) intros Hst'; inversion_clear Hst'. destruct (continue is') eqn: CONT'. * simpl; intros ((PRE1 & EVAL1 & REG1) & H1). eexists; intuition eauto; simpl; try congruence. @@ -398,14 +431,17 @@ Proof. intros st is CONT; rewrite CONT. intros (LOCAL & PC & NYE). rewrite <- PC. - inversion_SOME i; try_simplify_someHyps; intro Hi. - inversion_SOME st1. + inversion_SOME i; intro Hi. rewrite Hi. + inversion_SOME st1. intros Hst1 Hst' SEMis' (* SUCC *) . inversion_SOME is0. * destruct (continue is0) eqn:CONT0; eauto. - (* continue is0 = false -> EARLY EXIT *) - intros ISTEP Hst1 Hst'. - admit. + (*- (* continue is0 = true *) + intros; eapply IHpath; eauto. + destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. + - *) (* continue is0 = false -> EARLY EXIT *) + admit. * destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. + - generalize H; inversion_SOME v; try congruence. (* TODO: unprovable. an assumption is missing on default_succ ? *) Abort. - +*) -- cgit From 1ea60bc655d406575c181965d1fcae9cd587d701 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sun, 20 Oct 2019 14:57:19 +0200 Subject: define sem_option2_istate --- mppa_k1c/lib/RTLpathSE_theory.v | 66 ++++++++++++++++++++++++++++++++--------- 1 file changed, 52 insertions(+), 14 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 26fd8cae..4ba501db 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -92,6 +92,11 @@ Definition sem_local_istate (sge: sgenv) (st: local_istate) (rs0: regset) (m0: m /\ smem_eval sge st.(the_smem) rs0 m0 = Some m /\ forall (r:reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = Some (rs#r). +Definition has_aborted_on_basic (sge: sgenv) (st: local_istate) (rs0: regset) (m0: mem): Prop := + ~(st.(pre) sge rs0 m0) + \/ smem_eval sge st.(the_smem) rs0 m0 = None + \/ exists (r: reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = None. + (* Syntax and semantics of symbolic exit states *) Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. @@ -100,13 +105,16 @@ Definition seval_condition sge (cond: condition) (lsv: list_sval) (sm: smem) rs0 SOME m <- smem_eval sge sm rs0 m0 IN eval_condition cond args m. -Definition have_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := +Definition has_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := forall ext, List.In ext lx -> seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. +Definition has_aborted_on_test sge (lx: list exit_istate) rs0 m0: Prop := + exists ext, List.In ext lx /\ seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. + Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := SEE_intro ext lx': is_tail (ext::lx') lx -> - have_not_yet_exit sge lx' rs0 m0 -> + has_not_yet_exit sge lx' rs0 m0 -> seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> sem_local_istate sge ext.(exit_ist) rs0 m0 rs m -> ext.(cond_ifso) = pc -> @@ -135,14 +143,31 @@ Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: i then sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) - /\ have_not_yet_exit sge (get_exit st) rs0 m0 + /\ has_not_yet_exit sge (get_exit st) rs0 m0 else sem_early_exit sge (get_exit st) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). -(* The equality on istate below is not general enough... See [istate_eq] below... -Definition sem_option_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (ois: option istate): Prop := - forall is, ois = Some is <-> sem_istate sge st rs0 m0 is. +Definition has_aborted (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem): Prop := + has_aborted_on_basic sge (get_curr st) rs0 m0 + \/ has_aborted_on_test sge (get_exit st) rs0 m0. + + +Definition sem_option_istate sge (st: s_istate) rs0 m0 (ois: option istate): Prop := + match ois with + | Some is => sem_istate sge st rs0 m0 is + | None => has_aborted sge st rs0 m0 + end. + +Definition sem_option2_istate sge (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := + match ost with + | Some st => sem_option_istate sge st rs0 m0 ois + | None => ois=None + end. + + +(** TODO: Lemmes de correction/complétudes ci-dessous à généraliser en utilisant sem_option2_istate !!! *) + (** * Symbolic execution of one internal step *) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := @@ -198,7 +223,7 @@ Qed. Lemma symb_istep_correct sge i st rs0 m0 rs m st' is: sem_local_istate sge (get_curr st) rs0 m0 rs m -> - have_not_yet_exit sge (get_exit st) rs0 m0 -> + has_not_yet_exit sge (get_exit st) rs0 m0 -> symb_istep i st = Some st' -> step_istate sge.(the_ge) i sge.(the_sp) rs m = Some is -> sem_istate sge st' rs0 m0 is. @@ -236,7 +261,7 @@ Proof. * unfold seval_condition. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. - * intuition. unfold have_not_yet_exit in * |- *. simpl. + * intuition. unfold has_not_yet_exit in * |- *. simpl. intuition. subst. simpl. unfold seval_condition. erewrite list_sval_eval_inj; simpl; auto. @@ -263,6 +288,8 @@ Proof. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. + +(* TODO: remove [istat_eq] ! This is useless (use sem_option2_istate instead !) Definition istate_eq ist1 ist2 := ist1.(continue) = ist2.(continue) /\ ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ @@ -324,13 +351,13 @@ Proof. intro COND; rewrite COND. eexists; intuition eauto. simpl. generalize (REG r); rewrite REG1. congruence. Qed. - +*) (** * Symbolic execution of the internal steps of a path *) Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := match path with - | O => sist st st.(the_pc) (inl (get_curr st)) (* NOT CLEAR: returning [Some st] allows better symmetry => right_assoc below *) + | O => sist st st.(the_pc) (inl (get_curr st)) | S p => SOME i <- (fn_code f)!(st.(the_pc)) IN SOME st1 <- symb_istep i st IN @@ -375,7 +402,13 @@ Proof. try_simplify_someHyps; eauto. Qed. -(* ISSUE with path=0 +(* + +(* Remark: the properties are not true, because of path=0! + + if we use [symb_isteps 0%nat f st = Some st], this works... + +*) Lemma symb_isteps_right_assoc_decompose f path: forall st st', symb_isteps (S path) f st = Some st' -> @@ -407,7 +440,13 @@ Proof. Qed. *) -Lemma symb_isteps_complete sge path f rs0 m0 st' is': forall st is, + +(* PROPRIETE FAUSSE: + +il faut utiliser sem_option2_istate + + +Lemma symb_isteps_complete sge path f rs0 m0 st': forall st, is.(continue)=true -> sem_istate sge st rs0 m0 is -> symb_isteps path f st = Some st' -> @@ -442,6 +481,5 @@ Proof. admit. * destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. - generalize H; inversion_SOME v; try congruence. - (* TODO: unprovable. an assumption is missing on default_succ ? *) Abort. -*) +*) \ No newline at end of file -- cgit From 6e52f4acc4fbfaf0d24e4125ded4cad208a9926e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 21 Oct 2019 12:09:20 +0200 Subject: progress on symb_isteps... --- mppa_k1c/lib/RTLpathSE_theory.v | 229 ++++++++++++++++++---------------------- 1 file changed, 102 insertions(+), 127 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 4ba501db..cb7a6276 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -221,51 +221,95 @@ Proof. try_simplify_someHyps. Qed. -Lemma symb_istep_correct sge i st rs0 m0 rs m st' is: +Lemma symb_istep_preserv_early_exits i sge st rs0 m0 rs m pc st': + sem_early_exit sge (get_exit st) rs0 m0 rs m pc -> + symb_istep i st = Some st' -> + sem_early_exit sge (get_exit st') rs0 m0 rs m pc. +Proof. + intros H. + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + destruct H as [ext lx TAIL NYE COND INV PC]. + econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. +Qed. + +Lemma symb_istep_correct sge i st rs0 m0 rs m: sem_local_istate sge (get_curr st) rs0 m0 rs m -> has_not_yet_exit sge (get_exit st) rs0 m0 -> - symb_istep i st = Some st' -> - step_istate sge.(the_ge) i sge.(the_sp) rs m = Some is -> - sem_istate sge st' rs0 m0 is. + sem_option2_istate sge (symb_istep i st) rs0 m0 + (step_istate sge.(the_ge) i sge.(the_sp) rs m). Proof. intros (PRE & MEM & REG) NYE. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + destruct i; simpl; auto. + + (* Nop *) + unfold sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + (* Op *) - inversion_SOME v. try_simplify_someHyps. - intros EVALOP EVALMEM; intuition. - * exploit REG. try_simplify_someHyps. - * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl; auto. - erewrite list_sval_eval_inj; simpl; auto. - rewrite EVALMEM, EVALOP; auto. + inversion_SOME v; intros OP; simpl. + - unfold sem_istate, sem_local_istate; simpl; intuition. + * exploit REG. try_simplify_someHyps. + * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl; auto. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + - unfold has_aborted, has_aborted_on_basic; simpl. + left. right. right. + exists r. destruct (Pos.eq_dec r r); try congruence. + simpl. erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + (* LOAD *) - inversion_SOME a0. inversion_SOME v. try_simplify_someHyps. - intros LOAD EVALADD EVALMEM; intuition. - * exploit REG. try_simplify_someHyps. - * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl. + inversion_SOME a0; intros ADD. + { inversion_SOME v; intros LOAD; simpl. + - unfold sem_istate, sem_local_istate; simpl; intuition. + * exploit REG. try_simplify_someHyps. + * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. + subst; rewrite Regmap.gss; simpl. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + - unfold has_aborted, has_aborted_on_basic; simpl. + left. right. right. + exists r. destruct (Pos.eq_dec r r); try congruence. + simpl. erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. } + { unfold has_aborted, has_aborted_on_basic; simpl. + left. right. right. + exists r. destruct (Pos.eq_dec r r); try congruence. + simpl. erewrite list_sval_eval_inj; simpl; auto. + rewrite ADD; simpl; auto. } + + (* STORE *) + inversion_SOME a0; intros ADD. + { inversion_SOME m'; intros STORE; simpl. + - unfold sem_istate, sem_local_istate; simpl; intuition. + * congruence. + * erewrite list_sval_eval_inj; simpl; auto. + erewrite REG. + try_simplify_someHyps. + - unfold has_aborted, has_aborted_on_basic; simpl. + left. right. left. + erewrite list_sval_eval_inj; simpl; auto. + erewrite REG. + try_simplify_someHyps. } + { unfold has_aborted, has_aborted_on_basic; simpl. + left. right. left. erewrite list_sval_eval_inj; simpl; auto. - try_simplify_someHyps. - + (* STORE *) - inversion_SOME a0. inversion_SOME m'. try_simplify_someHyps. - intros STORE EVALADD EVALMEM; intuition. - * discriminate. - * erewrite list_sval_eval_inj; simpl; auto. - erewrite REG. - try_simplify_someHyps. - + (* COND *) + erewrite ADD; simpl; auto. } + + (* COND *) Local Hint Resolve is_tail_refl. Local Hint Unfold sem_local_istate. - inversion_SOME b. destruct b; try_simplify_someHyps. - intros EVALCOND EVALMEM; econstructor; eauto; simpl. - * unfold seval_condition. - erewrite list_sval_eval_inj; simpl; auto. - try_simplify_someHyps. - * intuition. unfold has_not_yet_exit in * |- *. simpl. - intuition. subst. simpl. - unfold seval_condition. - erewrite list_sval_eval_inj; simpl; auto. - try_simplify_someHyps. + inversion_SOME b; intros COND. + { destruct b; simpl; unfold sem_istate, sem_local_istate; simpl. + - intros; econstructor; eauto; simpl. + unfold seval_condition. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. + - intuition. unfold has_not_yet_exit in * |- *. simpl. + intuition. subst. simpl. + unfold seval_condition. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. } + { unfold has_aborted, has_aborted_on_test; simpl. + right. eexists; intuition eauto. simpl. + unfold seval_condition. + erewrite list_sval_eval_inj; simpl; auto. + try_simplify_someHyps. } Qed. Lemma symb_istep_correct_None sge i st rs0 m0 rs m: @@ -277,83 +321,6 @@ Proof. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. Qed. -Lemma symb_istep_preserv_early_exits i sge st rs0 m0 rs m pc st': - sem_early_exit sge (get_exit st) rs0 m0 rs m pc -> - symb_istep i st = Some st' -> - sem_early_exit sge (get_exit st') rs0 m0 rs m pc. -Proof. - intros H. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. - destruct H as [ext lx TAIL NYE COND INV PC]. - econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. -Qed. - - -(* TODO: remove [istat_eq] ! This is useless (use sem_option2_istate instead !) -Definition istate_eq ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). - -Definition istate_eq_option ist1 oist := - exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. - -Lemma symb_istep_complete sge i st rs0 m0 rs m st' is: - is.(continue) = true -> - sem_local_istate sge (get_curr st) rs0 m0 rs m -> - sem_istate sge st' rs0 m0 is -> - symb_istep i st = Some st' -> - istate_eq_option is (step_istate sge.(the_ge) i sge.(the_sp) rs m). -Proof. - unfold sem_istate, istate_eq_option, istate_eq; destruct is as [c pc rs1 m1]; simpl. - intro; subst. - intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC & NYE); subst. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. - + (* Nop *) - intros; eexists; intuition eauto; simpl. - generalize (REG1 r); rewrite REG; try congruence. - + (* Op *) - intros MEM. - generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. - erewrite list_sval_eval_inj; simpl; auto. - rewrite MEM. - try_simplify_someHyps. - eexists; intuition eauto; simpl. - generalize (REG1 r0). - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. - subst; rewrite Regmap.gss. auto. - + (* LOAD *) - intros MEM. - generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. - erewrite list_sval_eval_inj; simpl; auto. - rewrite MEM. - inversion_SOME addr. - intros EVAL LOAD. rewrite EVAL, LOAD; simpl. - eexists; intuition eauto; simpl. - generalize (REG1 r0). - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. - subst; rewrite Regmap.gss. auto. - + (* STORE *) - intros MEM. - generalize (REG r); destruct (Pos.eq_dec r r); try congruence; simpl. - intros Hr; rewrite Hr; simpl. - erewrite list_sval_eval_inj; simpl; auto. - inversion_SOME addr. - intros EVAL STORE. rewrite EVAL, STORE; simpl. - eexists; intuition eauto; simpl. - generalize (REG r0); rewrite REG1. congruence. - + (* COND *) - intros MEM. - exploit NYE. { simpl; eauto. } - simpl; unfold seval_condition; erewrite list_sval_eval_inj; simpl; auto. - rewrite MEM. - intro COND; rewrite COND. eexists; intuition eauto. simpl. - generalize (REG r); rewrite REG1. congruence. -Qed. -*) - - (** * Symbolic execution of the internal steps of a path *) Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := match path with @@ -378,29 +345,37 @@ Proof. inversion_SOME st1; eauto. Qed. -Lemma symb_isteps_correct_true sge path f rs0 m0 st' is': forall st is, + +Lemma symb_isteps_correct_true sge path f rs0 m0: forall st is, is.(continue)=true -> sem_istate sge st rs0 m0 is -> - symb_isteps path f st = Some st' -> - path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc) = Some is' -> - sem_istate sge st' rs0 m0 is'. + (* nth_default_succ (fn_code f) path st.(the_pc) <> None -> *) + sem_option2_istate sge (symb_isteps path f st) rs0 m0 + (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) + . Proof. - Local Hint Resolve symb_istep_correct symb_isteps_correct_false. + Local Hint Resolve symb_isteps_correct_false. induction path; simpl. - + intros st is; + + intros st is CONT INV; unfold sem_istate, sist in * |- *; - try_simplify_someHyps. + try_simplify_someHyps. simpl. destruct is; simpl in * |- *; subst; intuition auto. + intros st is CONT; unfold sem_istate at 1; rewrite CONT. intros (LOCAL & PC & NYE). rewrite <- PC. - inversion_SOME i; try_simplify_someHyps; intro Hi. - inversion_SOME st1. - inversion_SOME is0. - destruct (continue is0) eqn:CONT0; eauto. - (* continue is0 = false -> EARLY EXIT *) - try_simplify_someHyps; eauto. -Qed. + inversion_SOME i; intro Hi; rewrite Hi; simpl; auto. + exploit symb_istep_correct; eauto. + inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. + - inversion_SOME is1; intros His1;rewrite His1; simpl. + * destruct (continue is1) eqn:CONT1; eauto. + (* continue is0 = false -> EARLY EXIT *) + destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. + admit. (* PB ici: il faut l'hypothèse nth_default_succ ? *) + (* try_simplify_someHyps; eauto. *) + * destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; auto. + admit. + - intros His1;rewrite His1; simpl; auto. +Admitted. (* -- cgit From 15a4794c0e0d313b825b3ec94819d9d2c1e5d186 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 21 Oct 2019 14:58:35 +0200 Subject: Good theory for the symbolic execution of "internal nodes" in RTLpath. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Pour la suite du dev du verificateur de scheduling, repartir "top-down" en essayant de se brancher là-dessus ? --- mppa_k1c/lib/RTLpathSE_theory.v | 155 ++++++++++++++++++++++++---------------- 1 file changed, 92 insertions(+), 63 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index cb7a6276..a7929e59 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -163,11 +163,6 @@ Definition sem_option2_istate sge (ost: option s_istate) rs0 m0 (ois: option ist | None => ois=None end. - -(** TODO: Lemmes de correction/complétudes ci-dessous à généraliser en utilisant sem_option2_istate !!! -*) - - (** * Symbolic execution of one internal step *) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := @@ -232,6 +227,49 @@ Proof. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. +Lemma sreg_set_preserv_has_aborted sge st rs0 m0 r sv: + has_aborted_on_basic sge st rs0 m0 -> + has_aborted_on_basic sge (sreg_set st r sv) rs0 m0. +Proof. + unfold has_aborted_on_basic. 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 smem_set_preserv_has_aborted sge st rs0 m0 m: + has_aborted_on_basic sge st rs0 m0 -> + has_aborted_on_basic sge (smem_set st m) rs0 m0. +Proof. + unfold has_aborted_on_basic. simpl; intuition. +Qed. + +Lemma symb_istep_preserv_has_aborted i sge rs0 m0 st st': + symb_istep i st = Some st' -> + has_aborted sge st rs0 m0 -> has_aborted sge st' rs0 m0. +Proof. + Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted. + unfold has_aborted. + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps; intuition eauto. + (* COND *) + right. unfold has_aborted_on_test in * |- *. + destruct H as (ext & H1 & H2). + simpl; exists ext; intuition. +Qed. + +Lemma symb_istep_WF i st: + symb_istep i st = None -> default_succ i = None. +Proof. + destruct i; simpl; unfold sist; simpl; congruence. +Qed. + +Lemma symb_istep_default_succ i st st': + symb_istep i st = Some st' -> default_succ i = Some (st'.(the_pc)). +Proof. + destruct i; simpl; unfold sist; simpl; try congruence; + intro H; inversion_clear H; simpl; auto. +Qed. + Lemma symb_istep_correct sge i st rs0 m0 rs m: sem_local_istate sge (get_curr st) rs0 m0 rs m -> has_not_yet_exit sge (get_exit st) rs0 m0 -> @@ -345,44 +383,79 @@ Proof. inversion_SOME st1; eauto. Qed. +Lemma symb_isteps_preserv_has_aborted sge path f rs0 m0 st': forall st, + symb_isteps path f st = Some st' -> + has_aborted sge st rs0 m0 -> has_aborted sge st' rs0 m0. +Proof. + Local Hint Resolve symb_istep_preserv_has_aborted. + induction path; simpl. + + unfold sist; try_simplify_someHyps. + + intros st; inversion_SOME i. + inversion_SOME st1; eauto. +Qed. + +Lemma symb_isteps_WF path f: forall st, + symb_isteps path f st = None -> nth_default_succ (fn_code f) path st.(the_pc) = None. +Proof. + induction path; simpl. + + unfold sist. intuition congruence. + + intros st; destruct ((fn_code f) ! (the_pc st)); simpl; try tauto. + destruct (symb_istep i st) as [st1|] eqn: Hst1; simpl. + - intros; erewrite symb_istep_default_succ; eauto. + - intros; erewrite symb_istep_WF; eauto. +Qed. -Lemma symb_isteps_correct_true sge path f rs0 m0: forall st is, +(* remark: unused lemma ? *) +Lemma symb_isteps_default_succ path f st': forall st, + symb_isteps path f st = Some st' -> nth_default_succ (fn_code f) path st.(the_pc) = Some st'.(the_pc). +Proof. + induction path; simpl. + + unfold sist. intros st H. inversion_clear H; simpl; try congruence. + + intros st; destruct ((fn_code f) ! (the_pc st)); simpl; try congruence. + destruct (symb_istep i st) as [st1|] eqn: Hst1; simpl; try congruence. + intros; erewrite symb_istep_default_succ; eauto. +Qed. + +Lemma symb_isteps_correct_true sge path (f:function) rs0 m0: forall st is, is.(continue)=true -> sem_istate sge st rs0 m0 is -> - (* nth_default_succ (fn_code f) path st.(the_pc) <> None -> *) + nth_default_succ (fn_code f) path st.(the_pc) <> None -> sem_option2_istate sge (symb_isteps path f st) rs0 m0 (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) . Proof. - Local Hint Resolve symb_isteps_correct_false. + Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF. induction path; simpl. - + intros st is CONT INV; + + intros st is CONT INV WF; unfold sem_istate, sist in * |- *; try_simplify_someHyps. simpl. destruct is; simpl in * |- *; subst; intuition auto. + intros st is CONT; unfold sem_istate at 1; rewrite CONT. - intros (LOCAL & PC & NYE). + intros (LOCAL & PC & NYE) WF. rewrite <- PC. - inversion_SOME i; intro Hi; rewrite Hi; simpl; auto. + inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. exploit symb_istep_correct; eauto. inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - inversion_SOME is1; intros His1;rewrite His1; simpl. - * destruct (continue is1) eqn:CONT1; eauto. + * destruct (continue is1) eqn:CONT1. + (* continue is0 = true *) + intros; eapply IHpath; eauto. + destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. (* continue is0 = false -> EARLY EXIT *) destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. - admit. (* PB ici: il faut l'hypothèse nth_default_succ ? *) + destruct WF. erewrite symb_istep_default_succ; eauto. (* try_simplify_someHyps; eauto. *) - * destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; auto. - admit. + * destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. - intros His1;rewrite His1; simpl; auto. -Admitted. +Qed. (* -(* Remark: the properties are not true, because of path=0! +(* Remark: the properties below do not hold, because of path=0! + if we use [symb_isteps 0%nat f st = Some st], they hold... + + But these properties are probably not very useful for our goal. - if we use [symb_isteps 0%nat f st = Some st], this works... - *) Lemma symb_isteps_right_assoc_decompose f path: forall st st', @@ -414,47 +487,3 @@ Proof. subst; eauto. Qed. *) - - -(* PROPRIETE FAUSSE: - -il faut utiliser sem_option2_istate - - -Lemma symb_isteps_complete sge path f rs0 m0 st': forall st, - is.(continue)=true -> - sem_istate sge st rs0 m0 is -> - symb_isteps path f st = Some st' -> - sem_istate sge st' rs0 m0 is' -> - (* nth_default_succ (fn_code f) path st.(the_pc) <> None -> *) - istate_eq_option is' (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)). -Proof. - induction path. - + unfold sem_istate, istate_eq_option, istate_eq; simpl. - intros st is CONT; rewrite CONT. - intros ((PRE & EVAL & REG) & H & NYE). - (* unfold sist; *) intros Hst'; inversion_clear Hst'. - destruct (continue is') eqn: CONT'. - * simpl; intros ((PRE1 & EVAL1 & REG1) & H1). - eexists; intuition eauto; simpl; try congruence. - generalize (REG1 r); rewrite REG. congruence. - * intros [ext lx TAIL NYE1 COND INV PC]. - exploit NYE. { eapply is_tail_in; eauto. } - congruence. - + simpl. unfold sem_istate at 1. - intros st is CONT; rewrite CONT. - intros (LOCAL & PC & NYE). - rewrite <- PC. - inversion_SOME i; intro Hi. rewrite Hi. - inversion_SOME st1. intros Hst1 Hst' SEMis' (* SUCC *) . - inversion_SOME is0. - * destruct (continue is0) eqn:CONT0; eauto. - (*- (* continue is0 = true *) - intros; eapply IHpath; eauto. - destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. - - *) (* continue is0 = false -> EARLY EXIT *) - admit. - * destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. - - generalize H; inversion_SOME v; try congruence. -Abort. -*) \ No newline at end of file -- cgit From 99bbf17704e8295346aa2019bd1c86513f42e91d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 21 Oct 2019 15:10:02 +0200 Subject: rename: step_istate -> istep | path_step_istate -> isteps --- mppa_k1c/lib/RTLpath.v | 116 ++++++++++++++++++++++++------------------------- 1 file changed, 58 insertions(+), 58 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 4e25b2b4..008733a6 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -133,7 +133,7 @@ Coercion program_RTL: program >-> RTL.program. Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. -Definition step_istate (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := +Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with | Inop pc' => Some (ist true pc' rs m) | Iop op args res pc' => @@ -155,14 +155,14 @@ Definition step_istate (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m (** Execution of a path in a single step *) -Fixpoint path_step_istate ge (path:nat) (f: function) sp rs m pc: option istate := +Fixpoint isteps ge (path:nat) (f: function) sp rs m pc: option istate := match path with | O => Some (ist true pc rs m) | S p => SOME i <- (fn_code f)!pc IN - SOME st <- step_istate ge i sp rs m IN + SOME st <- istep ge i sp rs m IN if st.(continue) then - path_step_istate ge p f sp st.(the_rs) st.(the_mem) st.(the_pc) + isteps ge p f sp st.(the_rs) st.(the_mem) st.(the_pc) else Some st end. @@ -227,7 +227,7 @@ Coercion state_RTL: state >-> RTL.state. Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := | exec_istate i sp pc rs m st: (fn_code f)!pc = Some i -> - step_istate ge i sp rs m = Some st -> + istep ge i sp rs m = Some st -> path_last_step ge pge stack f sp pc rs m E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) | exec_Icall sp pc rs m sig ros args res pc' fd: @@ -263,11 +263,11 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := | exec_early_exit st: - path_step_istate ge path f sp rs m pc = Some st -> + isteps ge path f sp rs m pc = Some st -> st.(continue) = false -> path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) | exec_normal_exit st t s: - path_step_istate ge path f sp rs m pc = Some st -> + isteps ge path f sp rs m pc = Some st -> st.(continue) = true -> path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> path_step ge pge path stack f sp rs m pc t s. @@ -346,8 +346,8 @@ This way can be viewed as a correctness property: all transitions in RTLpath are Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. -Lemma step_istate_correct ge i stack (f:function) sp rs m st : - step_istate ge i sp rs m = Some st -> +Lemma istep_correct ge i stack (f:function) sp rs m st : + istep ge i sp rs m = Some st -> forall pc, (fn_code f)!pc = Some i -> RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. @@ -356,8 +356,8 @@ Qed. Local Hint Resolve star_refl. -Lemma path_step_istate_correct ge path stack f sp: forall rs m pc st, - path_step_istate ge path f sp rs m pc = Some st -> +Lemma isteps_correct ge path stack f sp: forall rs m pc st, + isteps ge path f sp rs m pc = Some st -> star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. induction path; simpl; try_simplify_someHyps. @@ -365,17 +365,17 @@ Proof. inversion_SOME st0; intros Hst0. destruct (continue st0) eqn:cont. + intros; eapply star_step. - - eapply step_istate_correct; eauto. + - eapply istep_correct; eauto. - simpl; eauto. - auto. + intros; simplify_someHyp; eapply star_step. - - eapply step_istate_correct; eauto. + - eapply istep_correct; eauto. - simpl; eauto. - auto. Qed. -Lemma path_step_istate_correct_early_exit ge path stack f sp: forall rs m pc st, - path_step_istate ge path f sp rs m pc = Some st -> +Lemma isteps_correct_early_exit ge path stack f sp: forall rs m pc st, + isteps ge path f sp rs m pc = Some st -> st.(continue) = false -> plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. @@ -384,12 +384,12 @@ Proof. inversion_SOME st0; intros Hst0. destruct (continue st0) eqn:cont. + intros; eapply plus_left. - - eapply step_istate_correct; eauto. - - eapply path_step_istate_correct; eauto. + - eapply istep_correct; eauto. + - eapply isteps_correct; eauto. - auto. + intros X; inversion X; subst. eapply plus_one. - eapply step_istate_correct; eauto. + eapply istep_correct; eauto. Qed. Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. @@ -424,13 +424,13 @@ Proof. intros (cuint & tf & H1 & H2 & H3); subst; auto. Qed. -Local Hint Resolve step_istate_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. +Local Hint Resolve istep_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. Lemma path_last_step_correct stack f sp pc rs m t s: path_last_step ge pge stack f sp pc rs m t s -> RTL.step ge (State stack f sp pc rs m) t s. Proof. - destruct 1; try (eapply step_istate_correct); simpl; eauto. + destruct 1; try (eapply istep_correct); simpl; eauto. Qed. Lemma path_step_correct path stack f sp pc rs m t s: @@ -438,9 +438,9 @@ Lemma path_step_correct path stack f sp pc rs m t s: plus RTL.step ge (State stack f sp pc rs m) t s. Proof. destruct 1. - + eapply path_step_istate_correct_early_exit; eauto. + + eapply isteps_correct_early_exit; eauto. + eapply plus_right. - eapply path_step_istate_correct; eauto. + eapply isteps_correct; eauto. eapply path_last_step_correct; eauto. auto. Qed. @@ -551,16 +551,16 @@ Qed. Local Hint Resolve fn_entry_point_wf. Local Opaque path_entry. -Lemma step_istate_successors ge i sp rs m st: - step_istate ge i sp rs m = Some st -> +Lemma istep_successors ge i sp rs m st: + istep ge i sp rs m = Some st -> In (the_pc st) (successors_instr i). Proof. destruct i; simpl; try congruence; simplify_SOME x. destruct x; simpl in * |- *; intuition congruence. Qed. -Lemma step_istate_normal_exit ge i sp rs m st: - step_istate ge i sp rs m = Some st -> +Lemma istep_normal_exit ge i sp rs m st: + istep ge i sp rs m = Some st -> st.(continue) = true -> default_succ i = Some st.(the_pc). Proof. @@ -568,9 +568,9 @@ Proof. destruct x; simpl in * |- *; try congruence. Qed. -Lemma path_step_istate_normal_exit ge path f sp: forall rs m pc st, +Lemma isteps_normal_exit ge path f sp: forall rs m pc st, st.(continue) = true -> - path_step_istate ge path f sp rs m pc = Some st -> + isteps ge path f sp rs m pc = Some st -> nth_default_succ (fn_code f) path pc = Some st.(the_pc). Proof. induction path; simpl. { try_simplify_someHyps. } @@ -579,18 +579,18 @@ Proof. inversion_SOME st0; intros Hst0. destruct (continue st0) eqn:X; try congruence. try_simplify_someHyps. - intros; erewrite step_istate_normal_exit; eauto. + intros; erewrite istep_normal_exit; eauto. Qed. (* TODO: the three following lemmas could maybe simplified by introducing an auxiliary - left-recursive definition equivalent to path_step_istate ? + left-recursive definition equivalent to isteps ? *) -Lemma path_step_istate_step_right ge path f sp: forall rs m pc st i, - path_step_istate ge path f sp rs m pc = Some st -> +Lemma isteps_step_right ge path f sp: forall rs m pc st i, + isteps ge path f sp rs m pc = Some st -> st.(continue) = true -> (fn_code f)!(st.(the_pc)) = Some i -> - step_istate ge i sp st.(the_rs) st.(the_mem) = path_step_istate ge (S path) f sp rs m pc. + istep ge i sp st.(the_rs) st.(the_mem) = isteps ge (S path) f sp rs m pc. Proof. induction path. + simpl; intros; try_simplify_someHyps. simplify_SOME st. @@ -603,15 +603,15 @@ Proof. * intros; congruence. Qed. -Lemma path_step_istate_inversion_early ge path f sp: forall rs m pc st, - path_step_istate ge path f sp rs m pc = Some st -> +Lemma isteps_inversion_early ge path f sp: forall rs m pc st, + isteps ge path f sp rs m pc = Some st -> (continue st)=false -> exists st0 i path0, (path > path0)%nat /\ - path_step_istate ge path0 f sp rs m pc = Some st0 /\ + isteps ge path0 f sp rs m pc = Some st0 /\ st0.(continue) = true /\ (fn_code f)!(st0.(the_pc)) = Some i /\ - step_istate ge i sp st0.(the_rs) st0.(the_mem) = Some st. + istep ge i sp st0.(the_rs) st0.(the_mem) = Some st. Proof. induction path as [|path]; simpl. - intros; try_simplify_someHyps; try congruence. @@ -629,11 +629,11 @@ Proof. omega. Qed. -Lemma path_step_istate_resize ge path0 path1 f sp rs m pc st: +Lemma isteps_resize ge path0 path1 f sp rs m pc st: (path0 <= path1)%nat -> - path_step_istate ge path0 f sp rs m pc = Some st -> + isteps ge path0 f sp rs m pc = Some st -> (continue st)=false -> - path_step_istate ge path1 f sp rs m pc = Some st. + isteps ge path1 f sp rs m pc = Some st. Proof. induction 1 as [|path1]; simpl; auto. intros PSTEP CONT. exploit IHle; auto. clear PSTEP IHle H path0. @@ -649,8 +649,8 @@ Inductive is_early_exit pc: instruction -> Prop := is_early_exit pc (Icond cond args pc ifnot) . (* TODO add jumptable here ? *) -Lemma step_istate_early_exit ge i sp rs m st : - step_istate ge i sp rs m = Some st -> +Lemma istep_early_exit ge i sp rs m st : + istep ge i sp rs m = Some st -> st.(continue) = false -> st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. Proof. @@ -726,7 +726,7 @@ Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> (idx <= path)%nat -> - path_step_istate ge (path-idx) f sp rs m pc = Some s2 -> + isteps ge (path-idx) f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> match_inst_states_goal idx s1 (State stack f sp pc rs m). @@ -738,18 +738,18 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := /\ wf_stackframe (stack_of s2). (** *** Auxiliary lemmas of completeness *) -Lemma step_istate_complete t i stack f sp rs m pc s': +Lemma istep_complete t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> (fn_code f)!pc = Some i -> default_succ i <> None -> - t = E0 /\ exists st, step_istate ge i sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + t = E0 /\ exists st, istep ge i sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. Qed. Lemma stuttering path idx stack f sp rs m pc st t s1': - path_step_istate ge (path-(S idx)) f sp rs m pc = Some st -> + isteps ge (path-(S idx)) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (S idx <= path)%nat -> st.(continue) = true -> @@ -759,17 +759,17 @@ Proof. intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path-(S idx))); omega || eauto. intros (i & pc' & Hi & Hpc & DUM). unfold nth_default_succ_inst in Hi. - erewrite path_step_istate_normal_exit in Hi; eauto. - exploit step_istate_complete; congruence || eauto. + erewrite isteps_normal_exit in Hi; eauto. + exploit istep_complete; congruence || eauto. intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. - erewrite <- path_step_istate_step_right; eauto. + erewrite <- isteps_step_right; eauto. Qed. Lemma normal_exit path stack f sp rs m pc st t s1': - path_step_istate ge path f sp rs m pc = Some st -> + isteps ge path f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> @@ -778,14 +778,14 @@ Lemma normal_exit path stack f sp rs m pc st t s1': (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' /\ (exists idx', match_states idx' s1' s2'). Proof. - Local Hint Resolve step_istate_successors list_nth_z_in. (* Hint for path_entry proofs *) + Local Hint Resolve istep_successors list_nth_z_in. (* Hint for path_entry proofs *) intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto. intros (i & Hi & SUCCS). unfold nth_default_succ_inst in Hi. - erewrite path_step_istate_normal_exit in Hi; eauto. + erewrite isteps_normal_exit in Hi; eauto. destruct (default_succ i) eqn:Hn0. + (* exec_istate *) - exploit step_istate_complete; congruence || eauto. + exploit istep_complete; congruence || eauto. intros (SILENT & st0 & STEP0 & EQ); subst. exploit (exec_istate ge pge); eauto. eexists; intuition eauto. @@ -834,7 +834,7 @@ Proof. Qed. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: - path_step_istate ge (path-idx) f sp rs m pc = Some st -> + isteps ge (path-idx) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (idx <= path)%nat -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> @@ -860,20 +860,20 @@ Proof. unfold match_states; exists idx; exists (State stack f sp pc rs m); intuition. + (* one or two path_step on early_exit *) - exploit (path_step_istate_resize ge (path - idx)%nat path); eauto; try omega. + exploit (isteps_resize ge (path - idx)%nat path); eauto; try omega. clear PSTEP; intros PSTEP. (* TODO for clarification: move the assert below into a separate lemma *) assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). { clear RSTEP. - exploit path_step_istate_inversion_early; eauto. + exploit isteps_inversion_early; eauto. intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). - exploit step_istate_early_exit; eauto. + exploit istep_early_exit; eauto. intros (X1 & X2 & EARLY_EXIT). destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. exploit (internal_node_path path0); omega || eauto. intros (i' & pc' & Hi' & Hpc' & ENTRY). unfold nth_default_succ_inst in Hi'. - erewrite path_step_istate_normal_exit in Hi'; eauto. + erewrite isteps_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; destruct (initialize_path (fn_path f) pc0); eauto. -- cgit From 4d94fbaaf42d5670c8dd4a0f6fbb0c9b71b781ad Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 21 Oct 2019 15:22:58 +0200 Subject: stronger notion of path_entry --- mppa_k1c/lib/RTLpath.v | 33 +++++++++++++++------------------ 1 file changed, 15 insertions(+), 18 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 008733a6..95818c1d 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -77,21 +77,18 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, Definition path_map: Type := PTree.t nat. -Definition path_entry (*c:code*) (pm: path_map) (n: node): Prop - := pm!n = None -> False. -(* NB: In theory, we could relax the above [False] into [c!n=None], but this seems PEDANTIC in practice ! - Indeed, in practice, the CFG will not have dandling node ! -*) +Definition path_entry (c:code) (pm: path_map) (n: node): Prop + := pm!n <> None /\ c!n <> None. Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := | wf_last_node i pc: c!pc = Some i -> - (forall n, List.In n (successors_instr i) -> path_entry (*c*) pm n) -> + (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> wellformed_path c pm 0 pc | wf_internal_node path i pc pc': c!pc = Some i -> default_succ i = Some pc' -> - (forall n, early_exit i = Some n -> path_entry (*c*) pm n) -> + (forall n, early_exit i = Some n -> path_entry c pm n) -> wellformed_path c pm path pc' -> wellformed_path c pm (S path) pc. @@ -108,7 +105,7 @@ Record function : Type := { fn_RTL:> RTL.function; fn_path: path_map; (* condition 1 below: the entry-point of the code is an entry-point of a path *) - fn_entry_point_wf: path_entry (*fn_RTL.(fn_code)*) fn_path fn_RTL.(fn_entrypoint); + fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); (* condition 2 below: the path_map is well-formed *) fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path }. @@ -516,7 +513,7 @@ Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := Lemma final_node_path f path pc: (fn_path f)!pc = Some path -> exists i, nth_default_succ_inst (fn_code f) path pc = Some i - /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n). + /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF. @@ -533,7 +530,7 @@ Lemma internal_node_path path f path0 pc: exists i pc', nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ - (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n). + (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } @@ -544,7 +541,7 @@ Proof. simplify_someHyps; eauto. Qed. -Lemma initialize_path (*c*) pm n (*i*): (* PEDANTIC: c!n = Some i -> *) path_entry (*c*) pm n -> exists path, pm!n = Some path. +Lemma initialize_path c pm n: path_entry c pm n -> exists path, pm!n = Some path. Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. @@ -691,7 +688,7 @@ Qed. (** *** Definition of well-formed stacks and of match_states *) Definition wf_stf (st: stackframe): Prop := match st with - | Stackframe res f sp pc rs => path_entry (*f.(fn_code)*) f.(fn_path) pc + | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc end. Definition wf_stackframe (stack: list stackframe): Prop := @@ -790,7 +787,7 @@ Proof. exploit (exec_istate ge pge); eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. + destruct (initialize_path (fn_code f) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -813,7 +810,7 @@ Proof. intros; exploit exec_Ibuiltin; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_path f) pc') as (path0 & Hpath0); eauto. + destruct (initialize_path (fn_code f) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -822,7 +819,7 @@ Proof. intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_path f) pc') as (path0 & Hpath0); eauto. + destruct (initialize_path (fn_code f) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -876,7 +873,7 @@ Proof. erewrite isteps_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; - destruct (initialize_path (fn_path f) pc0); eauto. + destruct (initialize_path (fn_code f) (fn_path f) pc0); eauto. } destruct HPATH0 as ([|path1] & Hpath1). * (* two step case *) @@ -912,7 +909,7 @@ Proof. eexists; constructor 1. * eapply exec_function_internal; eauto. * unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. + destruct (initialize_path (fn_code f) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. @@ -931,7 +928,7 @@ Proof. * unfold match_states, match_inst_states; simpl. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. - destruct (initialize_path (fn_path f0) pc0) as (path & Hpath); eauto. + destruct (initialize_path (fn_code f0) (fn_path f0) pc0) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. -- cgit From 9bcfc2f5ca0a53366c9eec45211b35db7daa2a02 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 21 Oct 2019 15:29:34 +0200 Subject: update wrt renaming in RTLpath --- mppa_k1c/lib/RTLpathSE_theory.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index a7929e59..d6959338 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -274,7 +274,7 @@ Lemma symb_istep_correct sge i st rs0 m0 rs m: sem_local_istate sge (get_curr st) rs0 m0 rs m -> has_not_yet_exit sge (get_exit st) rs0 m0 -> sem_option2_istate sge (symb_istep i st) rs0 m0 - (step_istate sge.(the_ge) i sge.(the_sp) rs m). + (istep sge.(the_ge) i sge.(the_sp) rs m). Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. @@ -353,7 +353,7 @@ Qed. Lemma symb_istep_correct_None sge i st rs0 m0 rs m: sem_local_istate sge (get_curr st) rs0 m0 rs m -> symb_istep i st = None -> - step_istate sge.(the_ge) i sge.(the_sp) rs m = None. + istep sge.(the_ge) i sge.(the_sp) rs m = None. Proof. intros (PRE & MEM & REG). destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -421,7 +421,7 @@ Lemma symb_isteps_correct_true sge path (f:function) rs0 m0: forall st is, sem_istate sge st rs0 m0 is -> nth_default_succ (fn_code f) path st.(the_pc) <> None -> sem_option2_istate sge (symb_isteps path f st) rs0 m0 - (path_step_istate sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) + (isteps sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) . Proof. Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF. -- cgit From a374f62201672ef3d5035c1c5899a0860caa93c6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 23 Oct 2019 12:27:35 +0200 Subject: a first skeleton of certified RTLpath scheduler (just to sketch ideas) --- mppa_k1c/lib/RTLpathScheduler.v | 147 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 147 insertions(+) create mode 100644 mppa_k1c/lib/RTLpathScheduler.v diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v new file mode 100644 index 00000000..76795fa4 --- /dev/null +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -0,0 +1,147 @@ +(** RTLpath Scheduling from an external oracle. + +This module is inspired from [Duplicate] + +*) + +Require Import AST Values RTL RTLpath Maps Globalenvs. +Require Import Coqlib Errors Op. + +Local Open Scope error_monad_scope. +Local Open Scope positive_scope. + + + +(** External oracle returning the new RTLpath function and a mapping of new path_entries to old path_entries + +NB: the new RTLpath function is generated from the fn_code, the fn_entrypoint and the fn_path +It requires to check that the path structure is wf ! + +*) +Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node). + +(* +Extract Constant untrusted_scheduler => "TODO". +*) + +Record xfunction : Type := + { fn_RTL: RTLpath.function; + fn_revmap: PTree.t node; + }. + +Definition xfundef := AST.fundef xfunction. +Definition xprogram := AST.program xfundef unit. +Definition xgenv := Genv.t xfundef unit. + +Definition fundef_RTL (fu: xfundef) : fundef := + match fu with + | Internal f => Internal (fn_RTL f) + | External ef => External ef + end. + +(* auxiliary definitions that should be use in [verify_match_path] implementation below ! *) +Definition verify_is_copy revmap n n' := + match revmap!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 revmap ln ln' := + match ln with + | n::ln => match ln' with + | n'::ln' => do u <- verify_is_copy revmap n n'; + verify_is_copy_list revmap 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. + +Definition verify_mapping_entrypoint (f: function) (xf: xfunction): res unit := + verify_is_copy (fn_revmap xf) (fn_entrypoint f) (fn_entrypoint (fn_RTL xf)). + + +(* This verifier remains to be implemented by symbolic execution *) +Parameter verify_match_path: function -> xfunction -> node -> node -> bool. + +(* +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. + +(** 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. + +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_mapping_mn_rec f xf lm := + match lm with + | nil => OK tt + | (pc2,pc1) :: lm => + if verify_match_path f xf pc1 pc2 + then + do u2 <- verify_mapping_mn_rec f xf lm; + OK tt + else + Error (msg "verify_match_path fails") + end. + +Definition verify_mapping_match_paths (f: function) (xf: xfunction) : res unit := + verify_mapping_mn_rec f xf (PTree.elements (fn_revmap xf)). + +(** Verifies that the [fn_revmap] of the translated function [xf] is giving correct information in regards to [f] *) +Definition verify_mapping (f: function) (xf: xfunction) : res unit := + do u <- verify_mapping_entrypoint f xf; + do v <- verify_mapping_match_paths f xf; OK tt. +(* TODO - verify the other axiom *) + +(** * Entry points *) + +(* TODO: à finir... + +Definition transf_function_aux (f: function) : res xfunction := + let (tcte, mp) := duplicate_aux f in + let (tc, te) := tcte in + let xf := {| fn_RTL := (mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te); fn_revmap := mp |} in + do u <- verify_mapping f xf; + OK xf. + +Theorem transf_function_aux_preserves: + forall f xf, + transf_function_aux f = OK xf -> + fn_sig f = fn_sig (fn_RTL xf) /\ fn_params f = fn_params (fn_RTL xf) /\ fn_stacksize f = fn_stacksize (fn_RTL xf). +Proof. + intros. unfold transf_function_aux in H. destruct (duplicate_aux _) as (tcte & mp). destruct tcte as (tc & te). monadInv H. + repeat (split; try reflexivity). +Qed. + +Remark transf_function_aux_fnsig: forall f xf, transf_function_aux f = OK xf -> fn_sig f = fn_sig (fn_RTL xf). + Proof. apply transf_function_aux_preserves. Qed. +Remark transf_function_aux_fnparams: forall f xf, transf_function_aux f = OK xf -> fn_params f = fn_params (fn_RTL xf). + Proof. apply transf_function_aux_preserves. Qed. +Remark transf_function_aux_fnstacksize: forall f xf, transf_function_aux f = OK xf -> fn_stacksize f = fn_stacksize (fn_RTL xf). + Proof. apply transf_function_aux_preserves. Qed. + +Definition transf_function (f: function) : res function := + do xf <- transf_function_aux f; + OK (fn_RTL xf). + +Definition transf_fundef (f: fundef) : res fundef := + transf_partial_fundef transf_function f. + +Definition transf_program (p: program) : res program := + transform_partial_program transf_fundef p. +*) \ No newline at end of file -- cgit From 0590d1be2a15927f5835833416541a9048568bb0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 5 Nov 2019 09:59:20 +0100 Subject: simplification of [s_istate] --- mppa_k1c/lib/RTLpathSE_theory.v | 72 ++++++++++++++--------------------------- 1 file changed, 25 insertions(+), 47 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index d6959338..3694ca9c 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -122,33 +122,19 @@ Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: m (** * Syntax and Semantics of symbolic internal state *) -Record s_istate := { the_pc: node; early_exit: list exit_istate; curr: local_istate + exit_istate }. - -Definition get_curr (st:s_istate): local_istate := - let (_,_,c) := st in - match c with - | inl x => x - | inr x => x.(exit_ist) - end. - -Definition get_exit (st:s_istate): list exit_istate := - let (_,_,c) := st in - match c with - | inl _ => st.(early_exit) - | inr cnd => cnd::(st.(early_exit)) - end. +Record s_istate := { the_pc: node; exits: list exit_istate; curr: local_istate }. Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(continue)) then - sem_local_istate sge (get_curr st) rs0 m0 is.(the_rs) is.(the_mem) + sem_local_istate sge st.(curr) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) - /\ has_not_yet_exit sge (get_exit st) rs0 m0 - else sem_early_exit sge (get_exit st) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). + /\ has_not_yet_exit sge st.(exits) rs0 m0 + else sem_early_exit sge st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). Definition has_aborted (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem): Prop := - has_aborted_on_basic sge (get_curr st) rs0 m0 - \/ has_aborted_on_test sge (get_exit st) rs0 m0. + has_aborted_on_basic sge st.(curr) rs0 m0 + \/ has_aborted_on_test sge st.(exits) rs0 m0. Definition sem_option_istate sge (st: s_istate) rs0 m0 (ois: option istate): Prop := @@ -175,33 +161,33 @@ Definition smem_set (st:local_istate) (sm:smem) := the_sreg:= st.(the_sreg); the_smem:= sm |}. -Definition sist (st: s_istate) (pc: node) (nxt: local_istate + exit_istate): option s_istate := - Some {| the_pc := pc; early_exit := get_exit st; curr:= nxt |}. +Definition sist (st: s_istate) (pc: node) (nxt: local_istate): option s_istate := + Some {| the_pc := pc; exits := st.(exits); curr:= nxt |}. Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with | Inop pc' => - sist st pc' (inl (get_curr st)) + sist st pc' st.(curr) | Iop op args dst pc' => - let prev := get_curr st in + let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in - sist st pc' (inl next) + sist st pc' next | Iload chunk addr args dst pc' => - let prev := get_curr st in + let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in - sist st pc' (inl next) + sist st pc' next | Istore chunk addr args src pc' => - let prev := get_curr st in + let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in - sist st pc' (inl next) + sist st pc' next | Icond cond args ifso ifnot => - let prev := get_curr st in + let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in - let next := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in - sist st ifnot (inr next) + let ex := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in + Some {| the_pc := ifnot; exits := ex::st.(exits); curr := prev |} | _ => None (* TODO jumptable ? *) end. @@ -217,9 +203,9 @@ Proof. Qed. Lemma symb_istep_preserv_early_exits i sge st rs0 m0 rs m pc st': - sem_early_exit sge (get_exit st) rs0 m0 rs m pc -> + sem_early_exit sge st.(exits) rs0 m0 rs m pc -> symb_istep i st = Some st' -> - sem_early_exit sge (get_exit st') rs0 m0 rs m pc. + sem_early_exit sge st'.(exits) rs0 m0 rs m pc. Proof. intros H. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -271,8 +257,8 @@ Proof. Qed. Lemma symb_istep_correct sge i st rs0 m0 rs m: - sem_local_istate sge (get_curr st) rs0 m0 rs m -> - has_not_yet_exit sge (get_exit st) rs0 m0 -> + sem_local_istate sge st.(curr) rs0 m0 rs m -> + has_not_yet_exit sge st.(exits) rs0 m0 -> sem_option2_istate sge (symb_istep i st) rs0 m0 (istep sge.(the_ge) i sge.(the_sp) rs m). Proof. @@ -351,7 +337,7 @@ Proof. Qed. Lemma symb_istep_correct_None sge i st rs0 m0 rs m: - sem_local_istate sge (get_curr st) rs0 m0 rs m -> + sem_local_istate sge (st.(curr)) rs0 m0 rs m -> symb_istep i st = None -> istep sge.(the_ge) i sge.(the_sp) rs m = None. Proof. @@ -362,7 +348,7 @@ Qed. (** * Symbolic execution of the internal steps of a path *) Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := match path with - | O => sist st st.(the_pc) (inl (get_curr st)) + | O => Some st | S p => SOME i <- (fn_code f)!(st.(the_pc)) IN SOME st1 <- symb_istep i st IN @@ -449,14 +435,7 @@ Proof. - intros His1;rewrite His1; simpl; auto. Qed. -(* - -(* Remark: the properties below do not hold, because of path=0! - if we use [symb_isteps 0%nat f st = Some st], they hold... - - But these properties are probably not very useful for our goal. - -*) +(** Auxilary properties, that are not useful here ? *) Lemma symb_isteps_right_assoc_decompose f path: forall st st', symb_isteps (S path) f st = Some st' -> @@ -486,4 +465,3 @@ Proof. inversion_SOME st1. intros Hst1; rewrite Hst1. subst; eauto. Qed. -*) -- cgit From d8f25c75d575b3d107add470490ceebf6cbea585 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 5 Nov 2019 17:34:49 +0100 Subject: notion of symbolic state (for full path execution) --- mppa_k1c/lib/RTLpathSE_theory.v | 74 ++++++++++++++++++++++++++++++++++++++++- 1 file changed, 73 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 3694ca9c..b99d7f27 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -22,7 +22,8 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ (** * Syntax and semantics of symbolic values *) -Record sgenv := { +Record sgenv := { + the_pge: RTLpath.genv; the_ge: RTL.genv; the_sp: val }. @@ -465,3 +466,74 @@ Proof. inversion_SOME st1. intros Hst1; rewrite Hst1. subst; eauto. Qed. + +(** * Symbolic (final) value of a path *) +Inductive sfval := + | Snone + | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) + (* NB: [res] the return register is hard-wired ! Is it restrictive ? *) +(* + | Stailcall: signature -> sval + ident -> list_sval -> sfval + | Sbuiltin: external_function -> list (builtin_arg sval) -> builtin_res sval -> sfval + | Sjumptable: sval -> list node -> sfval + | Sreturn: option sval -> sfval +*) +. + +Definition s_find_function (sge : sgenv) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := + match svos with + | inl sv => SOME v <- sval_eval sge sv rs0 m0 IN Genv.find_funct sge.(the_pge) v + | inr symb => SOME b <- Genv.find_symbol sge.(the_pge) symb IN Genv.find_funct_ptr sge.(the_pge) b + end. + +Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := + | exec_Snone rs m: + sfval_sem sge st stack f rs0 m0 Snone rs m E0 (State stack f sge.(the_sp) st.(the_pc) rs m) + | exec_Scall rs m sig svos lsv args res fd: + s_find_function sge svos rs0 m0 = Some fd -> + funsig fd = sig -> + list_sval_eval sge lsv rs0 m0 = Some args -> + sfval_sem sge st stack f rs0 m0 (Scall sig svos lsv res) rs m + E0 (Callstate (Stackframe res f sge.(the_sp) st.(the_pc) rs :: stack) fd args m) +(* TODO: + | exec_Itailcall stk pc rs m sig ros args fd m': + (fn_code f)!pc = Some(Itailcall sig ros args) -> + find_function pge ros rs = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Callstate stack fd rs##args m') + | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': + (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> + eval_builtin_args ge (fun r => rs#r) sp m args vargs -> + external_call ef ge vargs m t vres m' -> + path_last_step ge pge stack f sp pc rs m + t (State stack f sp pc' (regmap_setres res vres rs) m') + | exec_Ijumptable sp pc rs m arg tbl n pc': (* TODO remove jumptable from here ? *) + (fn_code f)!pc = Some(Ijumptable arg tbl) -> + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + path_last_step ge pge stack f sp pc rs m + E0 (State stack f sp pc' rs m) + | exec_Ireturn stk pc rs m or m': + (fn_code f)!pc = Some(Ireturn or) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m + E0 (Returnstate stack (regmap_optget or Vundef rs) m') +*) +. + +Record s_state := { internal:> s_istate; final: sfval }. + +Inductive sem_state (sge: sgenv) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := + | sem_have_early_exit is: + is.(continue) = false -> + sem_istate sge st rs0 m0 is -> + sem_state sge st stack f rs0 m0 E0 (State stack f sge.(the_sp) is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) + | sem_normal_exit is t s: + is.(continue) = true -> + sem_istate sge st rs0 m0 is -> + sfval_sem sge st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> + sem_state sge st stack f rs0 m0 t s + . + -- cgit From 6a8b1b9866dc78f821da52c4a746d7cb8f4215e1 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 6 Nov 2019 08:38:13 +0100 Subject: main function of the symbolic execution --- mppa_k1c/lib/RTLpathSE_theory.v | 47 ++++++++++++++++++++++++++++++++--------- 1 file changed, 37 insertions(+), 10 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index b99d7f27..efde229d 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -470,14 +470,14 @@ Qed. (** * Symbolic (final) value of a path *) Inductive sfval := | Snone - | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) + | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:node) (* NB: [res] the return register is hard-wired ! Is it restrictive ? *) (* | Stailcall: signature -> sval + ident -> list_sval -> sfval | Sbuiltin: external_function -> list (builtin_arg sval) -> builtin_res sval -> sfval | Sjumptable: sval -> list node -> sfval - | Sreturn: option sval -> sfval *) + | Sreturn: option sval -> sfval . Definition s_find_function (sge : sgenv) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := @@ -489,12 +489,12 @@ Definition s_find_function (sge : sgenv) (svos : sval + ident) (rs0: regset) (m0 Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: sfval_sem sge st stack f rs0 m0 Snone rs m E0 (State stack f sge.(the_sp) st.(the_pc) rs m) - | exec_Scall rs m sig svos lsv args res fd: + | exec_Scall rs m sig svos lsv args res pc fd: s_find_function sge svos rs0 m0 = Some fd -> funsig fd = sig -> list_sval_eval sge lsv rs0 m0 = Some args -> - sfval_sem sge st stack f rs0 m0 (Scall sig svos lsv res) rs m - E0 (Callstate (Stackframe res f sge.(the_sp) st.(the_pc) rs :: stack) fd args m) + sfval_sem sge st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res f sge.(the_sp) pc rs :: stack) fd args m) (* TODO: | exec_Itailcall stk pc rs m sig ros args fd m': (fn_code f)!pc = Some(Itailcall sig ros args) -> @@ -515,12 +515,13 @@ Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset list_nth_z tbl (Int.unsigned n) = Some pc' -> path_last_step ge pge stack f sp pc rs m E0 (State stack f sp pc' rs m) - | exec_Ireturn stk pc rs m or m': - (fn_code f)!pc = Some(Ireturn or) -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m - E0 (Returnstate stack (regmap_optget or Vundef rs) m') *) + | exec_Sreturn stk osv rs m m' v: + sge.(the_sp) = (Vptr stk Ptrofs.zero) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + match osv with Some sv => sval_eval sge sv rs0 m0 | None => Some Vundef end = Some v -> + sfval_sem sge st stack f rs0 m0 (Sreturn osv) rs m + E0 (Returnstate stack v m') . Record s_state := { internal:> s_istate; final: sfval }. @@ -537,3 +538,29 @@ Inductive sem_state (sge: sgenv) (st: s_state) stack f (rs0: regset) (m0: mem): sem_state sge st stack f rs0 m0 t s . +(** * Symbolic execution of final step *) +Definition symb_fstep (i: instruction) (prev: local_istate): sfval := + match i with + | Icall sig ros args res pc => + let svos := sum_left_map prev.(the_sreg) ros in + let sargs := list_val_inj (List.map prev.(the_sreg) args) in + Scall sig svos sargs res pc + | Ireturn or => + let sor := SOME r <- or IN Some (prev.(the_sreg) r) in + Sreturn sor + | _ => Snone + end. + +(** * Main function of the symbolic execution *) +Definition init_local_istate := {| pre:= fun _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. + +Definition init_istate pc := {| the_pc:= pc; exits:=nil; curr:= init_local_istate |}. + +Fixpoint symb_step (f: function) (pc:node): option s_state := + SOME path <- (fn_path f)!pc IN + SOME st <- symb_isteps path f (init_istate pc) IN + SOME i <- (fn_code f)!(st.(the_pc)) IN + Some (match symb_istep i st with + | Some st' => {| internal := st'; final := Snone |} + | None => {| internal := st; final := symb_fstep i st.(curr) |} + end). -- cgit From 3735ae5dea6f4876969eeaf7d4c5fc1d352bc68d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 6 Nov 2019 13:20:34 +0100 Subject: first steps in the proof of SE correctness --- mppa_k1c/lib/RTLpathSE_theory.v | 79 +++++++++++++++++++++++++++++++++++++++-- 1 file changed, 76 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index efde229d..c8f83a5c 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -392,7 +392,6 @@ Proof. - intros; erewrite symb_istep_WF; eauto. Qed. -(* remark: unused lemma ? *) Lemma symb_isteps_default_succ path f st': forall st, symb_isteps path f st = Some st' -> nth_default_succ (fn_code f) path st.(the_pc) = Some st'.(the_pc). Proof. @@ -436,7 +435,7 @@ Proof. - intros His1;rewrite His1; simpl; auto. Qed. -(** Auxilary properties, that are not useful here ? *) +(** REM: in the following two unused lemmas *) Lemma symb_isteps_right_assoc_decompose f path: forall st st', symb_isteps (S path) f st = Some st' -> @@ -551,12 +550,39 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := | _ => Snone end. +Lemma symb_fstep_correct i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: + (fn_code f) ! pc = Some i -> + pc = st.(the_pc) -> + sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> + path_last_step ge pge stack f sp pc rs m t s -> + symb_istep i st = None -> + sfval_sem (Build_sgenv pge ge sp) st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. +Proof. + intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + + (* Snone *) destruct i; simpl in Hi |- *; unfold sist in Hi; try congruence. + + (* Icall *) intros; eapply exec_Scall; auto. + - destruct ros; simpl in * |- *; auto. + rewrite REG; auto. + - erewrite list_sval_eval_inj; simpl; auto. + + (* Itaillcall *) admit. + + (* Ibuiltin *) admit. + + (* Ijumptable *) admit. + + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto. + destruct or; simpl; auto. +Admitted. + (** * Main function of the symbolic execution *) + Definition init_local_istate := {| pre:= fun _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. Definition init_istate pc := {| the_pc:= pc; exits:=nil; curr:= init_local_istate |}. -Fixpoint symb_step (f: function) (pc:node): option s_state := +Lemma init_sem_istate sge pc rs m: sem_istate sge (init_istate pc) rs m (ist true pc rs m). +Proof. + unfold sem_istate, sem_local_istate, has_not_yet_exit; simpl. intuition. +Qed. + +Definition symb_step (f: function) (pc:node): option s_state := SOME path <- (fn_path f)!pc IN SOME st <- symb_isteps path f (init_istate pc) IN SOME i <- (fn_code f)!(st.(the_pc)) IN @@ -564,3 +590,50 @@ Fixpoint symb_step (f: function) (pc:node): option s_state := | Some st' => {| internal := st'; final := Snone |} | None => {| internal := st; final := symb_fstep i st.(curr) |} end). + +Lemma final_node_path_simpl f path pc: + (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path pc <> None. +Proof. + intros; exploit final_node_path; eauto. + intros (i & NTH & DUM). + congruence. +Qed. + +Theorem symb_step_correct f pc pge ge sp path stack rs m t s: + (fn_path f)!pc = Some path -> + path_step ge pge path stack f sp rs m pc t s -> + exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge sp) st stack f rs m t s. +Proof. + Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. + intros PATH STEP; unfold symb_step; rewrite PATH; simpl. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist 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 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 sem_option2_istate; destruct (symb_isteps _ _ _) as [st|] eqn: Hst; try congruence); + (* intro SEM *) + (simpl; unfold sem_istate; simpl; rewrite CONT; intro SEM); + (* intro Hi' *) + ( assert (Hi': (fn_code f) ! (the_pc st) = Some i); + [ unfold nth_default_succ_inst in Hi; + exploit symb_isteps_default_succ; eauto; simpl; + intros DEF; rewrite DEF in Hi; auto + | clear Hi; rewrite Hi' ]); + (* eexists *) + (eexists; constructor; eauto). + - (* early exit *) + eapply sem_have_early_exit; eauto. + unfold sem_istate; simpl; rewrite CONT. + destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl; eauto. + - destruct SEM as (SEM & PC & HNYE). + destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl. + + (* normal exit on Snone *) admit. + + (* normal exit on non-Snone instruction *) + eapply sem_normal_exit; eauto. + * unfold sem_istate; simpl; rewrite CONT; intuition. + * simpl. eapply symb_fstep_correct; eauto. + rewrite PC; auto. +Admitted. + -- cgit From e48cfd6f21c0cb5ff9ae3a3c0ca339a6046a6eb4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 6 Nov 2019 16:21:07 +0100 Subject: Completeness is OK. But Correctness is a mess ! --- mppa_k1c/lib/RTLpathSE_theory.v | 98 ++++++++++++++++++++++++++++++++++++----- 1 file changed, 87 insertions(+), 11 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index c8f83a5c..63a76e8b 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -526,11 +526,11 @@ Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset Record s_state := { internal:> s_istate; final: sfval }. Inductive sem_state (sge: sgenv) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := - | sem_have_early_exit is: + | sem_early is: is.(continue) = false -> sem_istate sge st rs0 m0 is -> sem_state sge st stack f rs0 m0 E0 (State stack f sge.(the_sp) is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) - | sem_normal_exit is t s: + | sem_normal is t s: is.(continue) = true -> sem_istate sge st rs0 m0 is -> sfval_sem sge st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> @@ -550,7 +550,7 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := | _ => Snone end. -Lemma symb_fstep_correct i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: +Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> @@ -599,7 +599,68 @@ Proof. congruence. Qed. -Theorem symb_step_correct f pc pge ge sp path stack rs m t s: +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.(the_pc) -> + symb_istep i st = Some st' -> + path_last_step ge pge stack f sp pc rs m t s -> + exists ist, istep ge i sp rs m = Some ist /\ t = E0 /\ s = (State stack f sp ist.(RTLpath.the_pc) ist.(RTLpath.the_rs) ist.(the_mem)). +Proof. + intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl. +Qed. + +Lemma sem_istate_exclude_abort sge st rs m is: + has_aborted sge st rs m -> + sem_istate sge st rs m is -> False. +Proof. + (* DEFINITIONS PAS ASSEZ FORTE. A REVOIR! *) + intros [[AB|[AB|AB]]|AB]; + ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; + [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). + - (* ISSUE: PRE MONONOTIC !!! *) admit. +Admitted. + +(* +Lemma sem_istate_unicity sge st rs m ois is: + sem_option_istate sge st rs m ois -> + sem_istate sge st rs m is -> + ois = Some is. +Proof. + destruct ois as [is1|]; simpl. 2: { intros; exploit sem_istate_exclude_abort; eauto. destruct 1. } + unfold sem_istate. +Abort. +*) + +(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution *) +Theorem symb_step_correct f pc pge ge sp path stack st rs m t s: + (fn_path f)!pc = Some path -> + symb_step f pc = Some st -> + sem_state (Build_sgenv pge ge sp) st stack f rs m t s -> + path_step ge pge path stack f sp rs m pc t s. +Proof. + Local Hint Resolve init_sem_istate. + unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). + unfold nth_default_succ_inst in Hi. + destruct (symb_isteps path f (init_istate pc)) as [st0|] eqn: Hst0; simpl. + 2:{ (* absurd case *) + exploit symb_isteps_WF; eauto. + simpl; intros NDS; rewrite NDS in Hi; congruence. } + exploit symb_isteps_default_succ; eauto; simpl. + intros NDS; rewrite NDS in Hi. + rewrite Hi in SSTEP. + intros ISTEPS. try_simplify_someHyps. + destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. + + admit. + + destruct SEM as [is CONT SEM|]; simpl in * |- *. + - intros. eapply exec_early_exit; simpl; eauto. +Abort. + + +Theorem symb_step_complete f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> path_step ge pge path stack f sp rs m pc t s -> exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge sp) st stack f rs m t s. @@ -623,17 +684,32 @@ Proof. | clear Hi; rewrite Hi' ]); (* eexists *) (eexists; constructor; eauto). - - (* early exit *) - eapply sem_have_early_exit; eauto. + - (* early *) + eapply sem_early; eauto. unfold sem_istate; simpl; rewrite CONT. destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl; eauto. - destruct SEM as (SEM & PC & HNYE). destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl. - + (* normal exit on Snone *) admit. - + (* normal exit on non-Snone instruction *) - eapply sem_normal_exit; eauto. + + (* normal on Snone *) + rewrite <- PC in LAST. + exploit symb_path_last_step; eauto; simpl. + intros (ist & ISTEP & Ht & Hs); subst. + exploit symb_istep_correct; eauto. simpl. + erewrite Hst', ISTEP; simpl. + clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i. + intro SEM; destruct (ist.(continue)) eqn: CONT. + { (* continue ist = true *) + eapply sem_normal; simpl; eauto. + unfold sem_istate in SEM. + rewrite CONT in SEM. + destruct SEM as (SEM & PC & HNYE). + rewrite <- PC. + eapply exec_Snone. } + { eapply sem_early; eauto. } + + (* normal non-Snone instruction *) + eapply sem_normal; eauto. * unfold sem_istate; simpl; rewrite CONT; intuition. - * simpl. eapply symb_fstep_correct; eauto. + * simpl. eapply symb_fstep_complete; eauto. rewrite PC; auto. -Admitted. +Qed. -- cgit From 84f3168274378088f69d09467a6bb704d8160b6e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 7 Nov 2019 08:44:05 +0100 Subject: Paving the way for exactness proof... Try also to fix the confusion of terminology: - Correctness proof is OK (modulo remaining instructions to handle) - Exactness proof is in progress. --- mppa_k1c/lib/RTLpathSE_theory.v | 297 +++++++++++++++++++++++++++++++++------- 1 file changed, 245 insertions(+), 52 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 63a76e8b..4f7f3a6e 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -137,13 +137,14 @@ Definition has_aborted (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem): Prop has_aborted_on_basic sge st.(curr) rs0 m0 \/ has_aborted_on_test sge st.(exits) rs0 m0. - +(* TODO: est-ce vraiment utile ? *) Definition sem_option_istate sge (st: s_istate) rs0 m0 (ois: option istate): Prop := match ois with | Some is => sem_istate sge st rs0 m0 is | None => has_aborted sge st rs0 m0 end. +(* TODO: est-ce vraiment utile ? *) Definition sem_option2_istate sge (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := match ost with | Some st => sem_option_istate sge st rs0 m0 ois @@ -609,58 +610,10 @@ Proof. intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl. Qed. -Lemma sem_istate_exclude_abort sge st rs m is: - has_aborted sge st rs m -> - sem_istate sge st rs m is -> False. -Proof. - (* DEFINITIONS PAS ASSEZ FORTE. A REVOIR! *) - intros [[AB|[AB|AB]]|AB]; - ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; - [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). - - (* ISSUE: PRE MONONOTIC !!! *) admit. -Admitted. - -(* -Lemma sem_istate_unicity sge st rs m ois is: - sem_option_istate sge st rs m ois -> - sem_istate sge st rs m is -> - ois = Some is. -Proof. - destruct ois as [is1|]; simpl. 2: { intros; exploit sem_istate_exclude_abort; eauto. destruct 1. } - unfold sem_istate. -Abort. +(* NB: each concrete execution can be executed on the symbolic state (produced from [symb_step]) +(symb_step is a correct over-approximation) *) - -(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution *) -Theorem symb_step_correct f pc pge ge sp path stack st rs m t s: - (fn_path f)!pc = Some path -> - symb_step f pc = Some st -> - sem_state (Build_sgenv pge ge sp) st stack f rs m t s -> - path_step ge pge path stack f sp rs m pc t s. -Proof. - Local Hint Resolve init_sem_istate. - unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. - lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). - unfold nth_default_succ_inst in Hi. - destruct (symb_isteps path f (init_istate pc)) as [st0|] eqn: Hst0; simpl. - 2:{ (* absurd case *) - exploit symb_isteps_WF; eauto. - simpl; intros NDS; rewrite NDS in Hi; congruence. } - exploit symb_isteps_default_succ; eauto; simpl. - intros NDS; rewrite NDS in Hi. - rewrite Hi in SSTEP. - intros ISTEPS. try_simplify_someHyps. - destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. - + admit. - + destruct SEM as [is CONT SEM|]; simpl in * |- *. - - intros. eapply exec_early_exit; simpl; eauto. -Abort. - - -Theorem symb_step_complete f pc pge ge sp path stack rs m t s: +Theorem symb_step_correct f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> path_step ge pge path stack f sp rs m pc t s -> exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge sp) st stack f rs m t s. @@ -713,3 +666,243 @@ Proof. rewrite PC; auto. Qed. + +(** EXPERIMENTS: version 1. sem_istate est une fonction partielle... *) + +Definition istate_eq ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + +Definition istate_eq_option ist1 oist := + exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. + +Lemma sem_istate_exclude_abort sge st rs m is: + has_aborted sge st rs m -> + sem_istate sge st rs m is -> False. +Proof. + (* DEFINITIONS PAS ASSEZ FORTE. A REVOIR! *) + intros [[AB|[AB|AB]]|AB]; + ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; + [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). + - (* ISSUE: PRE MONONOTIC !!! *) admit. +Admitted. + + +Lemma has_not_yet_exit_cons_split sge ext lx rs0 m0: + has_not_yet_exit sge (ext::lx) rs0 m0 <-> + (seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge lx rs0 m0). +Proof. + unfold has_not_yet_exit; simpl; intuition (subst; auto). +Qed. + +Lemma exclude_early_NYE sge st rs0 m0 rs m pc: + sem_early_exit sge (exits st) rs0 m0 rs m pc -> + has_not_yet_exit sge (exits st) rs0 m0 -> + False. +Proof. + destruct 1. subst. +Admitted. + +Lemma sem_istate_exclude_incompatible_continue sge st rs m is1 is2: + is1.(continue) = true -> + is2.(continue) = false -> + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + False. +Proof. + Local Hint Resolve exclude_early_NYE. + unfold sem_istate. + intros CONT1 CONT2. + rewrite CONT1, CONT2; simpl. + intuition eauto. +Qed. + +Lemma sem_istate_determ_continue sge st rs m is1 is2: + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + is1.(continue) = is2.(continue). +Proof. + Local Hint Resolve sem_istate_exclude_incompatible_continue. + destruct (Bool.bool_dec is1.(continue) is2.(continue)) as [|H]; auto. + intros H1 H2. assert (absurd: False); intuition. + destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. +Qed. + +Lemma sem_local_istate_determ sge st rs0 m0 rs1 m1 rs2 m2: + sem_local_istate sge st rs0 m0 rs1 m1 -> + sem_local_istate sge st rs0 m0 rs2 m2 -> + (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + unfold sem_local_istate. intuition try congruence. + generalize (H5 r); rewrite H4; congruence. +Qed. + +Lemma sem_istate_determ sge st rs m is1 is2: + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + istate_eq is1 is2. +Proof. + unfold istate_eq. + intros SEM1 SEM2. + exploit (sem_istate_determ_continue sge st rs m is1 is2); eauto. + intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. + destruct (continue is2). + - generalize (sem_local_istate_determ sge (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); + intuition (try congruence). + - admit. +Admitted. + +Lemma sem_option_istate_determ sge st rs m ois is: + sem_option_istate sge st rs m ois -> + sem_istate sge st rs m is -> + istate_eq_option is ois. +Proof. + destruct ois as [is1|]; simpl; eauto. + - intros; eexists; intuition; eapply sem_istate_determ; eauto. + - intros; exploit sem_istate_exclude_abort; eauto. destruct 1. +Qed. + + +(** EXPERIMENTS: version 2. On prouve le déterminisme sur l'execution symbolique. + A priori, si la version 1 fonctionne, la version 2 aussi, mais sera plus lourde !!! + +Il faut mieux essayer d'abord la version 1.... + +*) + + +(* + +Lemma symb_istep_exact sge i st rs0 m0 rs m st' is: + is.(continue) = true -> + sem_local_istate sge st.(curr) rs0 m0 rs m -> + sem_istate sge st' rs0 m0 is -> + symb_istep i st = Some st' -> + istate_eq_option is (istep sge.(the_ge) i sge.(the_sp) rs m). +Proof. + unfold sem_istate, istate_eq_option, istate_eq; destruct is as [c pc rs1 m1]; simpl. + intro; subst. + intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC & NYE); subst. + destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. + + (* Nop *) + intros MEM; eexists; intuition eauto; simpl. + generalize (REG1 r); rewrite REG; try congruence. + + (* Op *) + intros MEM. + generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. + erewrite list_sval_eval_inj; simpl; auto. + rewrite MEM. + try_simplify_someHyps. + eexists; intuition eauto; simpl. + generalize (REG1 r0). + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. + subst; rewrite Regmap.gss. auto. + + (* LOAD *) + intros MEM. + generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. + erewrite list_sval_eval_inj; simpl; auto. + rewrite MEM. + inversion_SOME addr. + intros EVAL LOAD. rewrite EVAL, LOAD; simpl. + eexists; intuition eauto; simpl. + generalize (REG1 r0). + destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. + subst; rewrite Regmap.gss. auto. + + (* STORE *) + intros MEM. + generalize (REG r); destruct (Pos.eq_dec r r); try congruence; simpl. + intros Hr; rewrite Hr; simpl. + erewrite list_sval_eval_inj; simpl; auto. + inversion_SOME addr. + intros EVAL STORE. rewrite EVAL, STORE; simpl. + eexists; intuition eauto; simpl. + generalize (REG r0); rewrite REG1. congruence. + + (* COND *) + intros MEM. rewrite has_not_yet_exit_cons_split in NYE. + simpl in NYE. + destruct NYE as (COND & NYE). + unfold seval_condition in COND. + erewrite list_sval_eval_inj in COND; simpl; auto. + rewrite MEM in COND. + rewrite COND. eexists; intuition eauto. simpl. + generalize (REG r); rewrite REG1. congruence. +Qed. + +(* A REVOIR: essayer en right_assoc decompose ? *) +Lemma symb_isteps_exact sge path (f:function) rs0 m0 st' is': forall st is, + nth_default_succ (fn_code f) path st.(the_pc) <> None -> + is.(continue)=true -> + sem_istate sge st rs0 m0 is -> + symb_isteps path f st = Some st' -> + sem_istate sge st' rs0 m0 is' -> + istate_eq_option is' (isteps sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)). +Proof. + Local Hint Resolve symb_isteps_correct_false. + induction path. + + unfold sem_istate, istate_eq_option, istate_eq; simpl. + intros st is WF CONT; rewrite CONT. + intros ((PRE & EVAL & REG) & H & NYE). + intros Hst'; inversion Hst'. subst. + destruct (continue is') eqn: CONT'. + * simpl; intros ((PRE1 & EVAL1 & REG1) & PC & H1). + eexists; intuition eauto; simpl; try congruence. + generalize (REG1 r); rewrite REG. congruence. + * intros [ext lx TAIL NYE1 COND INV PC]. + exploit NYE. { eapply is_tail_in; eauto. } + congruence. + + simpl. unfold sem_istate at 1. + intros st is WF CONT; rewrite CONT. + intros (LOCAL & PC & NYE). + rewrite <- PC. + inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. + exploit symb_istep_correct; eauto. + inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. + intros SEM Hst' SEMis'. + inversion_SOME is0; intros His0; rewrite His0 in SEM; simpl in SEM. + * destruct (continue is0) eqn:CONT0; eauto. + - (* continue is0 = true *) + intros; eapply IHpath; simpl; eauto. + { intros; erewrite symb_isteps_default_succ; eauto. congruence. } + - (* continue is0 = false -> EARLY EXIT *) + admit. + * exploit sem_istate_exclude_abort; try tauto. + { eapply symb_isteps_preserv_has_aborted; eauto. } + { eauto. } +Abort. +*) +(** EXPERIMENTS: sketch du theorem final à prouver *) + +(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution +(symb_step is exact). + *) + +Definition match_state (s1 s2: state): Prop := s1 = s2. (* TODO: a generaliser ! *) + +Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: + (fn_path f)!pc = Some path -> + symb_step f pc = Some st -> + sem_state (Build_sgenv pge ge sp) st stack f rs m t s1 -> + exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ match_state s1 s2. +Proof. + Local Hint Resolve init_sem_istate. + unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). + unfold nth_default_succ_inst in Hi. + destruct (symb_isteps path f (init_istate pc)) as [st0|] eqn: Hst0; simpl. + 2:{ (* absurd case *) + exploit symb_isteps_WF; eauto. + simpl; intros NDS; rewrite NDS in Hi; congruence. } + exploit symb_isteps_default_succ; eauto; simpl. + intros NDS; rewrite NDS in Hi. + rewrite Hi in SSTEP. + intros ISTEPS. try_simplify_someHyps. + destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. + + admit. + + destruct SEM as [is CONT SEM|]; simpl in * |- *. + - intros. (* eapply exec_early_exit; simpl; eauto. *) +Abort. -- cgit From 51b6695f5fd765d68257f6e63bbf43db802783f2 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 12 Nov 2019 17:57:01 +0100 Subject: good sketch of the exact proof. --- mppa_k1c/lib/RTLpathSE_theory.v | 444 +++++++++++++++++++++------------------- 1 file changed, 229 insertions(+), 215 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 4f7f3a6e..cc2d2fde 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -151,6 +151,150 @@ Definition sem_option2_istate sge (ost: option s_istate) rs0 m0 (ois: option ist | None => ois=None end. +(** * An internal state represents a parallel program ! + + We prove below that the semantics [sem_option_istate] is deterministic. + + *) + +Definition istate_eq ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + +(* TODO: is it useful +Lemma has_not_yet_exit_cons_split sge ext lx rs0 m0: + has_not_yet_exit sge (ext::lx) rs0 m0 <-> + (seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge lx rs0 m0). +Proof. + unfold has_not_yet_exit; simpl; intuition (subst; auto). +Qed. +*) + +Lemma exclude_early_NYE sge st rs0 m0 rs m pc: + sem_early_exit sge (exits st) rs0 m0 rs m pc -> + has_not_yet_exit sge (exits st) rs0 m0 -> + False. +Proof. + Local Hint Resolve is_tail_in. + destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. + unfold has_not_yet_exit; intros NYE; rewrite NYE in EVAL; eauto. + try congruence. +Qed. + +Lemma sem_istate_exclude_incompatible_continue sge st rs m is1 is2: + is1.(continue) = true -> + is2.(continue) = false -> + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + False. +Proof. + Local Hint Resolve exclude_early_NYE. + unfold sem_istate. + intros CONT1 CONT2. + rewrite CONT1, CONT2; simpl. + intuition eauto. +Qed. + +Lemma sem_istate_determ_continue sge st rs m is1 is2: + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + is1.(continue) = is2.(continue). +Proof. + Local Hint Resolve sem_istate_exclude_incompatible_continue. + destruct (Bool.bool_dec is1.(continue) is2.(continue)) as [|H]; auto. + intros H1 H2. assert (absurd: False); intuition. + destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. +Qed. + +Lemma sem_local_istate_determ sge st rs0 m0 rs1 m1 rs2 m2: + sem_local_istate sge st rs0 m0 rs1 m1 -> + sem_local_istate sge st rs0 m0 rs2 m2 -> + (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + unfold sem_local_istate. 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. + induction 1 as [|i l1 l3 T1 IND]; simpl; auto. + intros T2; inversion T2; subst; auto. +Qed. + +Lemma exit_cond_determ sge rs0 m0 l1 l2: + is_tail l1 l2 -> forall ext1 lx1 ext2 lx2, + l1=(ext1 :: lx1) -> + l2=(ext2 :: lx2) -> + has_not_yet_exit sge lx1 rs0 m0 -> + seval_condition sge (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> + has_not_yet_exit sge 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. + unfold has_not_yet_exit in NYE. + rewrite NYE in EVAL; eauto. + try congruence. +Qed. + +Lemma sem_early_exit_determ sge lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : + sem_early_exit sge lx rs0 m0 rs1 m1 pc1 -> + sem_early_exit sge lx rs0 m0 rs2 m2 pc2 -> + pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. +Proof. + Local Hint Resolve exit_cond_determ eq_sym. + destruct 1 as [ext1 lx1 TAIL1 NYE1 EVAL1 SEM1 PC1]; subst. + destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. + assert (X:ext1=ext2). + - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. + - subst. destruct (sem_local_istate_determ sge (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. +Qed. + +Lemma sem_istate_determ sge st rs m is1 is2: + sem_istate sge st rs m is1 -> + sem_istate sge st rs m is2 -> + istate_eq is1 is2. +Proof. + unfold istate_eq. + intros SEM1 SEM2. + exploit (sem_istate_determ_continue sge st rs m is1 is2); eauto. + intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. + destruct (continue is2). + - destruct (sem_local_istate_determ sge (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); + intuition (try congruence). + - destruct (sem_early_exit_determ sge (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); + intuition. +Qed. + +Lemma sem_istate_exclude_abort sge st rs m is: + has_aborted sge st rs m -> + sem_istate sge st rs m is -> False. +Proof. + intros [[AB|[AB|AB]]|AB]; + ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; + [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). + (* TODO: découper cette grosse preuve à faire en lemmes *) +Admitted. + +Definition istate_eq_option ist1 oist := + exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. + +Lemma sem_option_istate_determ sge st rs m ois is: + sem_option_istate sge st rs m ois -> + sem_istate sge st rs m is -> + istate_eq_option is ois. +Proof. + destruct ois as [is1|]; simpl; eauto. + - intros; eexists; intuition; eapply sem_istate_determ; eauto. + - intros; exploit sem_istate_exclude_abort; eauto. destruct 1. +Qed. + (** * Symbolic execution of one internal step *) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := @@ -551,7 +695,7 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := | _ => Snone end. -Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: +Lemma symb_fstep_correct i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> @@ -572,6 +716,32 @@ Proof. destruct or; simpl; auto. Admitted. +Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: + (fn_code f) ! pc = Some i -> + pc = st.(the_pc) -> + sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> + sfval_sem (Build_sgenv pge ge sp) st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> + symb_istep i st = None -> + path_last_step ge pge stack f sp pc rs m t s. +Proof. + intros PC1 PC2 (PRE&MEM®) LAST Hi. + destruct i as [ | | | |(*Icall*)sig ros args res pc'| | | | |(*Ireturn*)or]; + subst; try_simplify_someHyps; try (unfold sist in Hi; try congruence); inversion LAST; subst; clear LAST; simpl in * |- *. + + (* Icall *) + erewrite list_sval_eval_inj in * |- ; simpl; try_simplify_someHyps; auto. + intros; eapply exec_Icall; eauto. + destruct ros; simpl in * |- *; auto. + rewrite REG in * |- ; auto. + + (* Itaillcall *) admit. + + (* Ibuiltin *) admit. + + (* Ijumptable *) admit. + + (* Ireturn *) + intros; subst. cutrewrite (v=regmap_optget or Vundef rs). + * eapply exec_Ireturn; eauto. + * intros; destruct or; simpl; congruence. +Admitted. + + (** * Main function of the symbolic execution *) Definition init_local_istate := {| pre:= fun _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. @@ -662,229 +832,49 @@ Proof. + (* normal non-Snone instruction *) eapply sem_normal; eauto. * unfold sem_istate; simpl; rewrite CONT; intuition. - * simpl. eapply symb_fstep_complete; eauto. + * simpl. eapply symb_fstep_correct; eauto. rewrite PC; auto. Qed. - -(** EXPERIMENTS: version 1. sem_istate est une fonction partielle... *) - -Definition istate_eq ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). - -Definition istate_eq_option ist1 oist := - exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. - -Lemma sem_istate_exclude_abort sge st rs m is: - has_aborted sge st rs m -> - sem_istate sge st rs m is -> False. -Proof. - (* DEFINITIONS PAS ASSEZ FORTE. A REVOIR! *) - intros [[AB|[AB|AB]]|AB]; - ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; - [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). - - (* ISSUE: PRE MONONOTIC !!! *) admit. -Admitted. - - -Lemma has_not_yet_exit_cons_split sge ext lx rs0 m0: - has_not_yet_exit sge (ext::lx) rs0 m0 <-> - (seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge lx rs0 m0). -Proof. - unfold has_not_yet_exit; simpl; intuition (subst; auto). -Qed. - -Lemma exclude_early_NYE sge st rs0 m0 rs m pc: - sem_early_exit sge (exits st) rs0 m0 rs m pc -> - has_not_yet_exit sge (exits st) rs0 m0 -> - False. -Proof. - destruct 1. subst. -Admitted. - -Lemma sem_istate_exclude_incompatible_continue sge st rs m is1 is2: - is1.(continue) = true -> - is2.(continue) = false -> - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> - False. -Proof. - Local Hint Resolve exclude_early_NYE. - unfold sem_istate. - intros CONT1 CONT2. - rewrite CONT1, CONT2; simpl. - intuition eauto. -Qed. - -Lemma sem_istate_determ_continue sge st rs m is1 is2: - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> - is1.(continue) = is2.(continue). -Proof. - Local Hint Resolve sem_istate_exclude_incompatible_continue. - destruct (Bool.bool_dec is1.(continue) is2.(continue)) as [|H]; auto. - intros H1 H2. assert (absurd: False); intuition. - destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. -Qed. - -Lemma sem_local_istate_determ sge st rs0 m0 rs1 m1 rs2 m2: - sem_local_istate sge st rs0 m0 rs1 m1 -> - sem_local_istate sge st rs0 m0 rs2 m2 -> - (forall r, rs1#r = rs2#r) /\ m1 = m2. -Proof. - unfold sem_local_istate. intuition try congruence. - generalize (H5 r); rewrite H4; congruence. -Qed. - -Lemma sem_istate_determ sge st rs m is1 is2: - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> - istate_eq is1 is2. -Proof. - unfold istate_eq. - intros SEM1 SEM2. - exploit (sem_istate_determ_continue sge st rs m is1 is2); eauto. - intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. - destruct (continue is2). - - generalize (sem_local_istate_determ sge (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); - intuition (try congruence). - - admit. -Admitted. - -Lemma sem_option_istate_determ sge st rs m ois is: - sem_option_istate sge st rs m ois -> - sem_istate sge st rs m is -> - istate_eq_option is ois. -Proof. - destruct ois as [is1|]; simpl; eauto. - - intros; eexists; intuition; eapply sem_istate_determ; eauto. - - intros; exploit sem_istate_exclude_abort; eauto. destruct 1. -Qed. - - -(** EXPERIMENTS: version 2. On prouve le déterminisme sur l'execution symbolique. - A priori, si la version 1 fonctionne, la version 2 aussi, mais sera plus lourde !!! - -Il faut mieux essayer d'abord la version 1.... - +(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution + (symb_step is exact). *) -(* - -Lemma symb_istep_exact sge i st rs0 m0 rs m st' is: - is.(continue) = true -> - sem_local_istate sge st.(curr) rs0 m0 rs m -> - sem_istate sge st' rs0 m0 is -> - symb_istep i st = Some st' -> - istate_eq_option is (istep sge.(the_ge) i sge.(the_sp) rs m). -Proof. - unfold sem_istate, istate_eq_option, istate_eq; destruct is as [c pc rs1 m1]; simpl. - intro; subst. - intros (PRE & MEM & REG) ((PRE1 & MEM1 & REG1) & PC & NYE); subst. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl in * |- *; try_simplify_someHyps. - + (* Nop *) - intros MEM; eexists; intuition eauto; simpl. - generalize (REG1 r); rewrite REG; try congruence. - + (* Op *) - intros MEM. - generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. - erewrite list_sval_eval_inj; simpl; auto. - rewrite MEM. - try_simplify_someHyps. - eexists; intuition eauto; simpl. - generalize (REG1 r0). - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. - subst; rewrite Regmap.gss. auto. - + (* LOAD *) - intros MEM. - generalize (REG1 r); destruct (Pos.eq_dec r r); try congruence; simpl. - erewrite list_sval_eval_inj; simpl; auto. - rewrite MEM. - inversion_SOME addr. - intros EVAL LOAD. rewrite EVAL, LOAD; simpl. - eexists; intuition eauto; simpl. - generalize (REG1 r0). - destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; congruence]. - subst; rewrite Regmap.gss. auto. - + (* STORE *) - intros MEM. - generalize (REG r); destruct (Pos.eq_dec r r); try congruence; simpl. - intros Hr; rewrite Hr; simpl. - erewrite list_sval_eval_inj; simpl; auto. - inversion_SOME addr. - intros EVAL STORE. rewrite EVAL, STORE; simpl. - eexists; intuition eauto; simpl. - generalize (REG r0); rewrite REG1. congruence. - + (* COND *) - intros MEM. rewrite has_not_yet_exit_cons_split in NYE. - simpl in NYE. - destruct NYE as (COND & NYE). - unfold seval_condition in COND. - erewrite list_sval_eval_inj in COND; simpl; auto. - rewrite MEM in COND. - rewrite COND. eexists; intuition eauto. simpl. - generalize (REG r); rewrite REG1. congruence. +Inductive smatch_state: state -> state -> Prop := + | smatch_state_refl s: smatch_state s s + | State_smatch stack f sp pc rs1 m rs2: + (forall r, rs1#r = rs2#r) -> + smatch_state (State stack f sp pc rs1 m) + (State stack f sp pc rs2 m) + | Callstate_smatch stack res f sp pc rs1 m rs2 fd args: + (forall r, rs1#r = rs2#r) -> + smatch_state (Callstate (Stackframe res f sp pc rs1 :: stack) fd args m) + (Callstate (Stackframe res f sp pc rs2 :: stack) fd args m). + +Lemma sfval_sem_smatch sge (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + sfval_sem sge st stack f rs0 m0 sv rs1 m t s -> + (forall r, rs1#r = rs2#r) -> + exists s', smatch_state s s' /\ sfval_sem sge st stack f rs0 m0 sv rs2 m t s'. +Proof. + destruct 1. + - (* Snone *) intros; eexists; econstructor. + + eapply State_smatch; eauto. + + eapply exec_Snone. + - (* Scall *) intros; eexists; econstructor. + + eapply Callstate_smatch; eauto. + + eapply exec_Scall; eauto. + - (* Sreturn *) intros; eexists; econstructor. + + eapply smatch_state_refl; eauto. + + eapply exec_Sreturn; eauto. Qed. -(* A REVOIR: essayer en right_assoc decompose ? *) -Lemma symb_isteps_exact sge path (f:function) rs0 m0 st' is': forall st is, - nth_default_succ (fn_code f) path st.(the_pc) <> None -> - is.(continue)=true -> - sem_istate sge st rs0 m0 is -> - symb_isteps path f st = Some st' -> - sem_istate sge st' rs0 m0 is' -> - istate_eq_option is' (isteps sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)). -Proof. - Local Hint Resolve symb_isteps_correct_false. - induction path. - + unfold sem_istate, istate_eq_option, istate_eq; simpl. - intros st is WF CONT; rewrite CONT. - intros ((PRE & EVAL & REG) & H & NYE). - intros Hst'; inversion Hst'. subst. - destruct (continue is') eqn: CONT'. - * simpl; intros ((PRE1 & EVAL1 & REG1) & PC & H1). - eexists; intuition eauto; simpl; try congruence. - generalize (REG1 r); rewrite REG. congruence. - * intros [ext lx TAIL NYE1 COND INV PC]. - exploit NYE. { eapply is_tail_in; eauto. } - congruence. - + simpl. unfold sem_istate at 1. - intros st is WF CONT; rewrite CONT. - intros (LOCAL & PC & NYE). - rewrite <- PC. - inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. - exploit symb_istep_correct; eauto. - inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - intros SEM Hst' SEMis'. - inversion_SOME is0; intros His0; rewrite His0 in SEM; simpl in SEM. - * destruct (continue is0) eqn:CONT0; eauto. - - (* continue is0 = true *) - intros; eapply IHpath; simpl; eauto. - { intros; erewrite symb_isteps_default_succ; eauto. congruence. } - - (* continue is0 = false -> EARLY EXIT *) - admit. - * exploit sem_istate_exclude_abort; try tauto. - { eapply symb_isteps_preserv_has_aborted; eauto. } - { eauto. } -Abort. -*) -(** EXPERIMENTS: sketch du theorem final à prouver *) - -(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution -(symb_step is exact). - *) - -Definition match_state (s1 s2: state): Prop := s1 = s2. (* TODO: a generaliser ! *) - Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> sem_state (Build_sgenv pge ge sp) st stack f rs m t s1 -> - exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ match_state s1 s2. + exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ + smatch_state s1 s2. Proof. Local Hint Resolve init_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. @@ -903,6 +893,30 @@ Proof. intros ISTEPS. try_simplify_someHyps. destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. + admit. - + destruct SEM as [is CONT SEM|]; simpl in * |- *. - - intros. (* eapply exec_early_exit; simpl; eauto. *) -Abort. + + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. + - (* early exit *) + intros. + exploit sem_option_istate_determ; eauto. + destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). + eexists. econstructor 1. + * eapply exec_early_exit; eauto. + * rewrite EQ2, EQ4; eapply State_smatch. auto. + - (* normal exit *) + intros. + exploit sem_option_istate_determ; eauto. + destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). + unfold sem_istate in SEM1. + rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). + exploit sfval_sem_smatch; 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 symb_fstep_complete; eauto. + * congruence. + * unfold sem_local_istate in * |- *. intuition try congruence. +Admitted. + + + -- cgit From 38e9b8f5723235f59543543bd274f57a57602695 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 14 Nov 2019 12:09:51 +0100 Subject: experiments a definition of "path_simu" --- configure | 2 +- mppa_k1c/lib/RTLpathSE_theory.v | 60 +++++-- mppa_k1c/lib/RTLpathScheduler.v | 390 +++++++++++++++++++++++++++++++--------- 3 files changed, 353 insertions(+), 99 deletions(-) diff --git a/configure b/configure index be6f2180..da20d0dc 100755 --- a/configure +++ b/configure @@ -849,7 +849,7 @@ ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure EXECUTE=k1-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __K1C_COS__ SIMU=k1-cluster -- -BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v\\ +BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathSE_theory.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index cc2d2fde..b7674916 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -841,31 +841,31 @@ Qed. *) -Inductive smatch_state: state -> state -> Prop := - | smatch_state_refl s: smatch_state s s - | State_smatch stack f sp pc rs1 m rs2: +Inductive equiv_state: state -> state -> Prop := + | equiv_state_refl s: equiv_state s s + | State_equiv stack f sp pc rs1 m rs2: (forall r, rs1#r = rs2#r) -> - smatch_state (State stack f sp pc rs1 m) + equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m) - | Callstate_smatch stack res f sp pc rs1 m rs2 fd args: + | Callstate_equiv stack res f sp pc rs1 m rs2 fd args: (forall r, rs1#r = rs2#r) -> - smatch_state (Callstate (Stackframe res f sp pc rs1 :: stack) fd args m) - (Callstate (Stackframe res f sp pc rs2 :: stack) fd args m). + equiv_state (Callstate (Stackframe res f sp pc rs1 :: stack) fd args m) + (Callstate (Stackframe res f sp pc rs2 :: stack) fd args m). -Lemma sfval_sem_smatch sge (f:function) st sv stack rs0 m0 t rs1 rs2 m s: +Lemma sfval_sem_equiv sge (f:function) st sv stack rs0 m0 t rs1 rs2 m s: sfval_sem sge st stack f rs0 m0 sv rs1 m t s -> (forall r, rs1#r = rs2#r) -> - exists s', smatch_state s s' /\ sfval_sem sge st stack f rs0 m0 sv rs2 m t s'. + exists s', equiv_state s s' /\ sfval_sem sge st stack f rs0 m0 sv rs2 m t s'. Proof. destruct 1. - (* Snone *) intros; eexists; econstructor. - + eapply State_smatch; eauto. + + eapply State_equiv; eauto. + eapply exec_Snone. - (* Scall *) intros; eexists; econstructor. - + eapply Callstate_smatch; eauto. + + eapply Callstate_equiv; eauto. + eapply exec_Scall; eauto. - (* Sreturn *) intros; eexists; econstructor. - + eapply smatch_state_refl; eauto. + + eapply equiv_state_refl; eauto. + eapply exec_Sreturn; eauto. Qed. @@ -874,7 +874,7 @@ Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: symb_step f pc = Some st -> sem_state (Build_sgenv pge ge sp) st stack f rs m t s1 -> exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ - smatch_state s1 s2. + equiv_state s1 s2. Proof. Local Hint Resolve init_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. @@ -892,7 +892,8 @@ Proof. rewrite Hi in SSTEP. intros ISTEPS. try_simplify_someHyps. destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. - + admit. + + (* normal exit on Snone instruction *) + admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. @@ -900,14 +901,14 @@ Proof. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). eexists. econstructor 1. * eapply exec_early_exit; eauto. - * rewrite EQ2, EQ4; eapply State_smatch. auto. - - (* normal exit *) + * rewrite EQ2, EQ4; eapply State_equiv. auto. + - (* normal exit non-Snone instruction *) intros. exploit sem_option_istate_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). unfold sem_istate in SEM1. rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). - exploit sfval_sem_smatch; eauto. + exploit sfval_sem_equiv; eauto. clear SEM2; destruct 1 as (s' & Ms' & SEM2). rewrite ! EQ4 in * |-; clear EQ4. rewrite ! EQ2 in * |-; clear EQ2. @@ -918,5 +919,30 @@ Proof. * unfold sem_local_istate in * |- *. intuition try congruence. Admitted. +(** * Simulation of RTLpath code w.r.t symbolic execution *) +Definition istate_simu (srce: node -> option node) ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + srce (ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + +Definition s_istate_simu (srce: node -> option node) (st1 st2: s_istate): Prop := + forall sge1 sge2 rs0 m0 is1 is2, istate_simu srce is1 is2 -> sem_istate sge1 st1 rs0 m0 is1 -> sem_istate sge2 st2 rs0 m0 is2. + +Inductive sfval_simu (srce: node -> option node): sfval -> sfval -> Prop := + | Snone_simu: sfval_simu srce Snone Snone + | Scall_simu sig svos lsv res pc1 pc2: + srce pc2 = Some pc1 -> + sfval_simu srce (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) + | Sreturn_simu os: + sfval_simu srce (Sreturn os) (Sreturn os). + +Definition s_state_simu (srce: node -> option node) (s1 s2: s_state): Prop := + s_istate_simu srce s1.(internal) s2.(internal) + /\ sfval_simu srce s1.(final) s2.(final). + +Definition path_simu (srce: node -> option node) (f1 f2: RTLpath.function) pc1 pc2: Prop := + forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. +(** See usage of [path_simu] in [RTLpathScheduler] *) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 76795fa4..c27d5317 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -1,11 +1,13 @@ (** RTLpath Scheduling from an external oracle. -This module is inspired from [Duplicate] +This module is inspired from [Duplicate] and [Duplicateproof] *) -Require Import AST Values RTL RTLpath Maps Globalenvs. -Require Import Coqlib Errors Op. +Require Import AST Linking Values Maps Globalenvs Smallstep. +Require Import Coqlib Maps Events Errors Op. +Require Import RTL RTLpath RTLpathSE_theory. + Local Open Scope error_monad_scope. Local Open Scope positive_scope. @@ -25,7 +27,7 @@ Extract Constant untrusted_scheduler => "TODO". *) Record xfunction : Type := - { fn_RTL: RTLpath.function; + { fn_RTL:> RTLpath.function; fn_revmap: PTree.t node; }. @@ -39,109 +41,335 @@ Definition fundef_RTL (fu: xfundef) : fundef := | External ef => External ef end. -(* auxiliary definitions that should be use in [verify_match_path] implementation below ! *) +(* Definition verify_is_copy revmap n n' := match revmap!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 revmap ln ln' := - match ln with - | n::ln => match ln' with - | n'::ln' => do u <- verify_is_copy revmap n n'; - verify_is_copy_list revmap 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. - Definition verify_mapping_entrypoint (f: function) (xf: xfunction): res unit := verify_is_copy (fn_revmap xf) (fn_entrypoint f) (fn_entrypoint (fn_RTL xf)). +*) +Record match_function (f1: RTLpath.function) (f2: xfunction): 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: f2.(fn_revmap)!(fn_entrypoint f2) = Some (fn_entrypoint f1); + pathentry_simu: + forall pc1 pc2, f2.(fn_revmap)!pc2 = Some pc1 -> path_entry (fn_code f1) (fn_path f1) pc1 -> path_entry (fn_code f2) (fn_path f2) pc2 /\ path_simu (fun n' => (fn_revmap f2)!n') f1 f2 pc1 pc2 +}. -(* This verifier remains to be implemented by symbolic execution *) -Parameter verify_match_path: function -> xfunction -> node -> node -> bool. +(* TODO: remove these two stub parameters *) +Parameter transf_function_aux: RTLpath.function -> res xfunction. -(* -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}. +Parameter transf_function_correct: forall f xf, transf_function_aux f = OK xf -> match_function f xf. + +Definition transf_function (f: function) : res function := + do xf <- transf_function_aux f; + OK (fn_RTL xf). + +Definition transf_fundef (f: fundef) : res fundef := + transf_partial_fundef transf_function f. + +Definition transf_program (p: program) : res program := + transform_partial_program transf_fundef p. + + +(** * Preservation proof *) + + +Inductive match_fundef: RTLpath.fundef -> RTLpath.fundef -> Prop := + | match_Internal f xf: match_function f xf -> match_fundef (Internal f) (Internal (fn_RTL xf)) + | match_External ef: match_fundef (External ef) (External ef). + + +Inductive match_stackframes: stackframe -> stackframe -> Prop := + | match_stackframe_intro: + forall res f sp pc rs xf pc' + (TRANSF: match_function f xf) + (DUPLIC: (fn_revmap xf)!pc' = Some pc), + match_stackframes (Stackframe res f sp pc rs) (Stackframe res (fn_RTL xf) sp pc' rs). + +Inductive match_states: state -> state -> Prop := + | match_states_intro: + forall st f sp pc rs m st' xf pc' + (STACKS: list_forall2 match_stackframes st st') + (TRANSF: match_function f xf) + (DUPLIC: (fn_revmap xf)!pc' = Some pc), + match_states (State st f sp pc rs m) (State st' (fn_RTL xf) sp pc' rs m) + | match_states_call: + forall 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: + forall st st' v m + (STACKS: list_forall2 match_stackframes st st'), + match_states (Returnstate st v m) (Returnstate st' v m). + + +Lemma transf_fundef_correct f f': + transf_fundef f = OK f' -> match_fundef f f'. Proof. - intros H H'. intros. decide equality. + intros TRANSF; destruct f; simpl; monadInv TRANSF. + + monadInv EQ. + eapply match_Internal; eapply transf_function_correct; eauto. + + eapply match_External. Qed. -(** 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}. +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. - apply (builtin_arg_eq Pos.eq_dec). -Defined. -Global Opaque builtin_arg_eq_pos. + intros. eapply match_transform_partial_program_contextual; eauto. +Qed. -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. -*) +Section PRESERVATION. -Fixpoint verify_mapping_mn_rec f xf lm := - match lm with - | nil => OK tt - | (pc2,pc1) :: lm => - if verify_match_path f xf pc1 pc2 - then - do u2 <- verify_mapping_mn_rec f xf lm; - OK tt - else - Error (msg "verify_match_path fails") - end. +Variable prog: program. +Variable tprog: program. -Definition verify_mapping_match_paths (f: function) (xf: xfunction) : res unit := - verify_mapping_mn_rec f xf (PTree.elements (fn_revmap xf)). +Hypothesis TRANSL: match_prog prog tprog. -(** Verifies that the [fn_revmap] of the translated function [xf] is giving correct information in regards to [f] *) -Definition verify_mapping (f: function) (xf: xfunction) : res unit := - do u <- verify_mapping_entrypoint f xf; - do v <- verify_mapping_match_paths f xf; OK tt. -(* TODO - verify the other axiom *) +Let pge := Genv.globalenv prog. +Let tpge := Genv.globalenv tprog. +Let ge := Genv.globalenv (program_RTL prog). +Let tge := Genv.globalenv (program_RTL tprog). -(** * Entry points *) +Lemma symbols_preserved_pge s: Genv.find_symbol tpge s = Genv.find_symbol pge s. +Proof. + rewrite <- (Genv.find_symbol_match TRANSL). reflexivity. +Qed. -(* TODO: à finir... +Lemma senv_sym x y: Senv.equiv x y -> Senv.equiv y x. +Proof. + unfold Senv.equiv. intuition congruence. +Qed. -Definition transf_function_aux (f: function) : res xfunction := - let (tcte, mp) := duplicate_aux f in - let (tc, te) := tcte in - let xf := {| fn_RTL := (mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te); fn_revmap := mp |} in - do u <- verify_mapping f xf; - OK xf. +Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z. +Proof. + unfold Senv.equiv. intuition congruence. +Qed. -Theorem transf_function_aux_preserves: - forall f xf, - transf_function_aux f = OK xf -> - fn_sig f = fn_sig (fn_RTL xf) /\ fn_params f = fn_params (fn_RTL xf) /\ fn_stacksize f = fn_stacksize (fn_RTL xf). +Lemma senv_preserved_pge: + Senv.equiv pge tpge. Proof. - intros. unfold transf_function_aux in H. destruct (duplicate_aux _) as (tcte & mp). destruct tcte as (tc & te). monadInv H. - repeat (split; try reflexivity). + eapply (Genv.senv_match TRANSL). Qed. -Remark transf_function_aux_fnsig: forall f xf, transf_function_aux f = OK xf -> fn_sig f = fn_sig (fn_RTL xf). - Proof. apply transf_function_aux_preserves. Qed. -Remark transf_function_aux_fnparams: forall f xf, transf_function_aux f = OK xf -> fn_params f = fn_params (fn_RTL xf). - Proof. apply transf_function_aux_preserves. Qed. -Remark transf_function_aux_fnstacksize: forall f xf, transf_function_aux f = OK xf -> fn_stacksize f = fn_stacksize (fn_RTL xf). - Proof. apply transf_function_aux_preserves. Qed. +Lemma senv_preserved: + Senv.equiv ge tge. +Proof. + eapply senv_transitivity. { eapply senv_sym; eapply RTLpath.senv_preserved. } + eapply senv_transitivity. { eapply senv_preserved_pge. } + eapply RTLpath.senv_preserved. +Qed. -Definition transf_function (f: function) : res function := - do xf <- transf_function_aux f; - OK (fn_RTL xf). -Definition transf_fundef (f: fundef) : res fundef := - transf_partial_fundef transf_function f. +Lemma functions_translated: + 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. -Definition transf_program (p: program) : res program := - transform_partial_program transf_fundef p. -*) \ No newline at end of file +Lemma function_ptr_translated_pge: + 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_ptr_translated: + forall v f, + Genv.find_funct_ptr ge v = Some f -> + exists f0 tf, + Genv.find_funct_ptr pge v = Some f0 /\ Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f0 = OK tf. +Proof. + intros; exploit find_funct_ptr_RTL_preserv; eauto. + intros (f0 & X & Y); subst. + exploit function_ptr_translated_pge; eauto. + intros (tf & Y & Z); eauto. +Qed. + +Lemma function_sig_translated: + forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. +Proof. + intros. destruct f. + - simpl in H. monadInv H. simpl. symmetry. monadInv EQ. + eapply preserv_fnsig. + eapply transf_function_correct. + assumption. + - simpl in H. monadInv H. (* monadInv EQ. *) reflexivity. +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_translated; eauto. intros (f0 & tf & FIND1 & FIND2 & TRANSF). + exists (RTL.Callstate nil (fn_RTL tf nil m0) nil m0). + eexists. split. + 2: { eapply match_states_call. eauto. + - unfold initial_state. econstructor. + + 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. + + exploit function_ptr_translated; eauto. + + destruct f. + * monadInv TRANSF. monadInv EQ. rewrite <- H3. symmetry; eapply transf_function_aux_preserves. assumption. + * monadInv TRANSF. (* monadInv EQ. *) assumption. + - constructor; eauto. constructor. apply transf_fundef_correct; auto. +Qed. + +(* TODO +Theorem transf_final_states: + forall 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. + +Theorem step_simulation: + forall s1 t s1', step pge s1 t s1' -> + forall s2 (MS: match_states s1 s2), + exists s2', + step tge s2 t s2' + /\ match_states s1' s2'. +Proof. + Local Hint Resolve transf_fundef_correct. + induction 1; intros; inv MS. +(* Inop *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). + inv H3. + eexists. split. + + eapply exec_Inop; eauto. + + constructor; eauto. +(* Iop *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + eexists. split. + + eapply exec_Iop; eauto. erewrite eval_operation_preserved; eauto. + + constructor; eauto. +(* Iload *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + eexists. split. + + eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto. + + constructor; auto. +(* Istore *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + eexists. split. + + eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto. + + constructor; auto. +(* Icall *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + destruct ros. + * simpl in H0. apply functions_translated in H0. + destruct H0 as (tf & cunit & TFUN & GFIND & LO). + eexists. split. + + eapply exec_Icall. eassumption. simpl. eassumption. + apply function_sig_translated. assumption. + + repeat (constructor; auto). + * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF). + eexists. split. + + eapply exec_Icall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS. + eassumption. apply function_sig_translated. assumption. + + repeat (constructor; auto). +(* Itailcall *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H10 & H11). inv H11. + pose symbols_preserved as SYMPRES. + destruct ros. + * simpl in H0. apply functions_translated in H0. + destruct H0 as (tf & cunit & TFUN & GFIND & LO). + eexists. split. + + eapply exec_Itailcall. eassumption. simpl. eassumption. + apply function_sig_translated. assumption. + erewrite <- preserv_fnstacksize; eauto. + + repeat (constructor; auto). + * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. + apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF). + eexists. split. + + eapply exec_Itailcall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS. + eassumption. apply function_sig_translated. assumption. + erewrite <- preserv_fnstacksize; eauto. + + repeat (constructor; auto). +(* Ibuiltin *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + 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. + + constructor; auto. +(* Icond *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + eexists. split. + + eapply exec_Icond; eauto. + + constructor; auto. destruct b; auto. +(* Ijumptable *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + exploit list_nth_z_revmap; eauto. intros (pc'1 & LNZ & REVM). + eexists. split. + + eapply exec_Ijumptable; eauto. + + constructor; auto. +(* Ireturn *) + - eapply revmap_correct in DUPLIC; eauto. + destruct DUPLIC as (i' & H2 & H3). inv H3. + pose symbols_preserved as SYMPRES. + eexists. split. + + eapply exec_Ireturn; eauto. erewrite <- preserv_fnstacksize; eauto. + + constructor; auto. +(* exec_function_internal *) + - inversion TRANSF as [f0 xf MATCHF|]; subst. eexists. split. + + eapply exec_function_internal. erewrite <- preserv_fnstacksize; eauto. + + erewrite preserv_fnparams; eauto. + econstructor; eauto. apply revmap_entrypoint. assumption. +(* exec_function_external *) + - inversion TRANSF as [|]; subst. eexists. split. + + 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. + + constructor. + + inv H1. constructor; assumption. +Qed. + +Theorem transf_program_correct: + forward_simulation (semantics prog) (semantics tprog). +Proof. + eapply forward_simulation_step with match_states. + - eapply senv_preserved. + - eapply transf_initial_states. + - eapply transf_final_states. + - eapply step_simulation. +Qed. +*) + +End PRESERVATION. -- cgit From 8de938a3e02b6f6332907f1d03fbeda374ecaa6e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 25 Nov 2019 13:55:06 +0100 Subject: progress on RTLpathScheduler --- mppa_k1c/lib/RTLpath.v | 20 ++++- mppa_k1c/lib/RTLpathScheduler.v | 161 +++++++++++++++++++++------------------- 2 files changed, 100 insertions(+), 81 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 95818c1d..e07822ca 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -291,8 +291,12 @@ Inductive step ge pge: state -> trace -> state -> Prop := step ge pge (Returnstate (Stackframe res f sp pc rs :: s) vres m) E0 (State s f sp pc (rs#res <- vres) m). -Definition initial_state (p: program) (st:state): Prop - := RTL.initial_state p st. +Inductive initial_state (p:program) : state -> Prop := + initial_state_intro (b : block) (f : fundef) (m0 : mem): + Genv.init_mem p = Some m0 -> + Genv.find_symbol (Genv.globalenv p) (prog_main p) = Some b -> + Genv.find_funct_ptr (Genv.globalenv p) b = Some f -> + funsig f = signature_main -> initial_state p (Callstate nil f nil m0). Definition final_state (st: state) (i:int): Prop := RTL.final_state st i. @@ -453,7 +457,12 @@ Theorem RTLpath_correct: forward_simulation (semantics p) (RTL.semantics p). Proof. eapply forward_simulation_plus with (match_states := fun s1 s2 => s2 = state_RTL s1); simpl; auto. - apply senv_preserved. - - unfold initial_state; eauto. + - destruct 1; intros; eexists; intuition eauto. econstructor; eauto. + + apply (Genv.init_mem_match match_prog_RTL); auto. + + rewrite (Genv.find_symbol_match match_prog_RTL). + rewrite (match_program_main match_prog_RTL); eauto. + + exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto. + intros (cunit & tf0 & XX); intuition subst; eauto. - unfold final_state; intros; subst; eauto. - intros; subst. eexists; intuition. eapply step_correct; eauto. @@ -959,9 +968,12 @@ Proof. eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states). constructor 1; simpl. - apply lt_wf. - - unfold initial_state, match_states, match_inst_states. destruct 1; simpl; exists O. + - unfold match_states, match_inst_states. destruct 1; simpl; exists O. destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto. exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto. + + apply (Genv.init_mem_match (match_RTL_prog p)); auto. + + rewrite (Genv.find_symbol_match (match_RTL_prog p)). + rewrite (match_program_main (match_RTL_prog p)); eauto. - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. destruct s2; simpl in * |- *; inversion H0; subst. constructor. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index c27d5317..dcdb3419 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -26,49 +26,30 @@ Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t Extract Constant untrusted_scheduler => "TODO". *) -Record xfunction : Type := - { fn_RTL:> RTLpath.function; - fn_revmap: PTree.t node; - }. - -Definition xfundef := AST.fundef xfunction. -Definition xprogram := AST.program xfundef unit. -Definition xgenv := Genv.t xfundef unit. - -Definition fundef_RTL (fu: xfundef) : fundef := - match fu with - | Internal f => Internal (fn_RTL f) - | External ef => External ef - end. - -(* -Definition verify_is_copy revmap n n' := - match revmap!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. - -Definition verify_mapping_entrypoint (f: function) (xf: xfunction): res unit := - verify_is_copy (fn_revmap xf) (fn_entrypoint f) (fn_entrypoint (fn_RTL xf)). -*) - -Record match_function (f1: RTLpath.function) (f2: xfunction): Prop := { +Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.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: f2.(fn_revmap)!(fn_entrypoint f2) = Some (fn_entrypoint f1); - pathentry_simu: - forall pc1 pc2, f2.(fn_revmap)!pc2 = Some pc1 -> path_entry (fn_code f1) (fn_path f1) pc1 -> path_entry (fn_code f2) (fn_path f2) pc2 /\ path_simu (fun n' => (fn_revmap f2)!n') f1 f2 pc1 pc2 + preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); + dupmap_correct: + forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_code f1) (fn_path f1) pc1 -> path_entry (fn_code f2) (fn_path f2) pc2 /\ path_simu (fun n' => dupmap!n') f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) -Parameter transf_function_aux: RTLpath.function -> res xfunction. +Parameter transf_function: RTLpath.function -> res RTLpath.function. -Parameter transf_function_correct: forall f xf, transf_function_aux f = OK xf -> match_function f xf. +Parameter transf_function_correct: + forall f f', transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. -Definition transf_function (f: function) : res function := - do xf <- transf_function_aux f; - OK (fn_RTL xf). +Theorem transf_function_preserves: + forall f f', + transf_function f = OK f' -> + fn_sig f = fn_sig f' /\ fn_params f = fn_params f' /\ fn_stacksize f = fn_stacksize f'. +Proof. + intros. exploit transf_function_correct; eauto. + destruct 1 as (dupmap & [SIG PARAM SIZE ENTRY CORRECT]). + intuition. +Qed. Definition transf_fundef (f: fundef) : res fundef := transf_partial_fundef transf_function f. @@ -79,26 +60,22 @@ Definition transf_program (p: program) : res program := (** * Preservation proof *) - Inductive match_fundef: RTLpath.fundef -> RTLpath.fundef -> Prop := - | match_Internal f xf: match_function f xf -> match_fundef (Internal f) (Internal (fn_RTL xf)) + | 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 -> stackframe -> Prop := - | match_stackframe_intro: - forall res f sp pc rs xf pc' - (TRANSF: match_function f xf) - (DUPLIC: (fn_revmap xf)!pc' = Some pc), - match_stackframes (Stackframe res f sp pc rs) (Stackframe res (fn_RTL xf) sp pc' rs). + | 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) (Stackframe res f' sp pc' rs). Inductive match_states: state -> state -> Prop := - | match_states_intro: - forall st f sp pc rs m st' xf pc' + | match_states_intro dupmap st f sp pc rs m st' f' pc' (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_function f xf) - (DUPLIC: (fn_revmap xf)!pc' = Some pc), - match_states (State st f sp pc rs m) (State st' (fn_RTL xf) sp pc' rs m) + (TRANSF: match_function dupmap f f') + (DUPLIC: dupmap!pc' = Some pc): + match_states (State st f sp pc rs m) (State st' f' sp pc' rs m) | match_states_call: forall st st' f f' args m (STACKS: list_forall2 match_stackframes st st') @@ -114,8 +91,9 @@ Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. Proof. intros TRANSF; destruct f; simpl; monadInv TRANSF. - + monadInv EQ. - eapply match_Internal; eapply transf_function_correct; eauto. + + exploit transf_function_correct; eauto. + intros (dupmap & MATCH_F). + eapply match_Internal; eauto. + eapply match_External. Qed. @@ -128,6 +106,13 @@ Proof. intros. eapply match_transform_partial_program_contextual; eauto. Qed. +(* +Lemma match_prog_RTL {LA: Linker fundef} {LV: Linker unit}: match_program (fun _ f tf => exists f' tf', tf = fundef_RTL f) eq (program_RTL p) (program_RTL tp). +Proof. + eapply match_transform_program; eauto. +Qed. +*) + Section PRESERVATION. Variable prog: program. @@ -137,14 +122,18 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. + +(* Let ge := Genv.globalenv (program_RTL prog). Let tge := Genv.globalenv (program_RTL tprog). +*) -Lemma symbols_preserved_pge s: Genv.find_symbol tpge s = Genv.find_symbol pge s. +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_sym x y: Senv.equiv x y -> Senv.equiv y x. Proof. unfold Senv.equiv. intuition congruence. @@ -154,21 +143,23 @@ Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x Proof. unfold Senv.equiv. intuition congruence. Qed. +*) -Lemma senv_preserved_pge: +Lemma senv_preserved: Senv.equiv pge tpge. Proof. eapply (Genv.senv_match TRANSL). Qed. -Lemma senv_preserved: +(* +Lemma senv_preserved_RTL: Senv.equiv ge tge. Proof. eapply senv_transitivity. { eapply senv_sym; eapply RTLpath.senv_preserved. } eapply senv_transitivity. { eapply senv_preserved_pge. } eapply RTLpath.senv_preserved. Qed. - +*) Lemma functions_translated: forall (v: val) (f: fundef), @@ -182,7 +173,7 @@ Proof. + eapply linkorder_refl. Qed. -Lemma function_ptr_translated_pge: +Lemma function_ptr_translated: forall v f, Genv.find_funct_ptr pge v = Some f -> exists tf, @@ -192,27 +183,40 @@ Proof. exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. Qed. -Lemma function_ptr_translated: +(* +Lemma function_ptr_translated_RTL1: forall v f, Genv.find_funct_ptr ge v = Some f -> exists f0 tf, - Genv.find_funct_ptr pge v = Some f0 /\ Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f0 = OK tf. + Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tpge v = Some tf. Proof. intros; exploit find_funct_ptr_RTL_preserv; eauto. intros (f0 & X & Y); subst. - exploit function_ptr_translated_pge; eauto. - intros (tf & Y & Z); eauto. + exploit function_ptr_translated; eauto. + intros (tf & Y & Z); eexists. eexists; intuition eauto. Qed. +Lemma function_ptr_translated_RTL2: + forall v f, + Genv.find_funct_ptr ge v = Some f -> + exists f0 tf, + Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tge v = Some (fundef_RTL tf). +Proof. + intros; exploit function_ptr_translated_RTL1; eauto. + intros (f0 & tf & X & Y & Z); eexists. eexists; intuition eauto. + exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto. + intros (cunit & tf0 & XX); intuition subst; eauto. +Qed. +*) + Lemma function_sig_translated: forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. Proof. intros. destruct f. - - simpl in H. monadInv H. simpl. symmetry. monadInv EQ. - eapply preserv_fnsig. - eapply transf_function_correct. + - simpl in H. monadInv H. simpl. symmetry. + eapply transf_function_preserves. assumption. - - simpl in H. monadInv H. (* monadInv EQ. *) reflexivity. + - simpl in H. monadInv H. reflexivity. Qed. Theorem transf_initial_states: @@ -220,29 +224,32 @@ Theorem transf_initial_states: exists s2, initial_state tprog s2 /\ match_states s1 s2. Proof. intros. inv H. - exploit function_ptr_translated; eauto. intros (f0 & tf & FIND1 & FIND2 & TRANSF). - exists (RTL.Callstate nil (fn_RTL tf nil m0) nil m0). - eexists. split. - 2: { eapply match_states_call. eauto. - - unfold initial_state. econstructor. - + eapply (Genv.init_mem_transf_partial TRANSL); eauto. + exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF). + exists (Callstate nil tf nil m0). + split. + - econstructor. + + intros; apply (Genv.init_mem_match TRANSL); assumption. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. symmetry. eapply match_program_main. eauto. - + exploit function_ptr_translated; eauto. + + eauto. + destruct f. - * monadInv TRANSF. monadInv EQ. rewrite <- H3. symmetry; eapply transf_function_aux_preserves. assumption. - * monadInv TRANSF. (* monadInv EQ. *) assumption. + * monadInv TRANSF. rewrite <- H3. symmetry; eapply transf_function_preserves. assumption. + * monadInv TRANSF. assumption. - constructor; eauto. constructor. apply transf_fundef_correct; auto. Qed. -(* TODO -Theorem transf_final_states: - forall s1 s2 r, - match_states s1 s2 -> final_state s1 r -> final_state s2 r. +Theorem transf_final_states s1 s2 r: + final_state s1 r -> match_states s1 s2 -> final_state s2 r. Proof. - intros. inv H0. inv H. inv STACKS. constructor. + unfold final_state. + intros H; inv H. + intros H; inv H; simpl in * |- *; try congruence. + inv H1. + destruct st; simpl in * |- *; try congruence. + inv STACKS. constructor. Qed. +(* TODO. Theorem step_simulation: forall s1 t s1', step pge s1 t s1' -> forall s2 (MS: match_states s1 s2), -- cgit From e8b750ba0576d9ce8f148eec827448690874c5e7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 25 Nov 2019 17:06:25 +0100 Subject: progress on RTLpathScheduler (issue on extensionality of rs) --- mppa_k1c/lib/RTLpath.v | 30 ++--- mppa_k1c/lib/RTLpathSE_theory.v | 2 +- mppa_k1c/lib/RTLpathScheduler.v | 291 ++++++++++++++++------------------------ 3 files changed, 134 insertions(+), 189 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index e07822ca..5de0f5a1 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -77,18 +77,18 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, Definition path_map: Type := PTree.t nat. -Definition path_entry (c:code) (pm: path_map) (n: node): Prop - := pm!n <> None /\ c!n <> None. +Definition path_entry (*c:code*) (pm: path_map) (n: node): Prop + := pm!n <> None (*/\ c!n <> None*). Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := | wf_last_node i pc: c!pc = Some i -> - (forall n, List.In n (successors_instr i) -> path_entry c pm n) -> + (forall n, List.In n (successors_instr i) -> path_entry (*c*) pm n) -> wellformed_path c pm 0 pc | wf_internal_node path i pc pc': c!pc = Some i -> default_succ i = Some pc' -> - (forall n, early_exit i = Some n -> path_entry c pm n) -> + (forall n, early_exit i = Some n -> path_entry (*c*) pm n) -> wellformed_path c pm path pc' -> wellformed_path c pm (S path) pc. @@ -105,7 +105,7 @@ Record function : Type := { fn_RTL:> RTL.function; fn_path: path_map; (* condition 1 below: the entry-point of the code is an entry-point of a path *) - fn_entry_point_wf: path_entry fn_RTL.(fn_code) fn_path fn_RTL.(fn_entrypoint); + fn_entry_point_wf: path_entry (*fn_RTL.(fn_code)*) fn_path fn_RTL.(fn_entrypoint); (* condition 2 below: the path_map is well-formed *) fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path }. @@ -522,7 +522,7 @@ Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := Lemma final_node_path f path pc: (fn_path f)!pc = Some path -> exists i, nth_default_succ_inst (fn_code f) path pc = Some i - /\ (forall n, List.In n (successors_instr i) -> path_entry (fn_code f) (fn_path f) n). + /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF. @@ -539,7 +539,7 @@ Lemma internal_node_path path f path0 pc: exists i pc', nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ - (forall n, early_exit i = Some n -> path_entry (fn_code f) (fn_path f) n). + (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } @@ -550,7 +550,7 @@ Proof. simplify_someHyps; eauto. Qed. -Lemma initialize_path c pm n: path_entry c pm n -> exists path, pm!n = Some path. +Lemma initialize_path (*c*) pm n: path_entry (*c*) pm n -> exists path, pm!n = Some path. Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. @@ -697,7 +697,7 @@ Qed. (** *** Definition of well-formed stacks and of match_states *) Definition wf_stf (st: stackframe): Prop := match st with - | Stackframe res f sp pc rs => path_entry f.(fn_code) f.(fn_path) pc + | Stackframe res f sp pc rs => path_entry (*f.(fn_code)*) f.(fn_path) pc end. Definition wf_stackframe (stack: list stackframe): Prop := @@ -796,7 +796,7 @@ Proof. exploit (exec_istate ge pge); eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_code f) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -819,7 +819,7 @@ Proof. intros; exploit exec_Ibuiltin; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_code f) (fn_path f) pc') as (path0 & Hpath0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -828,7 +828,7 @@ Proof. intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_code f) (fn_path f) pc') as (path0 & Hpath0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0; intuition eauto. econstructor; eauto. * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. @@ -882,7 +882,7 @@ Proof. erewrite isteps_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; - destruct (initialize_path (fn_code f) (fn_path f) pc0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. } destruct HPATH0 as ([|path1] & Hpath1). * (* two step case *) @@ -918,7 +918,7 @@ Proof. eexists; constructor 1. * eapply exec_function_internal; eauto. * unfold match_states, match_inst_states; simpl. - destruct (initialize_path (fn_code f) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. @@ -937,7 +937,7 @@ Proof. * unfold match_states, match_inst_states; simpl. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. - destruct (initialize_path (fn_code f0) (fn_path f0) pc0) as (path & Hpath); eauto. + destruct (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto. exists path. constructor; auto. econstructor; eauto. - cutrewrite (path-path=O)%nat; simpl; eauto. diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index b7674916..27325ea1 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -846,7 +846,7 @@ Inductive equiv_state: state -> state -> Prop := | State_equiv stack f sp pc rs1 m rs2: (forall r, rs1#r = rs2#r) -> equiv_state (State stack f sp pc rs1 m) - (State stack f sp pc rs2 m) + (State stack f sp pc rs2 m) | Callstate_equiv stack res f sp pc rs1 m rs2 fd args: (forall r, rs1#r = rs2#r) -> equiv_state (Callstate (Stackframe res f sp pc rs1 :: stack) fd args m) diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index dcdb3419..f5578c0f 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -32,7 +32,7 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := 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 -> path_entry (fn_code f1) (fn_path f1) pc1 -> path_entry (fn_code f2) (fn_path f2) pc2 /\ path_simu (fun n' => dupmap!n') f1 f2 pc1 pc2 + forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (*fn_code f1*) (fn_path f1) pc1 -> path_entry (*fn_code f2*) (fn_path f2) pc2 /\ path_simu (fun n' => dupmap!n') f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) @@ -65,17 +65,19 @@ Inductive match_fundef: RTLpath.fundef -> RTLpath.fundef -> Prop := | match_External ef: match_fundef (External ef) (External ef). Inductive match_stackframes: stackframe -> stackframe -> Prop := - | match_stackframe_intro dupmap res f sp pc rs f' pc' + | match_stackframe_intro dupmap res f sp pc rs1 rs2 f' pc' (TRANSF: match_function dupmap f f') - (DUPLIC: dupmap!pc' = Some pc): - match_stackframes (Stackframe res f sp pc rs) (Stackframe res f' sp pc' rs). + (DUPLIC: dupmap!pc' = Some pc) + (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + match_stackframes (Stackframe res f sp pc rs1) (Stackframe res f' sp pc' rs2). Inductive match_states: state -> state -> Prop := - | match_states_intro dupmap st f sp pc rs m st' f' pc' + | match_states_intro dupmap st f sp pc rs1 rs2 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) (State st' f' sp pc' rs m) + (DUPLIC: dupmap!pc' = Some pc) + (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + match_states (State st f sp pc rs1 m) (State st' f' sp pc' rs2 m) | match_states_call: forall st st' f f' args m (STACKS: list_forall2 match_stackframes st st') @@ -86,6 +88,18 @@ Inductive match_states: state -> state -> Prop := (STACKS: list_forall2 match_stackframes st st'), match_states (Returnstate st v m) (Returnstate st' v m). +Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3. +Proof. + Local Hint Resolve eq_trans. + destruct 1; intros EQ; inv EQ; try econstructor; eauto. + assert (TRIV: exists stk, stk=Stackframe res f0 sp pc rs1 :: stack); eauto. + destruct TRIV as (stk & Hstk). + rewrite <- Hstk in STACKS. + generalize STACKS Hstk; clear STACKS Hstk. + induction 1; simpl; congruence || econstructor; eauto. + - inv Hstk. inv H. econstructor; eauto. + - inv Hstk. inv STACKS; econstructor; eauto. +Qed. Lemma transf_fundef_correct f f': transf_fundef f = OK f' -> match_fundef f f'. @@ -106,13 +120,6 @@ Proof. intros. eapply match_transform_partial_program_contextual; eauto. Qed. -(* -Lemma match_prog_RTL {LA: Linker fundef} {LV: Linker unit}: match_program (fun _ f tf => exists f' tf', tf = fundef_RTL f) eq (program_RTL p) (program_RTL tp). -Proof. - eapply match_transform_program; eauto. -Qed. -*) - Section PRESERVATION. Variable prog: program. @@ -123,44 +130,17 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. -(* -Let ge := Genv.globalenv (program_RTL prog). -Let tge := Genv.globalenv (program_RTL tprog). -*) 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_sym x y: Senv.equiv x y -> Senv.equiv y x. -Proof. - unfold Senv.equiv. intuition congruence. -Qed. - -Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z. -Proof. - unfold Senv.equiv. intuition congruence. -Qed. -*) - Lemma senv_preserved: Senv.equiv pge tpge. Proof. eapply (Genv.senv_match TRANSL). Qed. - -(* -Lemma senv_preserved_RTL: - Senv.equiv ge tge. -Proof. - eapply senv_transitivity. { eapply senv_sym; eapply RTLpath.senv_preserved. } - eapply senv_transitivity. { eapply senv_preserved_pge. } - eapply RTLpath.senv_preserved. -Qed. -*) - Lemma functions_translated: forall (v: val) (f: fundef), Genv.find_funct pge v = Some f -> @@ -183,32 +163,6 @@ Proof. exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. Qed. -(* -Lemma function_ptr_translated_RTL1: - forall v f, - Genv.find_funct_ptr ge v = Some f -> - exists f0 tf, - Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tpge v = Some tf. -Proof. - intros; exploit find_funct_ptr_RTL_preserv; eauto. - intros (f0 & X & Y); subst. - exploit function_ptr_translated; eauto. - intros (tf & Y & Z); eexists. eexists; intuition eauto. -Qed. - -Lemma function_ptr_translated_RTL2: - forall v f, - Genv.find_funct_ptr ge v = Some f -> - exists f0 tf, - Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tge v = Some (fundef_RTL tf). -Proof. - intros; exploit function_ptr_translated_RTL1; eauto. - intros (f0 & tf & X & Y & Z); eexists. eexists; intuition eauto. - exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto. - intros (cunit & tf0 & XX); intuition subst; eauto. -Qed. -*) - Lemma function_sig_translated: forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. Proof. @@ -227,11 +181,10 @@ Proof. exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF). exists (Callstate nil tf nil m0). split. - - econstructor. + - econstructor; eauto. + intros; apply (Genv.init_mem_match TRANSL); assumption. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. symmetry. eapply match_program_main. eauto. - + eauto. + destruct f. * monadInv TRANSF. rewrite <- H3. symmetry; eapply transf_function_preserves. assumption. * monadInv TRANSF. assumption. @@ -249,124 +202,117 @@ Proof. inv STACKS. constructor. Qed. -(* TODO. -Theorem step_simulation: - forall s1 t s1', step pge s1 t s1' -> + +Let ge := Genv.globalenv (program_RTL prog). +Let tge := Genv.globalenv (program_RTL tprog). + +Lemma senv_sym x y: Senv.equiv x y -> Senv.equiv y x. +Proof. + unfold Senv.equiv. intuition congruence. +Qed. + +Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z. +Proof. + unfold Senv.equiv. intuition congruence. +Qed. + +Lemma senv_preserved_RTL: + Senv.equiv ge tge. +Proof. + eapply senv_transitivity. { eapply senv_sym; eapply RTLpath.senv_preserved. } + eapply senv_transitivity. { eapply senv_preserved. } + eapply RTLpath.senv_preserved. +Qed. + +(* +Lemma function_ptr_translated_RTL1: + forall v f, + Genv.find_funct_ptr ge v = Some f -> + exists f0 tf, + Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tpge v = Some tf. +Proof. + intros; exploit find_funct_ptr_RTL_preserv; eauto. + intros (f0 & X & Y); subst. + exploit function_ptr_translated; eauto. + intros (tf & Y & Z); eexists. eexists; intuition eauto. +Qed. + +Lemma function_ptr_translated_RTL2: + forall v f, + Genv.find_funct_ptr ge v = Some f -> + exists f0 tf, + Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tge v = Some (fundef_RTL tf). +Proof. + intros; exploit function_ptr_translated_RTL1; eauto. + intros (f0 & tf & X & Y & Z); eexists. eexists; intuition eauto. + exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto. + intros (cunit & tf0 & XX); intuition subst; eauto. +Qed. +*) + +Lemma symb_exec_path_simu dupmap f f' stk stk' sp st st' rs m t s: + match_function dupmap f f' -> + list_forall2 match_stackframes stk stk' -> + sem_state (Build_sgenv pge ge sp) st stk f rs m t s -> + s_state_simu (fun n' : node => dupmap ! n') st st' -> + exists s', sem_state (Build_sgenv tpge tge sp) st' stk' f' rs m t s' /\ match_states s s'. +Admitted. + +Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: + (fn_path f)!pc = Some path -> + path_step ge pge path stk f sp rs m pc t s -> + list_forall2 match_stackframes stk stk' -> + dupmap ! pc' = Some pc -> + match_function dupmap f f' -> + exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path' stk' f' sp rs m pc' t s' /\ match_states s s'. +Proof. + intros PATH STEP STACKS PC FUNCT. + exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } + unfold path_simu; intros (PATH' & SIMU). + exploit initialize_path; eauto. + clear PATH'; intros (path' & PATH'). + exists path'. + exploit (symb_step_correct f pc pge ge sp path stk rs m t s); eauto. + intros (st & SYMB & SEM); clear STEP. + exploit SIMU; eauto. + intros (st' & SYMB' & SIMU'); clear SIMU SYMB. + exploit symb_exec_path_simu; eauto. + intros (s0 & SEM0 & MATCH). + exploit symb_step_exact; eauto. + intros (s' & STEP' & EQ). + exists s'; intuition. + eapply match_states_equiv; eauto. +Qed. + +Lemma step_simulation: + forall s1 t s1', step ge pge s1 t s1' -> forall s2 (MS: match_states s1 s2), exists s2', - step tge s2 t s2' + step tge tpge s2 t s2' /\ match_states s1' s2'. Proof. Local Hint Resolve transf_fundef_correct. induction 1; intros; inv MS. -(* Inop *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). - inv H3. - eexists. split. - + eapply exec_Inop; eauto. - + constructor; eauto. -(* Iop *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - eexists. split. - + eapply exec_Iop; eauto. erewrite eval_operation_preserved; eauto. - + constructor; eauto. -(* Iload *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - eexists. split. - + eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto. - + constructor; auto. -(* Istore *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - eexists. split. - + eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto. - + constructor; auto. -(* Icall *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - destruct ros. - * simpl in H0. apply functions_translated in H0. - destruct H0 as (tf & cunit & TFUN & GFIND & LO). - eexists. split. - + eapply exec_Icall. eassumption. simpl. eassumption. - apply function_sig_translated. assumption. - + repeat (constructor; auto). - * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. - apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF). - eexists. split. - + eapply exec_Icall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS. - eassumption. apply function_sig_translated. assumption. - + repeat (constructor; auto). -(* Itailcall *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H10 & H11). inv H11. - pose symbols_preserved as SYMPRES. - destruct ros. - * simpl in H0. apply functions_translated in H0. - destruct H0 as (tf & cunit & TFUN & GFIND & LO). - eexists. split. - + eapply exec_Itailcall. eassumption. simpl. eassumption. - apply function_sig_translated. assumption. - erewrite <- preserv_fnstacksize; eauto. - + repeat (constructor; auto). - * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate. - apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF). - eexists. split. - + eapply exec_Itailcall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS. - eassumption. apply function_sig_translated. assumption. - erewrite <- preserv_fnstacksize; eauto. - + repeat (constructor; auto). -(* Ibuiltin *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - 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. - + constructor; auto. -(* Icond *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - eexists. split. - + eapply exec_Icond; eauto. - + constructor; auto. destruct b; auto. -(* Ijumptable *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - exploit list_nth_z_revmap; eauto. intros (pc'1 & LNZ & REVM). - eexists. split. - + eapply exec_Ijumptable; eauto. - + constructor; auto. -(* Ireturn *) - - eapply revmap_correct in DUPLIC; eauto. - destruct DUPLIC as (i' & H2 & H3). inv H3. - pose symbols_preserved as SYMPRES. - eexists. split. - + eapply exec_Ireturn; eauto. erewrite <- preserv_fnstacksize; eauto. - + constructor; auto. +(* exec_path *) + - intros; exploit exec_path_simulation; eauto. + intros (path' & s' & STEP). + exists s'; intuition eauto. + (* ISSUE with extensionality *) + admit. (* exec_function_internal *) - inversion TRANSF as [f0 xf MATCHF|]; subst. eexists. split. + eapply exec_function_internal. erewrite <- preserv_fnstacksize; eauto. + erewrite preserv_fnparams; eauto. - econstructor; eauto. apply revmap_entrypoint. assumption. + econstructor; eauto. apply preserv_entrypoint. assumption. (* exec_function_external *) - inversion TRANSF as [|]; subst. eexists. split. - + econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved. + + econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved_RTL. + constructor. assumption. (* exec_return *) - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split. + constructor. - + inv H1. constructor; assumption. -Qed. + + inv H1. econstructor; eauto. +Admitted. Theorem transf_program_correct: forward_simulation (semantics prog) (semantics tprog). @@ -374,9 +320,8 @@ Proof. eapply forward_simulation_step with match_states. - eapply senv_preserved. - eapply transf_initial_states. - - eapply transf_final_states. + - intros; eapply transf_final_states; eauto. - eapply step_simulation. Qed. -*) End PRESERVATION. -- cgit From ea29b4d7d173d6521e0a384a81b6908e62ee60aa Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 25 Nov 2019 18:05:34 +0100 Subject: progress on RTLpathScheduler (fix extensionality issue) --- mppa_k1c/lib/RTLpathSE_theory.v | 87 +++++++++++++++++++++++++++++++++-------- mppa_k1c/lib/RTLpathScheduler.v | 46 +++++++++++++++------- 2 files changed, 102 insertions(+), 31 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 27325ea1..e6e4a218 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -836,39 +836,87 @@ Proof. rewrite PC; auto. Qed. -(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution - (symb_step is exact). -*) - +(* 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 := - | equiv_state_refl s: equiv_state s s - | State_equiv stack f sp pc rs1 m rs2: - (forall r, rs1#r = rs2#r) -> - equiv_state (State stack f sp pc rs1 m) - (State stack f sp pc rs2 m) - | Callstate_equiv stack res f sp pc rs1 m rs2 fd args: - (forall r, rs1#r = rs2#r) -> - equiv_state (Callstate (Stackframe res f sp pc rs1 :: stack) fd args m) - (Callstate (Stackframe res f sp pc rs2 :: stack) fd args m). + | 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_stackframe 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. + induction stk; simpl; constructor; auto. +Qed. + +Lemma equiv_state_refl s: equiv_state s s. +Proof. + Local Hint Resolve equiv_stack_refl. + 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 sfval_sem_equiv sge (f:function) st sv stack rs0 m0 t rs1 rs2 m s: sfval_sem sge st stack f rs0 m0 sv rs1 m t s -> (forall r, rs1#r = rs2#r) -> exists s', equiv_state s s' /\ sfval_sem sge st stack f rs0 m0 sv rs2 m t s'. Proof. + Local Hint Resolve equiv_stack_refl. destruct 1. - (* Snone *) intros; eexists; econstructor. + eapply State_equiv; eauto. + eapply exec_Snone. - (* Scall *) intros; eexists; econstructor. - + eapply Callstate_equiv; eauto. - + eapply exec_Scall; eauto. + 2: { eapply exec_Scall; eauto. } + apply Call_equiv; auto. + repeat (constructor; auto). - (* Sreturn *) intros; eexists; econstructor. + eapply equiv_state_refl; eauto. + eapply exec_Sreturn; eauto. Qed. +(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution + (symb_step is exact). +*) Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> @@ -945,4 +993,11 @@ Definition s_state_simu (srce: node -> option node) (s1 s2: s_state): Prop := Definition path_simu (srce: node -> option node) (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. -(** See usage of [path_simu] in [RTLpathScheduler] *) \ No newline at end of file +(** See usage of [path_simu] in [RTLpathScheduler] *) + +(* TODO: à déplacer dans RTLpath ? *) +Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: + path_step ge pge path stk f sp rs1 m pc t s1 -> + (forall r : positive, rs1 !! r = rs2 !! r) -> + exists s2, path_step ge pge path stk f sp rs2 m pc t s2 /\ equiv_state s1 s2. +Admitted. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index f5578c0f..2417a43d 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -88,17 +88,27 @@ Inductive match_states: state -> state -> Prop := (STACKS: list_forall2 match_stackframes st st'), match_states (Returnstate st v m) (Returnstate st' v m). -Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3. +Lemma match_stackframes_equiv stf1 stf2 stf3: + match_stackframes stf1 stf2 -> equiv_stackframe stf2 stf3 -> match_stackframes stf1 stf3. Proof. - Local Hint Resolve eq_trans. destruct 1; intros EQ; inv EQ; try econstructor; eauto. - assert (TRIV: exists stk, stk=Stackframe res f0 sp pc rs1 :: stack); eauto. - destruct TRIV as (stk & Hstk). - rewrite <- Hstk in STACKS. - generalize STACKS Hstk; clear STACKS Hstk. - induction 1; simpl; congruence || econstructor; eauto. - - inv Hstk. inv H. econstructor; eauto. - - inv Hstk. inv STACKS; econstructor; eauto. + intros; eapply eq_trans; eauto. +Qed. + +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. + induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. +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. + destruct 1; intros EQ; inv EQ; econstructor; eauto. + intros; eapply eq_trans; eauto. Qed. Lemma transf_fundef_correct f f': @@ -294,11 +304,14 @@ Proof. Local Hint Resolve transf_fundef_correct. induction 1; intros; inv MS. (* exec_path *) - - intros; exploit exec_path_simulation; eauto. - intros (path' & s' & STEP). - exists s'; intuition eauto. - (* ISSUE with extensionality *) - admit. + - intros. + exploit exec_path_simulation; eauto. + intros (path' & s' & STEP'). clear H0. intuition. + exploit path_step_equiv; eauto. + intros (s2 & STEP & EQ). + exists s2; split. + + eapply exec_path; eauto. + + eapply match_states_equiv; eauto. (* exec_function_internal *) - inversion TRANSF as [f0 xf MATCHF|]; subst. eexists. split. + eapply exec_function_internal. erewrite <- preserv_fnstacksize; eauto. @@ -312,7 +325,10 @@ Proof. - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split. + constructor. + inv H1. econstructor; eauto. -Admitted. + intros r. destruct (Pos.eq_dec res' r). + * subst. rewrite ! Registers.Regmap.gss; auto. + * rewrite ! Registers.Regmap.gso; auto. +Qed. Theorem transf_program_correct: forward_simulation (semantics prog) (semantics tprog). -- cgit From ebba81108ca2fa28cc8f63c1749cdd251a35c13f Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 25 Nov 2019 18:56:00 +0100 Subject: progress on RTLpathScheduler --- mppa_k1c/lib/RTLpathSE_theory.v | 8 +++++++- mppa_k1c/lib/RTLpathScheduler.v | 23 +++++++++++++++++++++++ 2 files changed, 30 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index e6e4a218..573ce102 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -975,8 +975,11 @@ Definition istate_simu (srce: node -> option node) ist1 ist2 := (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ ist1.(the_mem) = ist2.(the_mem). -Definition s_istate_simu (srce: node -> option node) (st1 st2: s_istate): Prop := +(* TODO: what is it the good notion ? *) +Parameter s_istate_simu: forall (srce: node -> option node) (st1 st2: s_istate), Prop. +(* This one is may not the good one... forall sge1 sge2 rs0 m0 is1 is2, istate_simu srce is1 is2 -> sem_istate sge1 st1 rs0 m0 is1 -> sem_istate sge2 st2 rs0 m0 is2. +*) Inductive sfval_simu (srce: node -> option node): sfval -> sfval -> Prop := | Snone_simu: sfval_simu srce Snone Snone @@ -993,6 +996,9 @@ Definition s_state_simu (srce: node -> option node) (s1 s2: s_state): Prop := Definition path_simu (srce: node -> option node) (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. +(* IDEA: we will provide an efficient test for checking "path_simu" property...*) + + (** See usage of [path_simu] in [RTLpathScheduler] *) (* TODO: à déplacer dans RTLpath ? *) diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 2417a43d..ec64e65b 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -260,12 +260,35 @@ Proof. Qed. *) + +Lemma sem_istate_simu dupmap sp f f' sis sis' rs m is: + match_function dupmap f f' -> + sem_istate {| the_pge := pge; the_ge := ge; the_sp := sp |} sis rs m is -> + s_istate_simu (fun n' : node => dupmap ! n') sis sis' -> + exists is', + sem_istate {| the_pge := tpge; the_ge := tge; the_sp := sp |} sis' rs m is' /\ + istate_simu (fun n' : node => dupmap ! n') is is'. +Admitted. + Lemma symb_exec_path_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> sem_state (Build_sgenv pge ge sp) st stk f rs m t s -> s_state_simu (fun n' : node => dupmap ! n') st st' -> exists s', sem_state (Build_sgenv tpge tge sp) st' stk' f' rs m t s' /\ match_states s s'. +Proof. + intros FUNC STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. + - (* sem_early *) + exploit sem_istate_simu; eauto. + intros (is' & SEM' & (CONT' & PC' & RS' & M')). + exists (State stk' f' sp (RTLpath.the_pc is') (the_rs is') (the_mem is')). + split. + + eapply sem_early; auto. congruence. + + rewrite M'. econstructor; eauto. + - (* sem_normal *) + exploit sem_istate_simu; eauto. + intros (is' & SEM' & (CONT' & PC' & RS' & M')). + admit. Admitted. Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: -- cgit From 05ac2c67b08906425663297f1bac04b7e6735fd3 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 08:38:13 +0100 Subject: statement of the main theorem on simulation of symbolic executions --- mppa_k1c/lib/RTLpathSE_theory.v | 23 ++++++++++++----------- mppa_k1c/lib/RTLpathScheduler.v | 39 +++++++++++++++++++++++++++++---------- 2 files changed, 41 insertions(+), 21 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 573ce102..bbb1a9fc 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -969,37 +969,37 @@ Admitted. (** * Simulation of RTLpath code w.r.t symbolic execution *) -Definition istate_simu (srce: node -> option node) ist1 ist2 := +Definition istate_simu (srce: PTree.t node) ist1 ist2 := ist1.(continue) = ist2.(continue) /\ - srce (ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ + srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ ist1.(the_mem) = ist2.(the_mem). (* TODO: what is it the good notion ? *) -Parameter s_istate_simu: forall (srce: node -> option node) (st1 st2: s_istate), Prop. -(* This one is may not the good one... +Parameter s_istate_simu: forall (srce: PTree.t node) (st1 st2: s_istate), Prop. +(* This one is probably not the good one... forall sge1 sge2 rs0 m0 is1 is2, istate_simu srce is1 is2 -> sem_istate sge1 st1 rs0 m0 is1 -> sem_istate sge2 st2 rs0 m0 is2. *) -Inductive sfval_simu (srce: node -> option node): sfval -> sfval -> Prop := +Inductive sfval_simu (srce: PTree.t node): sfval -> sfval -> Prop := | Snone_simu: sfval_simu srce Snone Snone | Scall_simu sig svos lsv res pc1 pc2: - srce pc2 = Some pc1 -> + srce!pc2 = Some pc1 -> sfval_simu srce (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) | Sreturn_simu os: sfval_simu srce (Sreturn os) (Sreturn os). -Definition s_state_simu (srce: node -> option node) (s1 s2: s_state): Prop := +Definition s_state_simu srce (s1 s2: s_state): Prop := s_istate_simu srce s1.(internal) s2.(internal) /\ sfval_simu srce s1.(final) s2.(final). -Definition path_simu (srce: node -> option node) (f1 f2: RTLpath.function) pc1 pc2: Prop := +Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. -(* IDEA: we will provide an efficient test for checking "path_simu" property...*) +(* IDEA: we will provide an efficient test for checking "path_simu" property, by hash-consing. + See usage of [symb_step_simu] in [RTLpathScheduler]. - -(** See usage of [path_simu] in [RTLpathScheduler] *) +*) (* TODO: à déplacer dans RTLpath ? *) Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: @@ -1007,3 +1007,4 @@ Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: (forall r : positive, rs1 !! r = rs2 !! r) -> exists s2, path_step ge pge path stk f sp rs2 m pc t s2 /\ equiv_state s1 s2. Admitted. + diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index ec64e65b..b509df05 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -32,7 +32,7 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := 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 -> path_entry (*fn_code f1*) (fn_path f1) pc1 -> path_entry (*fn_code f2*) (fn_path f2) pc2 /\ path_simu (fun n' => dupmap!n') f1 f2 pc1 pc2 + forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (*fn_code f1*) (fn_path f1) pc1 -> path_entry (*fn_code f2*) (fn_path f2) pc2 /\ symb_step_simu dupmap f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) @@ -198,7 +198,9 @@ Proof. + destruct f. * monadInv TRANSF. rewrite <- H3. symmetry; eapply transf_function_preserves. assumption. * monadInv TRANSF. assumption. - - constructor; eauto. constructor. apply transf_fundef_correct; auto. + - constructor; eauto. + + constructor. + + apply transf_fundef_correct; auto. Qed. Theorem transf_final_states s1 s2 r: @@ -264,17 +266,28 @@ Qed. Lemma sem_istate_simu dupmap sp f f' sis sis' rs m is: match_function dupmap f f' -> sem_istate {| the_pge := pge; the_ge := ge; the_sp := sp |} sis rs m is -> - s_istate_simu (fun n' : node => dupmap ! n') sis sis' -> + s_istate_simu dupmap sis sis' -> exists is', sem_istate {| the_pge := tpge; the_ge := tge; the_sp := sp |} sis' rs m is' /\ - istate_simu (fun n' : node => dupmap ! n') is is'. + istate_simu dupmap is is'. Admitted. -Lemma symb_exec_path_simu dupmap f f' stk stk' sp st st' rs m t s: +Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: + match_function dupmap f f' -> + list_forall2 match_stackframes stk stk' -> + sfval_sem (Build_sgenv pge ge sp) st stk f rs0 m0 sv rs m t s -> + sfval_simu dupmap sv sv' -> + exists s', (* FIXME ? exists rs', m', *) + sfval_sem (Build_sgenv tpge tge sp) st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. +Admitted. + + +(* The main theorem on simulation of symbolic states ! *) +Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> sem_state (Build_sgenv pge ge sp) st stk f rs m t s -> - s_state_simu (fun n' : node => dupmap ! n') st st' -> + s_state_simu dupmap st st' -> exists s', sem_state (Build_sgenv tpge tge sp) st' stk' f' rs m t s' /\ match_states s s'. Proof. intros FUNC STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. @@ -288,8 +301,14 @@ Proof. - (* sem_normal *) exploit sem_istate_simu; eauto. intros (is' & SEM' & (CONT' & PC' & RS' & M')). - admit. -Admitted. + exploit sem_sfval_simu; eauto. + clear SEM2; intros (s0 & SEM2 & MATCH0). + exploit sfval_sem_equiv; eauto. + clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s1 & EQ & SEM2). + exists s1; split. + + eapply sem_normal; eauto. + + eapply match_states_equiv; eauto. +Qed. Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: (fn_path f)!pc = Some path -> @@ -301,7 +320,7 @@ Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: Proof. intros PATH STEP STACKS PC FUNCT. exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } - unfold path_simu; intros (PATH' & SIMU). + unfold symb_step_simu; intros (PATH' & SIMU). exploit initialize_path; eauto. clear PATH'; intros (path' & PATH'). exists path'. @@ -309,7 +328,7 @@ Proof. intros (st & SYMB & SEM); clear STEP. exploit SIMU; eauto. intros (st' & SYMB' & SIMU'); clear SIMU SYMB. - exploit symb_exec_path_simu; eauto. + exploit sem_symb_state_simu; eauto. intros (s0 & SEM0 & MATCH). exploit symb_step_exact; eauto. intros (s' & STEP' & EQ). -- cgit From 17d2d385a3b45b1498520159ad6ee640bfc6822d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 10:24:38 +0100 Subject: Scheduling proof reduced to properties of the simulation between internal symbolic states --- mppa_k1c/lib/RTLpathScheduler.v | 101 ++++++++++++++++++++++++++++------------ 1 file changed, 72 insertions(+), 29 deletions(-) diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index b509df05..949fe474 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -151,6 +151,7 @@ Lemma senv_preserved: Proof. eapply (Genv.senv_match TRANSL). Qed. + Lemma functions_translated: forall (v: val) (f: fundef), Genv.find_funct pge v = Some f -> @@ -236,51 +237,93 @@ Proof. eapply RTLpath.senv_preserved. Qed. -(* -Lemma function_ptr_translated_RTL1: - forall v f, - Genv.find_funct_ptr ge v = Some f -> - exists f0 tf, - Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tpge v = Some tf. +Lemma symbols_preserved_RTL s: Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. - intros; exploit find_funct_ptr_RTL_preserv; eauto. - intros (f0 & X & Y); subst. - exploit function_ptr_translated; eauto. - intros (tf & Y & Z); eexists. eexists; intuition eauto. + unfold tge, ge. erewrite RTLpath.symbols_preserved; eauto. + rewrite symbols_preserved. + erewrite RTLpath.symbols_preserved; eauto. Qed. -Lemma function_ptr_translated_RTL2: - forall v f, - Genv.find_funct_ptr ge v = Some f -> - exists f0 tf, - Genv.find_funct_ptr pge v = Some f0 /\ transf_fundef f0 = OK tf /\ Genv.find_funct_ptr tge v = Some (fundef_RTL tf). +Lemma sval_eval_translated sp sv rs0 m0: + sval_eval (Build_sgenv pge ge sp) sv rs0 m0 = sval_eval (Build_sgenv tpge tge sp) sv rs0 m0. Proof. - intros; exploit function_ptr_translated_RTL1; eauto. - intros (f0 & tf & X & Y & Z); eexists. eexists; intuition eauto. - exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto. - intros (cunit & tf0 & XX); intuition subst; eauto. + Local Hint Resolve symbols_preserved_RTL. + induction sv using sval_mut with (P0 := fun lsv => list_sval_eval (Build_sgenv pge ge sp) lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge sp) lsv rs0 m0) + (P1 := fun sm => smem_eval (Build_sgenv pge ge sp) sm rs0 m0 = smem_eval (Build_sgenv tpge tge sp) sm rs0 m0); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing tge sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (sval_eval _ _ _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing tge sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (smem_eval _ _ _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma list_sval_eval_translated sp lsv rs0 m0: + list_sval_eval (Build_sgenv pge ge sp) lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge sp) lsv rs0 m0. +Proof. + induction lsv; simpl; auto. + rewrite sval_eval_translated. destruct (sval_eval _ _ _ _); auto. + rewrite IHlsv; auto. +Qed. + +Lemma s_find_function_translated sp svos rs0 m0 fd: + s_find_function (Build_sgenv pge ge sp) svos rs0 m0 = Some fd -> + exists fd', s_find_function (Build_sgenv tpge tge sp) svos rs0 m0 = Some fd' + /\ transf_fundef fd = OK fd'. +Proof. + unfold s_find_function. destruct svos; simpl. + + rewrite sval_eval_translated. destruct (sval_eval _ _ _ _); try congruence. + intros; exploit functions_translated; eauto. + intros (fd' & cunit & X). eexists. intuition eauto. + + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. + intros; apply function_ptr_translated; auto. Qed. -*) +Lemma s_istate_simu_pc dupmap st st': s_istate_simu dupmap st st' -> dupmap ! (the_pc st') = Some (the_pc st). +Admitted. -Lemma sem_istate_simu dupmap sp f f' sis sis' rs m is: +Lemma sem_istate_simu dupmap sp f f' st st' rs m is: match_function dupmap f f' -> - sem_istate {| the_pge := pge; the_ge := ge; the_sp := sp |} sis rs m is -> - s_istate_simu dupmap sis sis' -> + sem_istate (Build_sgenv pge ge sp) st rs m is -> + s_istate_simu dupmap st st' -> exists is', - sem_istate {| the_pge := tpge; the_ge := tge; the_sp := sp |} sis' rs m is' /\ + sem_istate (Build_sgenv tpge tge sp) st' rs m is' /\ istate_simu dupmap is is'. Admitted. Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> - sfval_sem (Build_sgenv pge ge sp) st stk f rs0 m0 sv rs m t s -> + s_istate_simu dupmap st st' -> sfval_simu dupmap sv sv' -> - exists s', (* FIXME ? exists rs', m', *) - sfval_sem (Build_sgenv tpge tge sp) st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. -Admitted. - + sfval_sem (Build_sgenv pge ge sp) st stk f rs0 m0 sv rs m t s -> + exists s', sfval_sem (Build_sgenv tpge tge sp) st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. +Proof. + Local Hint Resolve s_istate_simu_pc transf_fundef_correct. + intros FUN STK SIS SFV. destruct SFV. + - (* Snone *) intros SEM; inv SEM; simpl. + eexists; split; econstructor; eauto. + - (* Scall *) intros SEM; inv SEM; simpl. + exploit s_find_function_translated; eauto. + intros (fd' & FIND & TRANSF). + erewrite <- function_sig_translated; eauto. + eexists; split; econstructor; eauto. + + erewrite <- list_sval_eval_translated; auto. + + simpl. repeat (econstructor; eauto). + - (* Sreturn *) intros SEM; inv SEM; simpl. + eexists; split; econstructor; eauto. + + erewrite <- preserv_fnstacksize; eauto. + + destruct os; auto. + erewrite <- sval_eval_translated; eauto. +Qed. (* The main theorem on simulation of symbolic states ! *) Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: -- cgit From 8d043a0c1040b324f49503d2b3ecf48669b3ab03 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 15:47:26 +0100 Subject: fix sfval_simu definition --- mppa_k1c/lib/RTLpathSE_theory.v | 16 +++++++--------- mppa_k1c/lib/RTLpathScheduler.v | 7 ++----- 2 files changed, 9 insertions(+), 14 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index bbb1a9fc..bfd45df9 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -977,21 +977,19 @@ Definition istate_simu (srce: PTree.t node) ist1 ist2 := (* TODO: what is it the good notion ? *) Parameter s_istate_simu: forall (srce: PTree.t node) (st1 st2: s_istate), Prop. -(* This one is probably not the good one... - forall sge1 sge2 rs0 m0 is1 is2, istate_simu srce is1 is2 -> sem_istate sge1 st1 rs0 m0 is1 -> sem_istate sge2 st2 rs0 m0 is2. -*) -Inductive sfval_simu (srce: PTree.t node): sfval -> sfval -> Prop := - | Snone_simu: sfval_simu srce Snone Snone +Inductive sfval_simu (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := + | Snone_simu: + srce!opc2 = Some opc1 -> sfval_simu srce opc1 opc2 Snone Snone | Scall_simu sig svos lsv res pc1 pc2: srce!pc2 = Some pc1 -> - sfval_simu srce (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) + sfval_simu srce opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) | Sreturn_simu os: - sfval_simu srce (Sreturn os) (Sreturn os). + sfval_simu srce opc1 opc2 (Sreturn os) (Sreturn os). Definition s_state_simu srce (s1 s2: s_state): Prop := - s_istate_simu srce s1.(internal) s2.(internal) - /\ sfval_simu srce s1.(final) s2.(final). + s_istate_simu srce s1.(internal) s2.(internal) + /\ sfval_simu srce s1.(the_pc) s2.(the_pc) s1.(final) s2.(final). Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 949fe474..6cca7ceb 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -287,9 +287,6 @@ Proof. intros; apply function_ptr_translated; auto. Qed. -Lemma s_istate_simu_pc dupmap st st': s_istate_simu dupmap st st' -> dupmap ! (the_pc st') = Some (the_pc st). -Admitted. - Lemma sem_istate_simu dupmap sp f f' st st' rs m is: match_function dupmap f f' -> sem_istate (Build_sgenv pge ge sp) st rs m is -> @@ -303,11 +300,11 @@ Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> s_istate_simu dupmap st st' -> - sfval_simu dupmap sv sv' -> + sfval_simu dupmap st.(the_pc) st'.(the_pc) sv sv' -> sfval_sem (Build_sgenv pge ge sp) st stk f rs0 m0 sv rs m t s -> exists s', sfval_sem (Build_sgenv tpge tge sp) st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. - Local Hint Resolve s_istate_simu_pc transf_fundef_correct. + Local Hint Resolve transf_fundef_correct. intros FUN STK SIS SFV. destruct SFV. - (* Snone *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. -- cgit From dddc36c06aeb04ee49704db117ac33c0368411b1 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 16:27:38 +0100 Subject: removing sp from sgenv --- mppa_k1c/lib/RTLpathSE_theory.v | 312 ++++++++++++++++++++-------------------- mppa_k1c/lib/RTLpathScheduler.v | 32 +++-- 2 files changed, 172 insertions(+), 172 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index bfd45df9..b3a7e6ad 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -8,7 +8,6 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. -Require Coq.Logic.FunctionalExtensionality. (* NOTE for the implementation of the symbolic execution. It is important to remove dependency of Op on memory. @@ -24,8 +23,7 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ Record sgenv := { the_pge: RTLpath.genv; - the_ge: RTL.genv; - the_sp: val + the_ge: RTL.genv }. (* symbolic value *) @@ -53,101 +51,99 @@ Fixpoint list_val_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. -Fixpoint sval_eval (sge: sgenv) (sv: sval) (rs0: regset) (m0: mem): option val := +Fixpoint sval_eval (sge: sgenv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := match sv with | Sinput r => Some (rs0#r) | Sop op l sm => - SOME args <- list_sval_eval sge l rs0 m0 IN - SOME m <- smem_eval sge sm rs0 m0 IN - eval_operation sge.(the_ge) sge.(the_sp) op args m + SOME args <- list_sval_eval sge sp l rs0 m0 IN + SOME m <- smem_eval sge sp sm rs0 m0 IN + eval_operation sge.(the_ge) sp op args m | Sload sm chunk addr lsv => - SOME args <- list_sval_eval sge lsv rs0 m0 IN - SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN - SOME m <- smem_eval sge sm rs0 m0 IN + SOME args <- list_sval_eval sge sp lsv rs0 m0 IN + SOME a <- eval_addressing sge.(the_ge) sp addr args IN + SOME m <- smem_eval sge sp sm rs0 m0 IN Mem.loadv chunk m a end -with list_sval_eval (sge: sgenv) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := +with list_sval_eval (sge: sgenv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := match lsv with | Snil => Some nil | Scons sv lsv' => - SOME v <- sval_eval sge sv rs0 m0 IN - SOME lv <- list_sval_eval sge lsv' rs0 m0 IN + SOME v <- sval_eval sge sp sv rs0 m0 IN + SOME lv <- list_sval_eval sge sp lsv' rs0 m0 IN Some (v::lv) end -with smem_eval (sge: sgenv) (sm: smem) (rs0: regset) (m0: mem): option mem := +with smem_eval (sge: sgenv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem := match sm with | Sinit => Some m0 | Sstore sm chunk addr lsv srce => - SOME args <- list_sval_eval sge lsv rs0 m0 IN - SOME a <- eval_addressing sge.(the_ge) sge.(the_sp) addr args IN - SOME m <- smem_eval sge sm rs0 m0 IN - SOME sv <- sval_eval sge srce rs0 m0 IN + SOME args <- list_sval_eval sge sp lsv rs0 m0 IN + SOME a <- eval_addressing sge.(the_ge) sp addr args IN + SOME m <- smem_eval sge sp sm rs0 m0 IN + SOME sv <- sval_eval sge sp srce rs0 m0 IN Mem.storev chunk m a sv end. (* Syntax and Semantics of local symbolic internal states *) -Record local_istate := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. +Record local_istate := { pre: sgenv -> val -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. -Definition sem_local_istate (sge: sgenv) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := - st.(pre) sge rs0 m0 - /\ smem_eval sge st.(the_smem) rs0 m0 = Some m - /\ forall (r:reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = Some (rs#r). +Definition sem_local_istate (sge: sgenv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := + st.(pre) sge sp rs0 m0 + /\ smem_eval sge sp st.(the_smem) rs0 m0 = Some m + /\ forall (r:reg), sval_eval sge sp (st.(the_sreg) r) rs0 m0 = Some (rs#r). -Definition has_aborted_on_basic (sge: sgenv) (st: local_istate) (rs0: regset) (m0: mem): Prop := - ~(st.(pre) sge rs0 m0) - \/ smem_eval sge st.(the_smem) rs0 m0 = None - \/ exists (r: reg), sval_eval sge (st.(the_sreg) r) rs0 m0 = None. +Definition has_aborted_on_basic (sge: sgenv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem): Prop := + ~(st.(pre) sge sp rs0 m0) + \/ smem_eval sge sp st.(the_smem) rs0 m0 = None + \/ exists (r: reg), sval_eval sge sp (st.(the_sreg) r) rs0 m0 = None. (* Syntax and semantics of symbolic exit states *) Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. -Definition seval_condition sge (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := - SOME args <- list_sval_eval sge lsv rs0 m0 IN - SOME m <- smem_eval sge sm rs0 m0 IN +Definition seval_condition sge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := + SOME args <- list_sval_eval sge sp lsv rs0 m0 IN + SOME m <- smem_eval sge sp sm rs0 m0 IN eval_condition cond args m. -Definition has_not_yet_exit sge (lx: list exit_istate) rs0 m0: Prop := - forall ext, List.In ext lx -> seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. +Definition has_not_yet_exit sge sp (lx: list exit_istate) rs0 m0: Prop := + forall ext, List.In ext lx -> seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. -Definition has_aborted_on_test sge (lx: list exit_istate) rs0 m0: Prop := - exists ext, List.In ext lx /\ seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. +Definition has_aborted_on_test sge sp (lx: list exit_istate) rs0 m0: Prop := + exists ext, List.In ext lx /\ seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. -Inductive sem_early_exit (sge: sgenv) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := +Inductive sem_early_exit (sge: sgenv) (sp:val) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := SEE_intro ext lx': is_tail (ext::lx') lx -> - has_not_yet_exit sge lx' rs0 m0 -> - seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> - sem_local_istate sge ext.(exit_ist) rs0 m0 rs m -> + has_not_yet_exit sge sp lx' rs0 m0 -> + seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> + sem_local_istate sge sp ext.(exit_ist) rs0 m0 rs m -> ext.(cond_ifso) = pc -> - sem_early_exit sge lx rs0 m0 rs m pc. + sem_early_exit sge sp lx rs0 m0 rs m pc. (** * Syntax and Semantics of symbolic internal state *) Record s_istate := { the_pc: node; exits: list exit_istate; curr: local_istate }. -Definition sem_istate (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := +Definition sem_istate (sge: sgenv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(continue)) then - sem_local_istate sge st.(curr) rs0 m0 is.(the_rs) is.(the_mem) + sem_local_istate sge sp st.(curr) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) - /\ has_not_yet_exit sge st.(exits) rs0 m0 - else sem_early_exit sge st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). + /\ has_not_yet_exit sge sp st.(exits) rs0 m0 + else sem_early_exit sge sp st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). -Definition has_aborted (sge: sgenv) (st: s_istate) (rs0: regset) (m0: mem): Prop := - has_aborted_on_basic sge st.(curr) rs0 m0 - \/ has_aborted_on_test sge st.(exits) rs0 m0. +Definition has_aborted (sge: sgenv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem): Prop := + has_aborted_on_basic sge sp st.(curr) rs0 m0 + \/ has_aborted_on_test sge sp st.(exits) rs0 m0. -(* TODO: est-ce vraiment utile ? *) -Definition sem_option_istate sge (st: s_istate) rs0 m0 (ois: option istate): Prop := +Definition sem_option_istate sge sp (st: s_istate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => sem_istate sge st rs0 m0 is - | None => has_aborted sge st rs0 m0 + | Some is => sem_istate sge sp st rs0 m0 is + | None => has_aborted sge sp st rs0 m0 end. -(* TODO: est-ce vraiment utile ? *) -Definition sem_option2_istate sge (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := +Definition sem_option2_istate sge sp (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => sem_option_istate sge st rs0 m0 ois + | Some st => sem_option_istate sge sp st rs0 m0 ois | None => ois=None end. @@ -164,17 +160,17 @@ Definition istate_eq ist1 ist2 := ist1.(the_mem) = ist2.(the_mem). (* TODO: is it useful -Lemma has_not_yet_exit_cons_split sge ext lx rs0 m0: - has_not_yet_exit sge (ext::lx) rs0 m0 <-> - (seval_condition sge ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge lx rs0 m0). +Lemma has_not_yet_exit_cons_split sge sp ext lx rs0 m0: + has_not_yet_exit sge sp (ext::lx) rs0 m0 <-> + (seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge sp lx rs0 m0). Proof. unfold has_not_yet_exit; simpl; intuition (subst; auto). Qed. *) -Lemma exclude_early_NYE sge st rs0 m0 rs m pc: - sem_early_exit sge (exits st) rs0 m0 rs m pc -> - has_not_yet_exit sge (exits st) rs0 m0 -> +Lemma exclude_early_NYE sge sp st rs0 m0 rs m pc: + sem_early_exit sge sp (exits st) rs0 m0 rs m pc -> + has_not_yet_exit sge sp (exits st) rs0 m0 -> False. Proof. Local Hint Resolve is_tail_in. @@ -183,11 +179,11 @@ Proof. try congruence. Qed. -Lemma sem_istate_exclude_incompatible_continue sge st rs m is1 is2: +Lemma sem_istate_exclude_incompatible_continue sge sp st rs m is1 is2: is1.(continue) = true -> is2.(continue) = false -> - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> + sem_istate sge sp st rs m is1 -> + sem_istate sge sp st rs m is2 -> False. Proof. Local Hint Resolve exclude_early_NYE. @@ -197,9 +193,9 @@ Proof. intuition eauto. Qed. -Lemma sem_istate_determ_continue sge st rs m is1 is2: - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> +Lemma sem_istate_determ_continue sge sp st rs m is1 is2: + sem_istate sge sp st rs m is1 -> + sem_istate sge sp st rs m is2 -> is1.(continue) = is2.(continue). Proof. Local Hint Resolve sem_istate_exclude_incompatible_continue. @@ -208,9 +204,9 @@ Proof. destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. Qed. -Lemma sem_local_istate_determ sge st rs0 m0 rs1 m1 rs2 m2: - sem_local_istate sge st rs0 m0 rs1 m1 -> - sem_local_istate sge st rs0 m0 rs2 m2 -> +Lemma sem_local_istate_determ sge sp st rs0 m0 rs1 m1 rs2 m2: + sem_local_istate sge sp st rs0 m0 rs1 m1 -> + sem_local_istate sge sp st rs0 m0 rs2 m2 -> (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. unfold sem_local_istate. intuition try congruence. @@ -225,13 +221,13 @@ Proof. intros T2; inversion T2; subst; auto. Qed. -Lemma exit_cond_determ sge rs0 m0 l1 l2: +Lemma exit_cond_determ sge sp rs0 m0 l1 l2: is_tail l1 l2 -> forall ext1 lx1 ext2 lx2, l1=(ext1 :: lx1) -> l2=(ext2 :: lx2) -> - has_not_yet_exit sge lx1 rs0 m0 -> - seval_condition sge (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> - has_not_yet_exit sge lx2 rs0 m0 -> + has_not_yet_exit sge sp lx1 rs0 m0 -> + seval_condition sge sp (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> + has_not_yet_exit sge sp lx2 rs0 m0 -> ext1=ext2. Proof. destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst; @@ -243,9 +239,9 @@ Proof. try congruence. Qed. -Lemma sem_early_exit_determ sge lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : - sem_early_exit sge lx rs0 m0 rs1 m1 pc1 -> - sem_early_exit sge lx rs0 m0 rs2 m2 pc2 -> +Lemma sem_early_exit_determ sge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : + sem_early_exit sge sp lx rs0 m0 rs1 m1 pc1 -> + sem_early_exit sge sp lx rs0 m0 rs2 m2 pc2 -> pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. Local Hint Resolve exit_cond_determ eq_sym. @@ -253,28 +249,28 @@ Proof. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. assert (X:ext1=ext2). - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. - - subst. destruct (sem_local_istate_determ sge (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. + - subst. destruct (sem_local_istate_determ sge sp (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. -Lemma sem_istate_determ sge st rs m is1 is2: - sem_istate sge st rs m is1 -> - sem_istate sge st rs m is2 -> +Lemma sem_istate_determ sge sp st rs m is1 is2: + sem_istate sge sp st rs m is1 -> + sem_istate sge sp st rs m is2 -> istate_eq is1 is2. Proof. unfold istate_eq. intros SEM1 SEM2. - exploit (sem_istate_determ_continue sge st rs m is1 is2); eauto. + exploit (sem_istate_determ_continue sge sp st rs m is1 is2); eauto. intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. destruct (continue is2). - - destruct (sem_local_istate_determ sge (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); + - destruct (sem_local_istate_determ sge sp (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); intuition (try congruence). - - destruct (sem_early_exit_determ sge (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); + - destruct (sem_early_exit_determ sge sp (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); intuition. Qed. -Lemma sem_istate_exclude_abort sge st rs m is: - has_aborted sge st rs m -> - sem_istate sge st rs m is -> False. +Lemma sem_istate_exclude_abort sge sp st rs m is: + has_aborted sge sp st rs m -> + sem_istate sge sp st rs m is -> False. Proof. intros [[AB|[AB|AB]]|AB]; ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; @@ -285,9 +281,9 @@ Admitted. Definition istate_eq_option ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -Lemma sem_option_istate_determ sge st rs m ois is: - sem_option_istate sge st rs m ois -> - sem_istate sge st rs m is -> +Lemma sem_option_istate_determ sge sp st rs m ois is: + sem_option_istate sge sp st rs m ois -> + sem_istate sge sp st rs m is -> istate_eq_option is ois. Proof. destruct ois as [is1|]; simpl; eauto. @@ -298,12 +294,12 @@ Qed. (** * Symbolic execution of one internal step *) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := - {| pre:=(fun sge rs m => sval_eval sge (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge rs m)); + {| pre:=(fun sge sp rs m => sval_eval sge sp (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge sp rs m)); the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; the_smem:= st.(the_smem)|}. Definition smem_set (st:local_istate) (sm:smem) := - {| pre:=(fun sge rs m => smem_eval sge st.(the_smem) rs m <> None /\ (st.(pre) sge rs m)); + {| pre:=(fun sge sp rs m => smem_eval sge sp st.(the_smem) rs m <> None /\ (st.(pre) sge sp rs m)); the_sreg:= st.(the_sreg); the_smem:= sm |}. @@ -337,9 +333,9 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := | _ => None (* TODO jumptable ? *) end. -Lemma list_sval_eval_inj sge l rs0 m0 (sreg: reg -> sval) rs: - (forall r : reg, sval_eval sge (sreg r) rs0 m0 = Some (rs # r)) -> - list_sval_eval sge (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Lemma list_sval_eval_inj sge sp l rs0 m0 (sreg: reg -> sval) rs: + (forall r : reg, sval_eval sge sp (sreg r) rs0 m0 = Some (rs # r)) -> + list_sval_eval sge sp (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; auto. inversion_SOME v. @@ -348,10 +344,10 @@ Proof. try_simplify_someHyps. Qed. -Lemma symb_istep_preserv_early_exits i sge st rs0 m0 rs m pc st': - sem_early_exit sge st.(exits) rs0 m0 rs m pc -> +Lemma symb_istep_preserv_early_exits i sge sp st rs0 m0 rs m pc st': + sem_early_exit sge sp st.(exits) rs0 m0 rs m pc -> symb_istep i st = Some st' -> - sem_early_exit sge st'.(exits) rs0 m0 rs m pc. + sem_early_exit sge sp st'.(exits) rs0 m0 rs m pc. Proof. intros H. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -359,9 +355,9 @@ Proof. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. -Lemma sreg_set_preserv_has_aborted sge st rs0 m0 r sv: - has_aborted_on_basic sge st rs0 m0 -> - has_aborted_on_basic sge (sreg_set st r sv) rs0 m0. +Lemma sreg_set_preserv_has_aborted sge sp st rs0 m0 r sv: + has_aborted_on_basic sge sp st rs0 m0 -> + has_aborted_on_basic sge sp (sreg_set st r sv) rs0 m0. Proof. unfold has_aborted_on_basic. simpl; intuition. destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST. @@ -369,16 +365,16 @@ Proof. - right. right. exists r1. rewrite HTEST. auto. Qed. -Lemma smem_set_preserv_has_aborted sge st rs0 m0 m: - has_aborted_on_basic sge st rs0 m0 -> - has_aborted_on_basic sge (smem_set st m) rs0 m0. +Lemma smem_set_preserv_has_aborted sge sp st rs0 m0 m: + has_aborted_on_basic sge sp st rs0 m0 -> + has_aborted_on_basic sge sp (smem_set st m) rs0 m0. Proof. unfold has_aborted_on_basic. simpl; intuition. Qed. -Lemma symb_istep_preserv_has_aborted i sge rs0 m0 st st': +Lemma symb_istep_preserv_has_aborted i sge sp rs0 m0 st st': symb_istep i st = Some st' -> - has_aborted sge st rs0 m0 -> has_aborted sge st' rs0 m0. + has_aborted sge sp st rs0 m0 -> has_aborted sge sp st' rs0 m0. Proof. Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted. unfold has_aborted. @@ -402,11 +398,10 @@ Proof. intro H; inversion_clear H; simpl; auto. Qed. -Lemma symb_istep_correct sge i st rs0 m0 rs m: - sem_local_istate sge st.(curr) rs0 m0 rs m -> - has_not_yet_exit sge st.(exits) rs0 m0 -> - sem_option2_istate sge (symb_istep i st) rs0 m0 - (istep sge.(the_ge) i sge.(the_sp) rs m). +Lemma symb_istep_correct sge sp i st rs0 m0 rs m: + sem_local_istate sge sp st.(curr) rs0 m0 rs m -> + has_not_yet_exit sge sp st.(exits) rs0 m0 -> + sem_option2_istate sge sp (symb_istep i st) rs0 m0 (istep sge.(the_ge) i sp rs m). Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. @@ -482,10 +477,10 @@ Proof. try_simplify_someHyps. } Qed. -Lemma symb_istep_correct_None sge i st rs0 m0 rs m: - sem_local_istate sge (st.(curr)) rs0 m0 rs m -> +Lemma symb_istep_correct_None sge sp i st rs0 m0 rs m: + sem_local_istate sge sp (st.(curr)) rs0 m0 rs m -> symb_istep i st = None -> - istep sge.(the_ge) i sge.(the_sp) rs m = None. + istep sge.(the_ge) i sp rs m = None. Proof. intros (PRE & MEM & REG). destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -501,11 +496,11 @@ Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := symb_isteps p f st1 end. -Lemma symb_isteps_correct_false sge path f rs0 m0 st' is: +Lemma symb_isteps_correct_false sge sp path f rs0 m0 st' is: is.(continue)=false -> - forall st, sem_istate sge st rs0 m0 is -> + forall st, sem_istate sge sp st rs0 m0 is -> symb_isteps path f st = Some st' -> - sem_istate sge st' rs0 m0 is. + sem_istate sge sp st' rs0 m0 is. Proof. Local Hint Resolve symb_istep_preserv_early_exits. intros CONT; unfold sem_istate; rewrite CONT. @@ -515,9 +510,9 @@ Proof. inversion_SOME st1; eauto. Qed. -Lemma symb_isteps_preserv_has_aborted sge path f rs0 m0 st': forall st, +Lemma symb_isteps_preserv_has_aborted sge sp path f rs0 m0 st': forall st, symb_isteps path f st = Some st' -> - has_aborted sge st rs0 m0 -> has_aborted sge st' rs0 m0. + has_aborted sge sp st rs0 m0 -> has_aborted sge sp st' rs0 m0. Proof. Local Hint Resolve symb_istep_preserv_has_aborted. induction path; simpl. @@ -547,12 +542,12 @@ Proof. intros; erewrite symb_istep_default_succ; eauto. Qed. -Lemma symb_isteps_correct_true sge path (f:function) rs0 m0: forall st is, +Lemma symb_isteps_correct_true sge sp path (f:function) rs0 m0: forall st is, is.(continue)=true -> - sem_istate sge st rs0 m0 is -> + sem_istate sge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(the_pc) <> None -> - sem_option2_istate sge (symb_isteps path f st) rs0 m0 - (isteps sge.(the_ge) path f sge.(the_sp) is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) + sem_option2_istate sge sp (symb_isteps path f st) rs0 m0 + (isteps sge.(the_ge) path f sp is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) . Proof. Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF. @@ -624,21 +619,21 @@ Inductive sfval := | Sreturn: option sval -> sfval . -Definition s_find_function (sge : sgenv) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := +Definition s_find_function (sge: sgenv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := match svos with - | inl sv => SOME v <- sval_eval sge sv rs0 m0 IN Genv.find_funct sge.(the_pge) v + | inl sv => SOME v <- sval_eval sge sp sv rs0 m0 IN Genv.find_funct sge.(the_pge) v | inr symb => SOME b <- Genv.find_symbol sge.(the_pge) symb IN Genv.find_funct_ptr sge.(the_pge) b end. -Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive sfval_sem (sge: sgenv) (sp:val) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: - sfval_sem sge st stack f rs0 m0 Snone rs m E0 (State stack f sge.(the_sp) st.(the_pc) rs m) + sfval_sem sge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(the_pc) rs m) | exec_Scall rs m sig svos lsv args res pc fd: - s_find_function sge svos rs0 m0 = Some fd -> + s_find_function sge sp svos rs0 m0 = Some fd -> funsig fd = sig -> - list_sval_eval sge lsv rs0 m0 = Some args -> - sfval_sem sge st stack f rs0 m0 (Scall sig svos lsv res pc) rs m - E0 (Callstate (Stackframe res f sge.(the_sp) pc rs :: stack) fd args m) + list_sval_eval sge sp lsv rs0 m0 = Some args -> + sfval_sem sge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) (* TODO: | exec_Itailcall stk pc rs m sig ros args fd m': (fn_code f)!pc = Some(Itailcall sig ros args) -> @@ -661,25 +656,25 @@ Inductive sfval_sem (sge: sgenv) (st: s_istate) stack (f: function) (rs0: regset E0 (State stack f sp pc' rs m) *) | exec_Sreturn stk osv rs m m' v: - sge.(the_sp) = (Vptr stk Ptrofs.zero) -> + sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - match osv with Some sv => sval_eval sge sv rs0 m0 | None => Some Vundef end = Some v -> - sfval_sem sge st stack f rs0 m0 (Sreturn osv) rs m + match osv with Some sv => sval_eval sge sp sv rs0 m0 | None => Some Vundef end = Some v -> + sfval_sem sge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . Record s_state := { internal:> s_istate; final: sfval }. -Inductive sem_state (sge: sgenv) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := +Inductive sem_state (sge: sgenv) (sp:val) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | sem_early is: is.(continue) = false -> - sem_istate sge st rs0 m0 is -> - sem_state sge st stack f rs0 m0 E0 (State stack f sge.(the_sp) is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) + sem_istate sge sp st rs0 m0 is -> + sem_state sge sp st stack f rs0 m0 E0 (State stack f sp is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) | sem_normal is t s: is.(continue) = true -> - sem_istate sge st rs0 m0 is -> - sfval_sem sge st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> - sem_state sge st stack f rs0 m0 t s + sem_istate sge sp st rs0 m0 is -> + sfval_sem sge sp st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> + sem_state sge sp st stack f rs0 m0 t s . (** * Symbolic execution of final step *) @@ -698,10 +693,10 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := Lemma symb_fstep_correct i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> - sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> + sem_local_istate (Build_sgenv pge ge) sp (curr st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> symb_istep i st = None -> - sfval_sem (Build_sgenv pge ge sp) st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. + sfval_sem (Build_sgenv pge ge) sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + (* Snone *) destruct i; simpl in Hi |- *; unfold sist in Hi; try congruence. @@ -719,8 +714,8 @@ Admitted. Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> - sem_local_istate (Build_sgenv pge ge sp) (curr st) rs0 m0 rs m -> - sfval_sem (Build_sgenv pge ge sp) st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> + sem_local_istate (Build_sgenv pge ge) sp (curr st) rs0 m0 rs m -> + sfval_sem (Build_sgenv pge ge) sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> symb_istep i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. @@ -744,11 +739,11 @@ Admitted. (** * Main function of the symbolic execution *) -Definition init_local_istate := {| pre:= fun _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. +Definition init_local_istate := {| pre:= fun _ _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. Definition init_istate pc := {| the_pc:= pc; exits:=nil; curr:= init_local_istate |}. -Lemma init_sem_istate sge pc rs m: sem_istate sge (init_istate pc) rs m (ist true pc rs m). +Lemma init_sem_istate sge sp pc rs m: sem_istate sge sp (init_istate pc) rs m (ist true pc rs m). Proof. unfold sem_istate, sem_local_istate, has_not_yet_exit; simpl. intuition. Qed. @@ -786,12 +781,12 @@ Qed. Theorem symb_step_correct f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> path_step ge pge path stack f sp rs m pc t s -> - exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge sp) st stack f rs m t s. + exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge) sp st stack f rs m t s. Proof. Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. intros PATH STEP; unfold symb_step; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true (Build_sgenv pge ge) sp path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST]; @@ -895,10 +890,10 @@ Proof. Qed. *) -Lemma sfval_sem_equiv sge (f:function) st sv stack rs0 m0 t rs1 rs2 m s: - sfval_sem sge st stack f rs0 m0 sv rs1 m t s -> +Lemma sfval_sem_equiv sge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + sfval_sem sge sp st stack f rs0 m0 sv rs1 m t s -> (forall r, rs1#r = rs2#r) -> - exists s', equiv_state s s' /\ sfval_sem sge st stack f rs0 m0 sv rs2 m t s'. + exists s', equiv_state s s' /\ sfval_sem sge sp st stack f rs0 m0 sv rs2 m t s'. Proof. Local Hint Resolve equiv_stack_refl. destruct 1. @@ -920,14 +915,14 @@ Qed. Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> - sem_state (Build_sgenv pge ge sp) st stack f rs m t s1 -> + sem_state (Build_sgenv pge ge) sp st stack f rs m t s1 -> exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ equiv_state s1 s2. Proof. Local Hint Resolve init_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true (Build_sgenv pge ge sp) path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true (Build_sgenv pge ge) sp path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). unfold nth_default_succ_inst in Hi. @@ -969,11 +964,9 @@ Admitted. (** * Simulation of RTLpath code w.r.t symbolic execution *) -Definition istate_simu (srce: PTree.t node) ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). +Section SymbStepSimuDef. + +(* Definition local_istate_simu st1 st2 := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. *) (* TODO: what is it the good notion ? *) Parameter s_istate_simu: forall (srce: PTree.t node) (st1 st2: s_istate), Prop. @@ -994,10 +987,11 @@ Definition s_state_simu srce (s1 s2: s_state): Prop := Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. -(* IDEA: we will provide an efficient test for checking "path_simu" property, by hash-consing. - See usage of [symb_step_simu] in [RTLpathScheduler]. - +(* IDEA: we will provide an efficient test for checking "symb_step_simu" property, by hash-consing. + See usage of [symb_step_simu] in [RTLpathScheduler]. *) +End SymbStepSimuDef. + (* TODO: à déplacer dans RTLpath ? *) Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 6cca7ceb..d8cdca71 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -4,7 +4,7 @@ This module is inspired from [Duplicate] and [Duplicateproof] *) -Require Import AST Linking Values Maps Globalenvs Smallstep. +Require Import AST Linking Values Maps Globalenvs Smallstep Registers. Require Import Coqlib Maps Events Errors Op. Require Import RTL RTLpath RTLpathSE_theory. @@ -245,11 +245,11 @@ Proof. Qed. Lemma sval_eval_translated sp sv rs0 m0: - sval_eval (Build_sgenv pge ge sp) sv rs0 m0 = sval_eval (Build_sgenv tpge tge sp) sv rs0 m0. + sval_eval (Build_sgenv pge ge) sp sv rs0 m0 = sval_eval (Build_sgenv tpge tge) sp sv rs0 m0. Proof. Local Hint Resolve symbols_preserved_RTL. - induction sv using sval_mut with (P0 := fun lsv => list_sval_eval (Build_sgenv pge ge sp) lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge sp) lsv rs0 m0) - (P1 := fun sm => smem_eval (Build_sgenv pge ge sp) sm rs0 m0 = smem_eval (Build_sgenv tpge tge sp) sm rs0 m0); simpl; auto. + induction sv using sval_mut with (P0 := fun lsv => list_sval_eval (Build_sgenv pge ge) sp lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge) sp lsv rs0 m0) + (P1 := fun sm => smem_eval (Build_sgenv pge ge) sp sm rs0 m0 = smem_eval (Build_sgenv tpge tge) sp sm rs0 m0); simpl; auto. + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. erewrite eval_operation_preserved; eauto. @@ -267,7 +267,7 @@ Proof. Qed. Lemma list_sval_eval_translated sp lsv rs0 m0: - list_sval_eval (Build_sgenv pge ge sp) lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge sp) lsv rs0 m0. + list_sval_eval (Build_sgenv pge ge) sp lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge) sp lsv rs0 m0. Proof. induction lsv; simpl; auto. rewrite sval_eval_translated. destruct (sval_eval _ _ _ _); auto. @@ -275,8 +275,8 @@ Proof. Qed. Lemma s_find_function_translated sp svos rs0 m0 fd: - s_find_function (Build_sgenv pge ge sp) svos rs0 m0 = Some fd -> - exists fd', s_find_function (Build_sgenv tpge tge sp) svos rs0 m0 = Some fd' + s_find_function (Build_sgenv pge ge) sp svos rs0 m0 = Some fd -> + exists fd', s_find_function (Build_sgenv tpge tge) sp svos rs0 m0 = Some fd' /\ transf_fundef fd = OK fd'. Proof. unfold s_find_function. destruct svos; simpl. @@ -287,12 +287,18 @@ Proof. intros; apply function_ptr_translated; auto. Qed. +Definition istate_simu (srce: PTree.t node) ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + Lemma sem_istate_simu dupmap sp f f' st st' rs m is: match_function dupmap f f' -> - sem_istate (Build_sgenv pge ge sp) st rs m is -> + sem_istate (Build_sgenv pge ge) sp st rs m is -> s_istate_simu dupmap st st' -> exists is', - sem_istate (Build_sgenv tpge tge sp) st' rs m is' /\ + sem_istate (Build_sgenv tpge tge) sp st' rs m is' /\ istate_simu dupmap is is'. Admitted. @@ -301,8 +307,8 @@ Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: list_forall2 match_stackframes stk stk' -> s_istate_simu dupmap st st' -> sfval_simu dupmap st.(the_pc) st'.(the_pc) sv sv' -> - sfval_sem (Build_sgenv pge ge sp) st stk f rs0 m0 sv rs m t s -> - exists s', sfval_sem (Build_sgenv tpge tge sp) st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + sfval_sem (Build_sgenv pge ge) sp st stk f rs0 m0 sv rs m t s -> + exists s', sfval_sem (Build_sgenv tpge tge) sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct. intros FUN STK SIS SFV. destruct SFV. @@ -326,9 +332,9 @@ Qed. Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> - sem_state (Build_sgenv pge ge sp) st stk f rs m t s -> + sem_state (Build_sgenv pge ge) sp st stk f rs m t s -> s_state_simu dupmap st st' -> - exists s', sem_state (Build_sgenv tpge tge sp) st' stk' f' rs m t s' /\ match_states s s'. + exists s', sem_state (Build_sgenv tpge tge) sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros FUNC STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) -- cgit From 9eb500b6bdf29b84ccc48ce5a2aa59bb7fb575db Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 17:00:38 +0100 Subject: progress on the interface of RTLpathSE and RTLpathScheduler --- mppa_k1c/lib/RTLpathSE_theory.v | 40 +++++++++++++++++++++++++--- mppa_k1c/lib/RTLpathScheduler.v | 59 +++++++++++------------------------------ 2 files changed, 51 insertions(+), 48 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index b3a7e6ad..33095a57 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -964,9 +964,43 @@ Admitted. (** * Simulation of RTLpath code w.r.t symbolic execution *) -Section SymbStepSimuDef. +Section SymbValPreserved. -(* Definition local_istate_simu st1 st2 := { pre: sgenv -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. *) +Variable sge sge': sgenv. + +Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol sge'.(the_ge) s = Genv.find_symbol sge.(the_ge) s. + +Lemma sval_eval_preserved sp sv rs0 m0: + sval_eval sge sp sv rs0 m0 = sval_eval sge' sp sv rs0 m0. +Proof. + Local Hint Resolve symbols_preserved_RTL. + induction sv using sval_mut with (P0 := fun lsv => list_sval_eval sge sp lsv rs0 m0 = list_sval_eval sge' sp lsv rs0 m0) + (P1 := fun sm => smem_eval sge sp sm rs0 m0 = smem_eval sge' sp sm rs0 m0); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. + erewrite eval_operation_preserved; eauto. + + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; auto. + + rewrite IHsv; clear IHsv. destruct (sval_eval _ _ _ _); auto. + rewrite IHsv0; auto. + + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsv; clear IHsv. destruct (smem_eval _ _ _ _); auto. + rewrite IHsv1; auto. +Qed. + +Lemma list_sval_eval_preserved sp lsv rs0 m0: + list_sval_eval sge sp lsv rs0 m0 = list_sval_eval sge' sp lsv rs0 m0. +Proof. + induction lsv; simpl; auto. + rewrite sval_eval_preserved. destruct (sval_eval _ _ _ _); auto. + rewrite IHlsv; auto. +Qed. + +End SymbValPreserved. (* TODO: what is it the good notion ? *) Parameter s_istate_simu: forall (srce: PTree.t node) (st1 st2: s_istate), Prop. @@ -990,8 +1024,6 @@ Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := (* IDEA: we will provide an efficient test for checking "symb_step_simu" property, by hash-consing. See usage of [symb_step_simu] in [RTLpathScheduler]. *) -End SymbStepSimuDef. - (* TODO: à déplacer dans RTLpath ? *) Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index d8cdca71..ca0c0634 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -152,7 +152,7 @@ Proof. eapply (Genv.senv_match TRANSL). Qed. -Lemma functions_translated: +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. @@ -164,7 +164,7 @@ Proof. + eapply linkorder_refl. Qed. -Lemma function_ptr_translated: +Lemma function_ptr_preserved: forall v f, Genv.find_funct_ptr pge v = Some f -> exists tf, @@ -174,7 +174,7 @@ Proof. exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto. Qed. -Lemma function_sig_translated: +Lemma function_sig_preserved: forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. Proof. intros. destruct f. @@ -189,7 +189,7 @@ Theorem transf_initial_states: exists s2, initial_state tprog s2 /\ match_states s1 s2. Proof. intros. inv H. - exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF). + exploit function_ptr_preserved; eauto. intros (tf & FIND & TRANSF). exists (Callstate nil tf nil m0). split. - econstructor; eauto. @@ -244,47 +244,19 @@ Proof. erewrite RTLpath.symbols_preserved; eauto. Qed. -Lemma sval_eval_translated sp sv rs0 m0: - sval_eval (Build_sgenv pge ge) sp sv rs0 m0 = sval_eval (Build_sgenv tpge tge) sp sv rs0 m0. -Proof. - Local Hint Resolve symbols_preserved_RTL. - induction sv using sval_mut with (P0 := fun lsv => list_sval_eval (Build_sgenv pge ge) sp lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge) sp lsv rs0 m0) - (P1 := fun sm => smem_eval (Build_sgenv pge ge) sp sm rs0 m0 = smem_eval (Build_sgenv tpge tge) sp sm rs0 m0); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. - erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing tge sp _ _); auto. - rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (sval_eval _ _ _ _); auto. - rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. - erewrite <- eval_addressing_preserved; eauto. - destruct (eval_addressing tge sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (smem_eval _ _ _ _); auto. - rewrite IHsv1; auto. -Qed. - -Lemma list_sval_eval_translated sp lsv rs0 m0: - list_sval_eval (Build_sgenv pge ge) sp lsv rs0 m0 = list_sval_eval (Build_sgenv tpge tge) sp lsv rs0 m0. -Proof. - induction lsv; simpl; auto. - rewrite sval_eval_translated. destruct (sval_eval _ _ _ _); auto. - rewrite IHlsv; auto. -Qed. - -Lemma s_find_function_translated sp svos rs0 m0 fd: +Lemma s_find_function_preserved sp svos rs0 m0 fd: s_find_function (Build_sgenv pge ge) sp svos rs0 m0 = Some fd -> exists fd', s_find_function (Build_sgenv tpge tge) sp svos rs0 m0 = Some fd' /\ transf_fundef fd = OK fd'. Proof. + Local Hint Resolve symbols_preserved_RTL. unfold s_find_function. destruct svos; simpl. - + rewrite sval_eval_translated. destruct (sval_eval _ _ _ _); try congruence. - intros; exploit functions_translated; eauto. + + rewrite (sval_eval_preserved (Build_sgenv pge ge) (Build_sgenv tpge tge)); auto. + destruct (sval_eval _ _ _ _); try congruence. + intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. - intros; apply function_ptr_translated; auto. + intros; apply function_ptr_preserved; auto. Qed. Definition istate_simu (srce: PTree.t node) ist1 ist2 := @@ -293,8 +265,7 @@ Definition istate_simu (srce: PTree.t node) ist1 ist2 := (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ ist1.(the_mem) = ist2.(the_mem). -Lemma sem_istate_simu dupmap sp f f' st st' rs m is: - match_function dupmap f f' -> +Lemma sem_istate_simu dupmap sp st st' rs m is: sem_istate (Build_sgenv pge ge) sp st rs m is -> s_istate_simu dupmap st st' -> exists is', @@ -315,17 +286,17 @@ Proof. - (* Snone *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. - (* Scall *) intros SEM; inv SEM; simpl. - exploit s_find_function_translated; eauto. + exploit s_find_function_preserved; eauto. intros (fd' & FIND & TRANSF). - erewrite <- function_sig_translated; eauto. + erewrite <- function_sig_preserved; eauto. eexists; split; econstructor; eauto. - + erewrite <- list_sval_eval_translated; auto. + + erewrite <- (list_sval_eval_preserved (Build_sgenv pge ge) (Build_sgenv tpge tge)); auto. + simpl. repeat (econstructor; eauto). - (* Sreturn *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. - erewrite <- sval_eval_translated; eauto. + erewrite <- sval_eval_preserved; eauto. Qed. (* The main theorem on simulation of symbolic states ! *) -- cgit From b18a6724a5669d2beb4d1d5cae9572ef274296e9 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 26 Nov 2019 18:15:42 +0100 Subject: cleaning the interface between RTLpathSE and RTLpathScheduler --- mppa_k1c/lib/RTLpathSE_theory.v | 341 +++++++++++++++++++++------------------- mppa_k1c/lib/RTLpathScheduler.v | 30 ++-- 2 files changed, 195 insertions(+), 176 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 33095a57..7f94ae58 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -20,12 +20,6 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ (** * Syntax and semantics of symbolic values *) - -Record sgenv := { - the_pge: RTLpath.genv; - the_ge: RTL.genv -}. - (* symbolic value *) Inductive sval := | Sinput (r: reg) @@ -51,99 +45,99 @@ Fixpoint list_val_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. -Fixpoint sval_eval (sge: sgenv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := +Fixpoint sval_eval (ge: RTL.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 <- list_sval_eval sge sp l rs0 m0 IN - SOME m <- smem_eval sge sp sm rs0 m0 IN - eval_operation sge.(the_ge) sp op args m + SOME args <- list_sval_eval ge sp l rs0 m0 IN + SOME m <- smem_eval ge sp sm rs0 m0 IN + eval_operation ge sp op args m | Sload sm chunk addr lsv => - SOME args <- list_sval_eval sge sp lsv rs0 m0 IN - SOME a <- eval_addressing sge.(the_ge) sp addr args IN - SOME m <- smem_eval sge sp sm rs0 m0 IN + SOME args <- list_sval_eval ge sp lsv rs0 m0 IN + SOME a <- eval_addressing ge sp addr args IN + SOME m <- smem_eval ge sp sm rs0 m0 IN Mem.loadv chunk m a end -with list_sval_eval (sge: sgenv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := +with list_sval_eval (ge: RTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := match lsv with | Snil => Some nil | Scons sv lsv' => - SOME v <- sval_eval sge sp sv rs0 m0 IN - SOME lv <- list_sval_eval sge sp lsv' rs0 m0 IN + SOME v <- sval_eval ge sp sv rs0 m0 IN + SOME lv <- list_sval_eval ge sp lsv' rs0 m0 IN Some (v::lv) end -with smem_eval (sge: sgenv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem := +with smem_eval (ge: RTL.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 <- list_sval_eval sge sp lsv rs0 m0 IN - SOME a <- eval_addressing sge.(the_ge) sp addr args IN - SOME m <- smem_eval sge sp sm rs0 m0 IN - SOME sv <- sval_eval sge sp srce rs0 m0 IN + SOME args <- list_sval_eval ge sp lsv rs0 m0 IN + SOME a <- eval_addressing ge sp addr args IN + SOME m <- smem_eval ge sp sm rs0 m0 IN + SOME sv <- sval_eval ge sp srce rs0 m0 IN Mem.storev chunk m a sv end. (* Syntax and Semantics of local symbolic internal states *) -Record local_istate := { pre: sgenv -> val -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. +Record local_istate := { pre: RTL.genv -> val -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. -Definition sem_local_istate (sge: sgenv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := - st.(pre) sge sp rs0 m0 - /\ smem_eval sge sp st.(the_smem) rs0 m0 = Some m - /\ forall (r:reg), sval_eval sge sp (st.(the_sreg) r) rs0 m0 = Some (rs#r). +Definition sem_local_istate (ge: RTL.genv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := + st.(pre) ge sp rs0 m0 + /\ smem_eval ge sp st.(the_smem) rs0 m0 = Some m + /\ forall (r:reg), sval_eval ge sp (st.(the_sreg) r) rs0 m0 = Some (rs#r). -Definition has_aborted_on_basic (sge: sgenv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem): Prop := - ~(st.(pre) sge sp rs0 m0) - \/ smem_eval sge sp st.(the_smem) rs0 m0 = None - \/ exists (r: reg), sval_eval sge sp (st.(the_sreg) r) rs0 m0 = None. +Definition has_aborted_on_basic (ge: RTL.genv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem): Prop := + ~(st.(pre) ge sp rs0 m0) + \/ smem_eval ge sp st.(the_smem) rs0 m0 = None + \/ exists (r: reg), sval_eval ge sp (st.(the_sreg) r) rs0 m0 = None. (* Syntax and semantics of symbolic exit states *) Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. -Definition seval_condition sge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := - SOME args <- list_sval_eval sge sp lsv rs0 m0 IN - SOME m <- smem_eval sge sp sm rs0 m0 IN +Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := + SOME args <- list_sval_eval ge sp lsv rs0 m0 IN + SOME m <- smem_eval ge sp sm rs0 m0 IN eval_condition cond args m. -Definition has_not_yet_exit sge sp (lx: list exit_istate) rs0 m0: Prop := - forall ext, List.In ext lx -> seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. +Definition has_not_yet_exit ge sp (lx: list exit_istate) rs0 m0: Prop := + forall ext, List.In ext lx -> seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. -Definition has_aborted_on_test sge sp (lx: list exit_istate) rs0 m0: Prop := - exists ext, List.In ext lx /\ seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. +Definition has_aborted_on_test ge sp (lx: list exit_istate) rs0 m0: Prop := + exists ext, List.In ext lx /\ seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. -Inductive sem_early_exit (sge: sgenv) (sp:val) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := +Inductive sem_early_exit (ge: RTL.genv) (sp:val) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := SEE_intro ext lx': is_tail (ext::lx') lx -> - has_not_yet_exit sge sp lx' rs0 m0 -> - seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> - sem_local_istate sge sp ext.(exit_ist) rs0 m0 rs m -> + has_not_yet_exit ge sp lx' rs0 m0 -> + seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> + sem_local_istate ge sp ext.(exit_ist) rs0 m0 rs m -> ext.(cond_ifso) = pc -> - sem_early_exit sge sp lx rs0 m0 rs m pc. + sem_early_exit ge sp lx rs0 m0 rs m pc. (** * Syntax and Semantics of symbolic internal state *) Record s_istate := { the_pc: node; exits: list exit_istate; curr: local_istate }. -Definition sem_istate (sge: sgenv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := +Definition sem_istate (ge: RTL.genv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(continue)) then - sem_local_istate sge sp st.(curr) rs0 m0 is.(the_rs) is.(the_mem) + sem_local_istate ge sp st.(curr) rs0 m0 is.(the_rs) is.(the_mem) /\ st.(the_pc) = is.(RTLpath.the_pc) - /\ has_not_yet_exit sge sp st.(exits) rs0 m0 - else sem_early_exit sge sp st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). + /\ has_not_yet_exit ge sp st.(exits) rs0 m0 + else sem_early_exit ge sp st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). -Definition has_aborted (sge: sgenv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem): Prop := - has_aborted_on_basic sge sp st.(curr) rs0 m0 - \/ has_aborted_on_test sge sp st.(exits) rs0 m0. +Definition has_aborted (ge: RTL.genv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem): Prop := + has_aborted_on_basic ge sp st.(curr) rs0 m0 + \/ has_aborted_on_test ge sp st.(exits) rs0 m0. -Definition sem_option_istate sge sp (st: s_istate) rs0 m0 (ois: option istate): Prop := +Definition sem_option_istate ge sp (st: s_istate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => sem_istate sge sp st rs0 m0 is - | None => has_aborted sge sp st rs0 m0 + | Some is => sem_istate ge sp st rs0 m0 is + | None => has_aborted ge sp st rs0 m0 end. -Definition sem_option2_istate sge sp (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := +Definition sem_option2_istate ge sp (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => sem_option_istate sge sp st rs0 m0 ois + | Some st => sem_option_istate ge sp st rs0 m0 ois | None => ois=None end. @@ -160,17 +154,17 @@ Definition istate_eq ist1 ist2 := ist1.(the_mem) = ist2.(the_mem). (* TODO: is it useful -Lemma has_not_yet_exit_cons_split sge sp ext lx rs0 m0: - has_not_yet_exit sge sp (ext::lx) rs0 m0 <-> - (seval_condition sge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit sge sp lx rs0 m0). +Lemma has_not_yet_exit_cons_split ge sp ext lx rs0 m0: + has_not_yet_exit ge sp (ext::lx) rs0 m0 <-> + (seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit ge sp lx rs0 m0). Proof. unfold has_not_yet_exit; simpl; intuition (subst; auto). Qed. *) -Lemma exclude_early_NYE sge sp st rs0 m0 rs m pc: - sem_early_exit sge sp (exits st) rs0 m0 rs m pc -> - has_not_yet_exit sge sp (exits st) rs0 m0 -> +Lemma exclude_early_NYE ge sp st rs0 m0 rs m pc: + sem_early_exit ge sp (exits st) rs0 m0 rs m pc -> + has_not_yet_exit ge sp (exits st) rs0 m0 -> False. Proof. Local Hint Resolve is_tail_in. @@ -179,11 +173,11 @@ Proof. try congruence. Qed. -Lemma sem_istate_exclude_incompatible_continue sge sp st rs m is1 is2: +Lemma sem_istate_exclude_incompatible_continue ge sp st rs m is1 is2: is1.(continue) = true -> is2.(continue) = false -> - sem_istate sge sp st rs m is1 -> - sem_istate sge sp st rs m is2 -> + sem_istate ge sp st rs m is1 -> + sem_istate ge sp st rs m is2 -> False. Proof. Local Hint Resolve exclude_early_NYE. @@ -193,9 +187,9 @@ Proof. intuition eauto. Qed. -Lemma sem_istate_determ_continue sge sp st rs m is1 is2: - sem_istate sge sp st rs m is1 -> - sem_istate sge sp st rs m is2 -> +Lemma sem_istate_determ_continue ge sp st rs m is1 is2: + sem_istate ge sp st rs m is1 -> + sem_istate ge sp st rs m is2 -> is1.(continue) = is2.(continue). Proof. Local Hint Resolve sem_istate_exclude_incompatible_continue. @@ -204,9 +198,9 @@ Proof. destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. Qed. -Lemma sem_local_istate_determ sge sp st rs0 m0 rs1 m1 rs2 m2: - sem_local_istate sge sp st rs0 m0 rs1 m1 -> - sem_local_istate sge sp st rs0 m0 rs2 m2 -> +Lemma sem_local_istate_determ ge sp st rs0 m0 rs1 m1 rs2 m2: + sem_local_istate ge sp st rs0 m0 rs1 m1 -> + sem_local_istate ge sp st rs0 m0 rs2 m2 -> (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. unfold sem_local_istate. intuition try congruence. @@ -221,13 +215,13 @@ Proof. intros T2; inversion T2; subst; auto. Qed. -Lemma exit_cond_determ sge sp rs0 m0 l1 l2: +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) -> - has_not_yet_exit sge sp lx1 rs0 m0 -> - seval_condition sge sp (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> - has_not_yet_exit sge sp lx2 rs0 m0 -> + has_not_yet_exit ge sp lx1 rs0 m0 -> + seval_condition ge sp (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> + has_not_yet_exit ge sp lx2 rs0 m0 -> ext1=ext2. Proof. destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst; @@ -239,9 +233,9 @@ Proof. try congruence. Qed. -Lemma sem_early_exit_determ sge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : - sem_early_exit sge sp lx rs0 m0 rs1 m1 pc1 -> - sem_early_exit sge sp lx rs0 m0 rs2 m2 pc2 -> +Lemma sem_early_exit_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : + sem_early_exit ge sp lx rs0 m0 rs1 m1 pc1 -> + sem_early_exit ge sp lx rs0 m0 rs2 m2 pc2 -> pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. Local Hint Resolve exit_cond_determ eq_sym. @@ -249,28 +243,28 @@ Proof. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. assert (X:ext1=ext2). - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. - - subst. destruct (sem_local_istate_determ sge sp (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. + - subst. destruct (sem_local_istate_determ ge sp (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. -Lemma sem_istate_determ sge sp st rs m is1 is2: - sem_istate sge sp st rs m is1 -> - sem_istate sge sp st rs m is2 -> +Lemma sem_istate_determ ge sp st rs m is1 is2: + sem_istate ge sp st rs m is1 -> + sem_istate ge sp st rs m is2 -> istate_eq is1 is2. Proof. unfold istate_eq. intros SEM1 SEM2. - exploit (sem_istate_determ_continue sge sp st rs m is1 is2); eauto. + exploit (sem_istate_determ_continue ge sp st rs m is1 is2); eauto. intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. destruct (continue is2). - - destruct (sem_local_istate_determ sge sp (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); + - destruct (sem_local_istate_determ ge sp (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); intuition (try congruence). - - destruct (sem_early_exit_determ sge sp (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); + - destruct (sem_early_exit_determ ge sp (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); intuition. Qed. -Lemma sem_istate_exclude_abort sge sp st rs m is: - has_aborted sge sp st rs m -> - sem_istate sge sp st rs m is -> False. +Lemma sem_istate_exclude_abort ge sp st rs m is: + has_aborted ge sp st rs m -> + sem_istate ge sp st rs m is -> False. Proof. intros [[AB|[AB|AB]]|AB]; ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; @@ -281,9 +275,9 @@ Admitted. Definition istate_eq_option ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -Lemma sem_option_istate_determ sge sp st rs m ois is: - sem_option_istate sge sp st rs m ois -> - sem_istate sge sp st rs m is -> +Lemma sem_option_istate_determ ge sp st rs m ois is: + sem_option_istate ge sp st rs m ois -> + sem_istate ge sp st rs m is -> istate_eq_option is ois. Proof. destruct ois as [is1|]; simpl; eauto. @@ -294,12 +288,12 @@ Qed. (** * Symbolic execution of one internal step *) Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := - {| pre:=(fun sge sp rs m => sval_eval sge sp (st.(the_sreg) r) rs m <> None /\ (st.(pre) sge sp rs m)); + {| pre:=(fun ge sp rs m => sval_eval ge sp (st.(the_sreg) r) rs m <> None /\ (st.(pre) ge sp rs m)); the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; the_smem:= st.(the_smem)|}. Definition smem_set (st:local_istate) (sm:smem) := - {| pre:=(fun sge sp rs m => smem_eval sge sp st.(the_smem) rs m <> None /\ (st.(pre) sge sp rs m)); + {| pre:=(fun ge sp rs m => smem_eval ge sp st.(the_smem) rs m <> None /\ (st.(pre) ge sp rs m)); the_sreg:= st.(the_sreg); the_smem:= sm |}. @@ -333,9 +327,9 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := | _ => None (* TODO jumptable ? *) end. -Lemma list_sval_eval_inj sge sp l rs0 m0 (sreg: reg -> sval) rs: - (forall r : reg, sval_eval sge sp (sreg r) rs0 m0 = Some (rs # r)) -> - list_sval_eval sge sp (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Lemma list_sval_eval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: + (forall r : reg, sval_eval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> + list_sval_eval ge sp (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; auto. inversion_SOME v. @@ -344,10 +338,10 @@ Proof. try_simplify_someHyps. Qed. -Lemma symb_istep_preserv_early_exits i sge sp st rs0 m0 rs m pc st': - sem_early_exit sge sp st.(exits) rs0 m0 rs m pc -> +Lemma symb_istep_preserv_early_exits i ge sp st rs0 m0 rs m pc st': + sem_early_exit ge sp st.(exits) rs0 m0 rs m pc -> symb_istep i st = Some st' -> - sem_early_exit sge sp st'.(exits) rs0 m0 rs m pc. + sem_early_exit ge sp st'.(exits) rs0 m0 rs m pc. Proof. intros H. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -355,9 +349,9 @@ Proof. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. -Lemma sreg_set_preserv_has_aborted sge sp st rs0 m0 r sv: - has_aborted_on_basic sge sp st rs0 m0 -> - has_aborted_on_basic sge sp (sreg_set st r sv) rs0 m0. +Lemma sreg_set_preserv_has_aborted ge sp st rs0 m0 r sv: + has_aborted_on_basic ge sp st rs0 m0 -> + has_aborted_on_basic ge sp (sreg_set st r sv) rs0 m0. Proof. unfold has_aborted_on_basic. simpl; intuition. destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST. @@ -365,16 +359,16 @@ Proof. - right. right. exists r1. rewrite HTEST. auto. Qed. -Lemma smem_set_preserv_has_aborted sge sp st rs0 m0 m: - has_aborted_on_basic sge sp st rs0 m0 -> - has_aborted_on_basic sge sp (smem_set st m) rs0 m0. +Lemma smem_set_preserv_has_aborted ge sp st rs0 m0 m: + has_aborted_on_basic ge sp st rs0 m0 -> + has_aborted_on_basic ge sp (smem_set st m) rs0 m0. Proof. unfold has_aborted_on_basic. simpl; intuition. Qed. -Lemma symb_istep_preserv_has_aborted i sge sp rs0 m0 st st': +Lemma symb_istep_preserv_has_aborted i ge sp rs0 m0 st st': symb_istep i st = Some st' -> - has_aborted sge sp st rs0 m0 -> has_aborted sge sp st' rs0 m0. + has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. Proof. Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted. unfold has_aborted. @@ -398,10 +392,10 @@ Proof. intro H; inversion_clear H; simpl; auto. Qed. -Lemma symb_istep_correct sge sp i st rs0 m0 rs m: - sem_local_istate sge sp st.(curr) rs0 m0 rs m -> - has_not_yet_exit sge sp st.(exits) rs0 m0 -> - sem_option2_istate sge sp (symb_istep i st) rs0 m0 (istep sge.(the_ge) i sp rs m). +Lemma symb_istep_correct ge sp i st rs0 m0 rs m: + sem_local_istate ge sp st.(curr) rs0 m0 rs m -> + has_not_yet_exit ge sp st.(exits) rs0 m0 -> + sem_option2_istate ge sp (symb_istep i st) rs0 m0 (istep ge i sp rs m). Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. @@ -477,10 +471,10 @@ Proof. try_simplify_someHyps. } Qed. -Lemma symb_istep_correct_None sge sp i st rs0 m0 rs m: - sem_local_istate sge sp (st.(curr)) rs0 m0 rs m -> +Lemma symb_istep_correct_None ge sp i st rs0 m0 rs m: + sem_local_istate ge sp (st.(curr)) rs0 m0 rs m -> symb_istep i st = None -> - istep sge.(the_ge) i sp rs m = None. + istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. @@ -496,11 +490,11 @@ Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := symb_isteps p f st1 end. -Lemma symb_isteps_correct_false sge sp path f rs0 m0 st' is: +Lemma symb_isteps_correct_false ge sp path f rs0 m0 st' is: is.(continue)=false -> - forall st, sem_istate sge sp st rs0 m0 is -> + forall st, sem_istate ge sp st rs0 m0 is -> symb_isteps path f st = Some st' -> - sem_istate sge sp st' rs0 m0 is. + sem_istate ge sp st' rs0 m0 is. Proof. Local Hint Resolve symb_istep_preserv_early_exits. intros CONT; unfold sem_istate; rewrite CONT. @@ -510,9 +504,9 @@ Proof. inversion_SOME st1; eauto. Qed. -Lemma symb_isteps_preserv_has_aborted sge sp path f rs0 m0 st': forall st, +Lemma symb_isteps_preserv_has_aborted ge sp path f rs0 m0 st': forall st, symb_isteps path f st = Some st' -> - has_aborted sge sp st rs0 m0 -> has_aborted sge sp st' rs0 m0. + has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. Proof. Local Hint Resolve symb_istep_preserv_has_aborted. induction path; simpl. @@ -542,12 +536,12 @@ Proof. intros; erewrite symb_istep_default_succ; eauto. Qed. -Lemma symb_isteps_correct_true sge sp path (f:function) rs0 m0: forall st is, +Lemma symb_isteps_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(continue)=true -> - sem_istate sge sp st rs0 m0 is -> + sem_istate ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(the_pc) <> None -> - sem_option2_istate sge sp (symb_isteps path f st) rs0 m0 - (isteps sge.(the_ge) path f sp is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) + sem_option2_istate ge sp (symb_isteps path f st) rs0 m0 + (isteps ge path f sp is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) . Proof. Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF. @@ -619,20 +613,20 @@ Inductive sfval := | Sreturn: option sval -> sfval . -Definition s_find_function (sge: sgenv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := +Definition s_find_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := match svos with - | inl sv => SOME v <- sval_eval sge sp sv rs0 m0 IN Genv.find_funct sge.(the_pge) v - | inr symb => SOME b <- Genv.find_symbol sge.(the_pge) symb IN Genv.find_funct_ptr sge.(the_pge) b + | inl sv => SOME v <- sval_eval ge sp sv rs0 m0 IN Genv.find_funct pge v + | inr symb => SOME b <- Genv.find_symbol pge symb IN Genv.find_funct_ptr pge b end. -Inductive sfval_sem (sge: sgenv) (sp:val) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive sfval_sem (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: - sfval_sem sge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(the_pc) rs m) + sfval_sem pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(the_pc) rs m) | exec_Scall rs m sig svos lsv args res pc fd: - s_find_function sge sp svos rs0 m0 = Some fd -> + s_find_function pge ge sp svos rs0 m0 = Some fd -> funsig fd = sig -> - list_sval_eval sge sp lsv rs0 m0 = Some args -> - sfval_sem sge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + list_sval_eval ge sp lsv rs0 m0 = Some args -> + sfval_sem pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) (* TODO: | exec_Itailcall stk pc rs m sig ros args fd m': @@ -658,23 +652,23 @@ Inductive sfval_sem (sge: sgenv) (sp:val) (st: s_istate) stack (f: function) (rs | 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 => sval_eval sge sp sv rs0 m0 | None => Some Vundef end = Some v -> - sfval_sem sge sp st stack f rs0 m0 (Sreturn osv) rs m + match osv with Some sv => sval_eval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> + sfval_sem pge ge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . Record s_state := { internal:> s_istate; final: sfval }. -Inductive sem_state (sge: sgenv) (sp:val) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := +Inductive sem_state pge (ge: RTL.genv) (sp:val) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | sem_early is: is.(continue) = false -> - sem_istate sge sp st rs0 m0 is -> - sem_state sge sp st stack f rs0 m0 E0 (State stack f sp is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) + sem_istate ge sp st rs0 m0 is -> + sem_state pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) | sem_normal is t s: is.(continue) = true -> - sem_istate sge sp st rs0 m0 is -> - sfval_sem sge sp st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> - sem_state sge sp st stack f rs0 m0 t s + sem_istate ge sp st rs0 m0 is -> + sfval_sem pge ge sp st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> + sem_state pge ge sp st stack f rs0 m0 t s . (** * Symbolic execution of final step *) @@ -690,13 +684,13 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := | _ => Snone end. -Lemma symb_fstep_correct i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: +Lemma symb_fstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> - sem_local_istate (Build_sgenv pge ge) sp (curr st) rs0 m0 rs m -> + sem_local_istate ge sp (curr st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> symb_istep i st = None -> - sfval_sem (Build_sgenv pge ge) sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. + sfval_sem pge ge sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + (* Snone *) destruct i; simpl in Hi |- *; unfold sist in Hi; try congruence. @@ -714,8 +708,8 @@ Admitted. Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(the_pc) -> - sem_local_istate (Build_sgenv pge ge) sp (curr st) rs0 m0 rs m -> - sfval_sem (Build_sgenv pge ge) sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> + sem_local_istate ge sp (curr st) rs0 m0 rs m -> + sfval_sem pge ge sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> symb_istep i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. @@ -743,7 +737,7 @@ Definition init_local_istate := {| pre:= fun _ _ _ _ => True; the_sreg:= fun r = Definition init_istate pc := {| the_pc:= pc; exits:=nil; curr:= init_local_istate |}. -Lemma init_sem_istate sge sp pc rs m: sem_istate sge sp (init_istate pc) rs m (ist true pc rs m). +Lemma init_sem_istate ge sp pc rs m: sem_istate ge sp (init_istate pc) rs m (ist true pc rs m). Proof. unfold sem_istate, sem_local_istate, has_not_yet_exit; simpl. intuition. Qed. @@ -781,12 +775,12 @@ Qed. Theorem symb_step_correct f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> path_step ge pge path stack f sp rs m pc t s -> - exists st, symb_step f pc = Some st /\ sem_state (Build_sgenv pge ge) sp st stack f rs m t s. + exists st, symb_step f pc = Some st /\ sem_state pge ge sp st stack f rs m t s. Proof. Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. intros PATH STEP; unfold symb_step; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true (Build_sgenv pge ge) sp path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST]; @@ -890,10 +884,10 @@ Proof. Qed. *) -Lemma sfval_sem_equiv sge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: - sfval_sem sge sp st stack f rs0 m0 sv rs1 m t s -> +Lemma sfval_sem_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + sfval_sem 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' /\ sfval_sem sge sp st stack f rs0 m0 sv rs2 m t s'. + exists s', equiv_state s s' /\ sfval_sem pge ge sp st stack f rs0 m0 sv rs2 m t s'. Proof. Local Hint Resolve equiv_stack_refl. destruct 1. @@ -915,14 +909,14 @@ Qed. Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> - sem_state (Build_sgenv pge ge) sp st stack f rs m t s1 -> + sem_state pge ge sp st stack f rs m t s1 -> exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ equiv_state s1 s2. Proof. Local Hint Resolve init_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true (Build_sgenv pge ge) sp path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). unfold nth_default_succ_inst in Hi. @@ -966,16 +960,16 @@ Admitted. Section SymbValPreserved. -Variable sge sge': sgenv. +Variable ge ge': RTL.genv. -Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol sge'.(the_ge) s = Genv.find_symbol sge.(the_ge) s. +Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. Lemma sval_eval_preserved sp sv rs0 m0: - sval_eval sge sp sv rs0 m0 = sval_eval sge' sp sv rs0 m0. + sval_eval ge sp sv rs0 m0 = sval_eval ge' sp sv rs0 m0. Proof. Local Hint Resolve symbols_preserved_RTL. - induction sv using sval_mut with (P0 := fun lsv => list_sval_eval sge sp lsv rs0 m0 = list_sval_eval sge' sp lsv rs0 m0) - (P1 := fun sm => smem_eval sge sp sm rs0 m0 = smem_eval sge' sp sm rs0 m0); simpl; auto. + induction sv using sval_mut with (P0 := fun lsv => list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0) + (P1 := fun sm => smem_eval ge sp sm rs0 m0 = smem_eval ge' sp sm rs0 m0); simpl; auto. + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. erewrite eval_operation_preserved; eauto. @@ -993,18 +987,47 @@ Proof. Qed. Lemma list_sval_eval_preserved sp lsv rs0 m0: - list_sval_eval sge sp lsv rs0 m0 = list_sval_eval sge' sp lsv rs0 m0. + list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0. Proof. induction lsv; simpl; auto. rewrite sval_eval_preserved. destruct (sval_eval _ _ _ _); auto. rewrite IHlsv; auto. Qed. +Lemma smem_eval_preserved sp sm rs0 m0: + smem_eval ge sp sm rs0 m0 = smem_eval ge' sp sm rs0 m0. +Proof. + induction sm; simpl; auto. + rewrite list_sval_eval_preserved. destruct (list_sval_eval _ _ _ _); auto. + erewrite <- eval_addressing_preserved; eauto. + destruct (eval_addressing _ sp _ _); auto. + rewrite IHsm; clear IHsm. destruct (smem_eval _ _ _ _); auto. + rewrite sval_eval_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 (list_sval_eval _ _ _ _); auto. + rewrite smem_eval_preserved; auto. +Qed. + End SymbValPreserved. -(* TODO: what is it the good notion ? *) -Parameter s_istate_simu: forall (srce: PTree.t node) (st1 st2: s_istate), Prop. +Definition istate_simu (srce: PTree.t node) ist1 ist2 := + ist1.(continue) = ist2.(continue) /\ + srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ + (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ + ist1.(the_mem) = ist2.(the_mem). + +(* NOTE: a pure semantic definition on [s_istate], for a total freedom in refinements *) +Definition s_istate_simu (srce: PTree.t node) (st1 st2: s_istate): Prop := + forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + forall rs m is1, sem_istate ge1 sp st1 rs m is1 -> + exists is2, sem_istate ge2 sp st2 rs m is2 /\ istate_simu srce is1 is2. +(* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: srce!opc2 = Some opc1 -> sfval_simu srce opc1 opc2 Snone Snone @@ -1026,7 +1049,7 @@ Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := *) (* TODO: à déplacer dans RTLpath ? *) -Lemma path_step_equiv ge pge path stk f sp rs1 m pc t s1 rs2: +Lemma path_step_equiv pge ge path stk f sp rs1 m pc t s1 rs2: path_step ge pge path stk f sp rs1 m pc t s1 -> (forall r : positive, rs1 !! r = rs2 !! r) -> exists s2, path_step ge pge path stk f sp rs2 m pc t s2 /\ equiv_state s1 s2. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index ca0c0634..bfd2384c 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -245,13 +245,13 @@ Proof. Qed. Lemma s_find_function_preserved sp svos rs0 m0 fd: - s_find_function (Build_sgenv pge ge) sp svos rs0 m0 = Some fd -> - exists fd', s_find_function (Build_sgenv tpge tge) sp svos rs0 m0 = Some fd' + s_find_function pge ge sp svos rs0 m0 = Some fd -> + exists fd', s_find_function tpge tge sp svos rs0 m0 = Some fd' /\ transf_fundef fd = OK fd'. Proof. Local Hint Resolve symbols_preserved_RTL. unfold s_find_function. destruct svos; simpl. - + rewrite (sval_eval_preserved (Build_sgenv pge ge) (Build_sgenv tpge tge)); auto. + + rewrite (sval_eval_preserved ge tge); auto. destruct (sval_eval _ _ _ _); try congruence. intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. @@ -259,27 +259,23 @@ Proof. intros; apply function_ptr_preserved; auto. Qed. -Definition istate_simu (srce: PTree.t node) ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). - Lemma sem_istate_simu dupmap sp st st' rs m is: - sem_istate (Build_sgenv pge ge) sp st rs m is -> + sem_istate ge sp st rs m is -> s_istate_simu dupmap st st' -> exists is', - sem_istate (Build_sgenv tpge tge) sp st' rs m is' /\ + sem_istate tge sp st' rs m is' /\ istate_simu dupmap is is'. -Admitted. +Proof. + intros SEM X; eapply X; eauto. +Qed. Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> s_istate_simu dupmap st st' -> sfval_simu dupmap st.(the_pc) st'.(the_pc) sv sv' -> - sfval_sem (Build_sgenv pge ge) sp st stk f rs0 m0 sv rs m t s -> - exists s', sfval_sem (Build_sgenv tpge tge) sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + sfval_sem pge ge sp st stk f rs0 m0 sv rs m t s -> + exists s', sfval_sem tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct. intros FUN STK SIS SFV. destruct SFV. @@ -290,7 +286,7 @@ Proof. intros (fd' & FIND & TRANSF). erewrite <- function_sig_preserved; eauto. eexists; split; econstructor; eauto. - + erewrite <- (list_sval_eval_preserved (Build_sgenv pge ge) (Build_sgenv tpge tge)); auto. + + erewrite <- (list_sval_eval_preserved ge tge); auto. + simpl. repeat (econstructor; eauto). - (* Sreturn *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. @@ -303,9 +299,9 @@ Qed. Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> - sem_state (Build_sgenv pge ge) sp st stk f rs m t s -> + sem_state pge ge sp st stk f rs m t s -> s_state_simu dupmap st st' -> - exists s', sem_state (Build_sgenv tpge tge) sp st' stk' f' rs m t s' /\ match_states s s'. + exists s', sem_state tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros FUNC STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) -- cgit From b2fedcff4c2e95d8de14c380446d2de9b2a874ec Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 21 Dec 2019 08:38:24 +0100 Subject: Starting a Liveness checker at RTLpath --- mppa_k1c/lib/RTLpath.v | 90 ++++++----- mppa_k1c/lib/RTLpathLivenessChecker.v | 291 ++++++++++++++++++++++++++++++++++ mppa_k1c/lib/RTLpathSE_theory.v | 18 +-- mppa_k1c/lib/RTLpathScheduler.v | 4 +- 4 files changed, 351 insertions(+), 52 deletions(-) create mode 100644 mppa_k1c/lib/RTLpathLivenessChecker.v diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 5de0f5a1..486ec68d 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -75,7 +75,9 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, *) -Definition path_map: Type := PTree.t nat. +Record path_info := { path_size: nat; path_uses: Regset.t }. + +Definition path_map: Type := PTree.t path_info. Definition path_entry (*c:code*) (pm: path_map) (n: node): Prop := pm!n <> None (*/\ c!n <> None*). @@ -94,7 +96,7 @@ Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := (* all paths defined from the path_map are wellformed *) Definition wellformed_path_map (c:code) (pm: path_map): Prop := - forall n path, pm!n = Some path -> wellformed_path c pm path n. + forall n path, pm!n = Some path -> wellformed_path c pm path.(path_size) n. (** We "extend" the notion of RTL program with the additional structure for path. @@ -272,7 +274,7 @@ Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop Inductive step ge pge: state -> trace -> state -> Prop := | exec_path path stack f sp rs m pc t s: (fn_path f)!pc = Some path -> - path_step ge pge path stack f sp rs m pc t s -> + path_step ge pge path.(path_size) stack f sp rs m pc t s -> step ge pge (State stack f sp pc rs m) t s | exec_function_internal s f args m m' stk: Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) -> @@ -521,30 +523,32 @@ Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := Lemma final_node_path f path pc: (fn_path f)!pc = Some path -> - exists i, nth_default_succ_inst (fn_code f) path pc = Some i + exists i, nth_default_succ_inst (fn_code f) path.(path_size) pc = Some i /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF. - exploit (wellformed_suffix_path (fn_code f) (fn_path f) path O); omega || eauto. + set (ps:=path.(path_size)). + exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps O); omega || eauto. destruct 1 as (pc' & NTH_SUCC & WF'); auto. - cutrewrite (path - 0 = path)%nat in NTH_SUCC; try omega. + cutrewrite (ps - 0 = ps)%nat in NTH_SUCC; try omega. unfold nth_default_succ_inst. inversion WF'; clear WF'; subst. simplify_someHyps; eauto. Qed. Lemma internal_node_path path f path0 pc: (fn_path f)!pc = (Some path0) -> - (path < path0)%nat -> + (path < path0.(path_size))%nat -> exists i pc', nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. - intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) path0 (path0-path)); eauto. { omega. } + set (ps:=path0.(path_size)). + intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps (ps-path)); eauto. { omega. } destruct 1 as (pc' & NTH_SUCC & WF'). - cutrewrite (path0 - (path0 - path) = path)%nat in NTH_SUCC; try omega. + cutrewrite (ps - (ps - path) = path)%nat in NTH_SUCC; try omega. unfold nth_default_succ_inst. inversion WF'; clear WF'; subst. { omega. } simplify_someHyps; eauto. @@ -731,8 +735,8 @@ Definition is_inst (s: RTL.state): bool := Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> - (idx <= path)%nat -> - isteps ge (path-idx) f sp rs m pc = Some s2 -> + (idx <= path.(path_size))%nat -> + isteps ge (path.(path_size)-idx) f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> match_inst_states_goal idx s1 (State stack f sp pc rs m). @@ -755,14 +759,14 @@ Proof. Qed. Lemma stuttering path idx stack f sp rs m pc st t s1': - isteps ge (path-(S idx)) f sp rs m pc = Some st -> + isteps ge (path.(path_size)-(S idx)) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> - (S idx <= path)%nat -> + (S idx <= path.(path_size))%nat -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). Proof. - intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path-(S idx))); omega || eauto. + intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path.(path_size)-(S idx))); omega || eauto. intros (i & pc' & Hi & Hpc & DUM). unfold nth_default_succ_inst in Hi. erewrite isteps_normal_exit in Hi; eauto. @@ -770,12 +774,12 @@ Proof. intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. - cutrewrite (path - idx = S (path - (S idx)))%nat; try omega. + set (ps:=path.(path_size)). cutrewrite (ps - idx = S (ps - (S idx)))%nat; try omega. erewrite <- isteps_step_right; eauto. Qed. Lemma normal_exit path stack f sp rs m pc st t s1': - isteps ge path f sp rs m pc = Some st -> + isteps ge path.(path_size) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> @@ -797,9 +801,9 @@ Proof. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. - exists path0; intuition eauto. + exists (path0.(path_size)); intuition eauto. econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. * simpl; eauto. + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. - (* Icall *) @@ -820,18 +824,18 @@ Proof. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. - exists path0; intuition eauto. + exists path0.(path_size); intuition eauto. econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. * simpl; eauto. - (* Ijumptable *) intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. - exists path0; intuition eauto. + exists path0.(path_size); intuition eauto. econstructor; eauto. - * cutrewrite (path0-path0=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. * simpl; eauto. - (* Ireturn *) intros; exploit exec_Ireturn; eauto. @@ -840,16 +844,17 @@ Proof. Qed. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: - isteps ge (path-idx) f sp rs m pc = Some st -> + isteps ge (path.(path_size)-idx) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> - (idx <= path)%nat -> + (idx <= path.(path_size))%nat -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> wf_stackframe stack -> exists idx' s2', - (path_step ge pge path stack f sp rs m pc t s2' + (path_step ge pge path.(path_size) stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) - \/ (path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - /\ (fn_path f)!(the_pc st) = Some O /\ path_step ge pge 0 stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') + \/ (exists path', path_step ge pge path.(path_size) stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + /\ (fn_path f)!(the_pc st) = Some path' /\ path'.(path_size) = O + /\ path_step ge pge path'.(path_size) stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') ) /\ match_states idx' s1' s2'. Proof. @@ -857,7 +862,7 @@ Proof. intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. + (* path_step on normal_exit *) - cutrewrite (path-0=path)%nat in PSTEP; [|omega]. + cutrewrite (path.(path_size)-0=path.(path_size))%nat in PSTEP; [|omega]. exploit normal_exit; eauto. intros (s2' & LSTEP & (idx' & MATCH)). exists idx'; exists s2'; intuition eauto. @@ -866,7 +871,7 @@ Proof. unfold match_states; exists idx; exists (State stack f sp pc rs m); intuition. + (* one or two path_step on early_exit *) - exploit (isteps_resize ge (path - idx)%nat path); eauto; try omega. + exploit (isteps_resize ge (path.(path_size) - idx)%nat path.(path_size)); eauto; try omega. clear PSTEP; intros PSTEP. (* TODO for clarification: move the assert below into a separate lemma *) assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). @@ -884,21 +889,23 @@ Proof. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. } - destruct HPATH0 as ([|path1] & Hpath1). + destruct HPATH0 as (path1 & Hpath1). + destruct (path1.(path_size)) as [|ps] eqn:Hpath1size. * (* two step case *) - exploit (normal_exit 0); simpl; eauto. + exploit (normal_exit path1); try rewrite Hpath1size; simpl; eauto. simpl; intros (s2' & LSTEP & (idx' & MATCH)). exists idx'. exists s2'. constructor; auto. - right. right. intuition eauto. + right. right. eexists; intuition eauto. (* now, prove the last step *) - exploit exec_normal_exit. 4:{ eauto. } + rewrite Hpath1size; exploit exec_normal_exit. 4:{ eauto. } - simpl; eauto. - simpl; eauto. - simpl; eauto. * (* single step case *) - exploit (stuttering (S path1) path1 stack f sp (the_rs st) (the_mem st) (the_pc st)); simpl; auto. - - { cutrewrite (path1-path1=0)%nat; try omega. simpl; eauto. } - - simpl; auto. + exploit (stuttering path1 ps stack f sp (the_rs st) (the_mem st) (the_pc st)); simpl; auto. + - { rewrite Hpath1size; cutrewrite (S ps-S ps=0)%nat; try omega. simpl; eauto. } + - omega. + - simpl; eauto. - simpl; eauto. - intuition subst. repeat eexists; intuition eauto. @@ -919,9 +926,9 @@ Proof. * eapply exec_function_internal; eauto. * unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. - exists path. constructor; auto. + exists path.(path_size). constructor; auto. econstructor; eauto. - - cutrewrite (path-path=O)%nat; simpl; eauto. + - set (ps:=path.(path_size)). cutrewrite (ps-ps=O)%nat; simpl; eauto. omega. - simpl; auto. + (* exec_function_external *) @@ -938,9 +945,9 @@ Proof. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. destruct (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto. - exists path. constructor; auto. + exists path.(path_size). constructor; auto. econstructor; eauto. - - cutrewrite (path-path=O)%nat; simpl; eauto. + - set (ps:=path.(path_size)). cutrewrite (ps-ps=O)%nat; simpl; eauto. omega. - simpl; auto. Qed. @@ -957,7 +964,8 @@ Proof. intros STEP; exploit path_step_complete; eauto. intros (idx' & s2' & H0 & H1). eexists; eexists; eauto. - destruct H0 as [H0|[H0|H0]]; intuition subst; eauto. + destruct H0 as [H0|[H0|(path'&H0)]]; intuition subst; eauto. + - intros; exploit step_noninst_complete; eauto. intros (s2' & STEP & (idx0 & MATCH)). exists idx0; exists s2'; intuition auto. diff --git a/mppa_k1c/lib/RTLpathLivenessChecker.v b/mppa_k1c/lib/RTLpathLivenessChecker.v new file mode 100644 index 00000000..2b20fb82 --- /dev/null +++ b/mppa_k1c/lib/RTLpathLivenessChecker.v @@ -0,0 +1,291 @@ +(** A Liveness checker on RTLpath *) + + +Require Import Coqlib. +Require Import Maps. +Require Import Lattice. +Require Import AST. +Require Import Op. +Require Import Registers. +Require Import Globalenvs Smallstep RTL RTLpath. +Require Import Bool. + +Local Open Scope lazy_bool_scope. + +Local Notation ext alive := (fun r => Regset.In r alive). + +Notation "'ASSERT' A 'IN' B" := (if A then B else None) + (at level 200, A at level 100, B at level 200) + : option_monad_scope. + +Local Open Scope option_monad_scope. + +Local Hint Resolve Regset.mem_2 Regset.subset_2. + +Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. +Proof. + destruct b1; simpl; intuition. +Qed. + +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. + +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 exit_checker {A} (f: function) (alive: Regset.t) (pc: node) (v:A): option A := + SOME path <- f.(fn_path)!pc IN + ASSERT Regset.subset path.(path_uses) alive IN + Some v. + +Definition iinst_checker (f: function) (alive: Regset.t) (i: instruction): option (Regset.t * node) := + match i with + | Inop pc' => Some (alive, pc') + | Iop op args dst pc' => + ASSERT list_mem args alive IN + Some (Regset.add dst alive, pc') + | Iload chunk addr args dst pc' => + ASSERT list_mem args alive IN + Some (Regset.add dst alive, pc') + | Istore chunk addr args src pc' => + ASSERT Regset.mem src alive IN + ASSERT list_mem args alive IN + Some (alive, pc') + | Icond cond args ifso ifnot => + ASSERT list_mem args alive IN + exit_checker f alive ifso (alive, ifnot) + | _ => None (* TODO jumptable ? *) + end. + +Fixpoint ipath_checker (path:nat) (f: function) (alive: Regset.t) (pc:node): option (Regset.t * node) := + match path with + | O => Some (alive, pc) + | S p => + SOME i <- f.(fn_code)!pc IN + SOME res <- iinst_checker f alive i IN + ipath_checker p f (fst res) (snd res) + end. + +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. + +Fixpoint reg_list_add (rl: list reg) (alive: Regset.t) {struct rl} : Regset.t := + match rl with + | nil => alive + | r1 :: rs => reg_list_add rs (Regset.add r1 alive) + end. + +Fixpoint exit_list_checker (f: function) (alive: Regset.t) (l: list node): bool := + match l with + | nil => true + | pc::l' => exit_checker f alive pc tt &&& exit_list_checker f alive l' + end. + +Definition inst_checker (f: function) (alive: Regset.t) (i: instruction): option unit := + match i with + | Icall sig ros args res pc' => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN + exit_checker f (Regset.add res alive) pc' tt + | Itailcall sig ros args => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN + Some tt + | Ibuiltin ef args res pc' => + ASSERT list_mem (params_of_builtin_args args) alive IN + exit_checker f (reg_list_add (params_of_builtin_res res) alive) pc' tt + | Ijumptable arg tbl => + ASSERT Regset.mem arg alive IN + ASSERT exit_list_checker f alive tbl IN + Some tt + | Ireturn optarg => + ASSERT (reg_option_mem optarg) alive IN + Some tt + | _ => + SOME res <- iinst_checker f alive i IN + exit_checker f (fst res) (snd res) tt + end. + +Definition path_checker (f: function) (pc: node) (path:path_info): bool := + (SOME res <- ipath_checker (path.(path_size)) f (path.(path_uses)) pc IN + SOME i <- f.(fn_code)!(snd res) IN + inst_checker f (fst res) i) &&& true. + + + + + +(* AUTRES LEMMES *) + + +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. + +(* +Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). +Proof. + intuition. + - eapply Regset.remove_3; eauto. + - eapply Regset.remove_1; [| eauto]. auto. + - eapply Regset.remove_2; eauto. +Qed. + +Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). +Proof. + induction args; simpl. + + intuition. + + intros; rewrite IHargs, regset_add_spec. intuition. +Qed. +*) + + +Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := + forall r, (alive r) -> rs1#r = rs2#r. + +Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs. +Proof. + unfold eqlive_reg; auto. +Qed. + +Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1. +Proof. + unfold eqlive_reg; intros; symmetry; auto. +Qed. + +Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3. +Proof. + unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. +Qed. + +Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). +Proof. + unfold eqlive_reg; intros NOTALIVE r0 ALIVE. + destruct (Pos.eq_dec r r0) as [H|H]. + - subst; tauto. + - rewrite Regmap.gso; auto. +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. + + + +(* ESSAIS + + +Section PRESERVATION. + +Variable prog: program. + +Let ge := Genv.globalenv prog. + +(** ESSAI pour voir: + +Lemme faux, juste pour se faire la main sur transfer et step... + +REM: peux-être faut-il un transfer plus faible qui rend ce lemme vrai ? +Sinon la preuve de simulation au niveau RTLpath risque d'être infaisable ? + +*) +Lemma step_transfer_wrong i stk f sp pc rs1 m m' t rs1' rs2 (after: Regset.t): + f.(fn_code)!pc = Some i -> + step ge (State stk f sp pc rs1 m) t (State stk f sp pc rs1' m') -> + (eqlive_reg (ext (transfer f pc after)) rs1 rs2) -> + exists rs2', step ge (State stk f sp pc rs2 m) t (State stk f sp pc rs2' m') /\ (eqlive_reg (ext after) rs1' rs2'). +Proof. + intros PC STEP. unfold transfer; rewrite PC. + inv STEP. + + (* Inop *) rewrite H6 in PC. inv PC. intros. eexists; split. + - eapply exec_Inop; eauto. + - auto. + + (* Iop *) rewrite H6 in PC. inv PC. destruct (Regset.mem res after) eqn:RES. + - eexists; split. + * eapply exec_Iop; eauto. + erewrite eqlive_reg_list; eauto. + { eapply eqlive_reg_symmetry; eauto. } + intros r; rewrite reg_list_live_spec. intuition. + * eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros; rewrite reg_list_live_spec, regset_remove_spec. + intuition. + - (* cas faux: ici, on voudrait juste dire que le res est inutile pour la suite du calcul *) +Abort. + + + +End PRESERVATION. + + +Definition is_live (live: PMap.t Regset.t) (f: RTL.function): Prop + := False. + + +Lemma is_live_eqlive_reg pc' f live pc i rs1 rs2: + f.(fn_code)!pc = Some i -> + is_live live f -> + In pc' (successors_instr i) -> + eqlive_reg (ext live#pc) rs1 rs2 -> + eqlive_reg (ext (transfer f pc' live#pc')) rs1 rs2. +Proof. + unfold is_live; intro PC; intros. + eapply eqlive_reg_monotonic; eauto. + eapply Liveness.analyze_solution; eauto. +Qed. + +Ltac exploit_liveness pc' := refine (modusponens _ _ (is_live_eqlive_reg pc' _ _ _ _ _ _ _ _ _ _) _); [ eauto | eauto | try (simpl; intuition) | eauto | eauto]. + +(* Y aura-t-il besoin de chose comme ça ? *) +Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := + | eqlive_stackframe_intro live res f sp pc rs1 rs2 + (LIVE: is_live live f) + (EQUIV: eqlive_reg (ext (live#pc)) rs1 rs2): + eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + +Inductive eqlive_states: state -> state -> Prop := + | eqlive_states_intro + live st1 st2 f sp pc rs1 rs2 m + (STACKS: list_forall2 eqlive_stackframes st1 st2) + (LIVE: is_live live f) + (EQUIV: eqlive_reg (ext (live#pc)) 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 + (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). + +*) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 7f94ae58..f4958082 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -744,7 +744,7 @@ Qed. Definition symb_step (f: function) (pc:node): option s_state := SOME path <- (fn_path f)!pc IN - SOME st <- symb_isteps path f (init_istate pc) IN + SOME st <- symb_isteps path.(path_size) f (init_istate pc) IN SOME i <- (fn_code f)!(st.(the_pc)) IN Some (match symb_istep i st with | Some st' => {| internal := st'; final := Snone |} @@ -752,7 +752,7 @@ Definition symb_step (f: function) (pc:node): option s_state := end). Lemma final_node_path_simpl f path pc: - (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path pc <> None. + (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(path_size) pc <> None. Proof. intros; exploit final_node_path; eauto. intros (i & NTH & DUM). @@ -774,15 +774,15 @@ Qed. *) Theorem symb_step_correct f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> - path_step ge pge path stack f sp rs m pc t s -> + path_step ge pge path.(path_size) stack f sp rs m pc t s -> exists st, symb_step f pc = Some st /\ sem_state pge ge sp st stack f rs m t s. Proof. Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. intros PATH STEP; unfold symb_step; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path.(path_size) f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). + (destruct (nth_default_succ_inst (fn_code f) path.(path_size) 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 sem_option2_istate; destruct (symb_isteps _ _ _) as [st|] eqn: Hst; try congruence); @@ -910,17 +910,17 @@ Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> sem_state pge ge sp st stack f rs m t s1 -> - exists s2, path_step ge pge path stack f sp rs m pc t s2 /\ + exists s2, path_step ge pge path.(path_size) stack f sp rs m pc t s2 /\ equiv_state s1 s2. Proof. Local Hint Resolve init_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path.(path_size) f rs m (init_istate pc) (ist 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 pc) as [i|] eqn: Hi; [clear WF|congruence]). + (destruct (nth_default_succ_inst (fn_code f) path.(path_size) pc) as [i|] eqn: Hi; [clear WF|congruence]). unfold nth_default_succ_inst in Hi. - destruct (symb_isteps path f (init_istate pc)) as [st0|] eqn: Hst0; simpl. + destruct (symb_isteps path.(path_size) f (init_istate pc)) as [st0|] eqn: Hst0; simpl. 2:{ (* absurd case *) exploit symb_isteps_WF; eauto. simpl; intros NDS; rewrite NDS in Hi; congruence. } diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index bfd2384c..102e8e4c 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -325,11 +325,11 @@ Qed. Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: (fn_path f)!pc = Some path -> - path_step ge pge path stk f sp rs m pc t s -> + path_step ge pge path.(path_size) stk f sp rs m pc t s -> list_forall2 match_stackframes stk stk' -> dupmap ! pc' = Some pc -> match_function dupmap f f' -> - exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path' stk' f' sp rs m pc' t s' /\ match_states s s'. + exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(path_size) stk' f' sp rs m pc' t s' /\ match_states s s'. Proof. intros PATH STEP STACKS PC FUNCT. exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } -- cgit From 428763b9c30fb63d96071ed532512d95c9eba828 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 23 Dec 2019 23:17:43 +0100 Subject: RTLpath: liveness + wf checker --- mppa_k1c/lib/RTLpath.v | 82 +++--- mppa_k1c/lib/RTLpathLivegen.v | 479 ++++++++++++++++++++++++++++++++++ mppa_k1c/lib/RTLpathLivenessChecker.v | 291 --------------------- mppa_k1c/lib/RTLpathSE_theory.v | 18 +- mppa_k1c/lib/RTLpathScheduler.v | 4 +- 5 files changed, 539 insertions(+), 335 deletions(-) create mode 100644 mppa_k1c/lib/RTLpathLivegen.v delete mode 100644 mppa_k1c/lib/RTLpathLivenessChecker.v diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 486ec68d..eced2635 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -31,6 +31,10 @@ Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end) (at level 200, X ident, A at level 100, B at level 200) : option_monad_scope. +Notation "'ASSERT' A 'IN' B" := (if A then B else None) + (at level 200, A at level 100, B at level 200) + : option_monad_scope. + Local Open Scope option_monad_scope. (** * Syntax of RTLpath programs *) @@ -75,7 +79,10 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, *) -Record path_info := { path_size: nat; path_uses: Regset.t }. +Record path_info := { + psize: nat; (* number minus 1 of instructions in the path *) + input_regs: Regset.t +}. Definition path_map: Type := PTree.t path_info. @@ -96,7 +103,7 @@ Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop := (* all paths defined from the path_map are wellformed *) Definition wellformed_path_map (c:code) (pm: path_map): Prop := - forall n path, pm!n = Some path -> wellformed_path c pm path.(path_size) n. + forall n path, pm!n = Some path -> wellformed_path c pm path.(psize) n. (** We "extend" the notion of RTL program with the additional structure for path. @@ -274,7 +281,7 @@ Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop Inductive step ge pge: state -> trace -> state -> Prop := | exec_path path stack f sp rs m pc t s: (fn_path f)!pc = Some path -> - path_step ge pge path.(path_size) stack f sp rs m pc t s -> + path_step ge pge path.(psize) stack f sp rs m pc t s -> step ge pge (State stack f sp pc rs m) t s | exec_function_internal s f args m m' stk: Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) -> @@ -317,9 +324,18 @@ Proof. intros; destruct e; simpl; auto. Qed. +Lemma destruct_ASSERT B (P: option B -> Prop) (e: bool) (x: option B): + (e = true -> P x) -> (e = false -> P None) -> (P (ASSERT e IN x)). +Proof. + intros; destruct e; simpl; auto. +Qed. + Ltac inversion_SOME x := try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]). +Ltac inversion_ASSERT := + try (eapply destruct_ASSERT; [ idtac | simpl; try congruence ]). + Ltac simplify_someHyp := match goal with | H: None = Some _ |- _ => inversion H; clear H; subst @@ -523,12 +539,12 @@ Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction := Lemma final_node_path f path pc: (fn_path f)!pc = Some path -> - exists i, nth_default_succ_inst (fn_code f) path.(path_size) pc = Some i + exists i, nth_default_succ_inst (fn_code f) path.(psize) pc = Some i /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. intro WF. - set (ps:=path.(path_size)). + set (ps:=path.(psize)). exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps O); omega || eauto. destruct 1 as (pc' & NTH_SUCC & WF'); auto. cutrewrite (ps - 0 = ps)%nat in NTH_SUCC; try omega. @@ -538,14 +554,14 @@ Qed. Lemma internal_node_path path f path0 pc: (fn_path f)!pc = (Some path0) -> - (path < path0.(path_size))%nat -> + (path < path0.(psize))%nat -> exists i pc', nth_default_succ_inst (fn_code f) path pc = Some i /\ default_succ i = Some pc' /\ (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n). Proof. intros; exploit fn_path_wf; eauto. - set (ps:=path0.(path_size)). + set (ps:=path0.(psize)). intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps (ps-path)); eauto. { omega. } destruct 1 as (pc' & NTH_SUCC & WF'). cutrewrite (ps - (ps - path) = path)%nat in NTH_SUCC; try omega. @@ -735,8 +751,8 @@ Definition is_inst (s: RTL.state): bool := Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := | State_match path stack f sp pc rs m s2: (fn_path f)!pc = Some path -> - (idx <= path.(path_size))%nat -> - isteps ge (path.(path_size)-idx) f sp rs m pc = Some s2 -> + (idx <= path.(psize))%nat -> + isteps ge (path.(psize)-idx) f sp rs m pc = Some s2 -> s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> match_inst_states_goal idx s1 (State stack f sp pc rs m). @@ -759,14 +775,14 @@ Proof. Qed. Lemma stuttering path idx stack f sp rs m pc st t s1': - isteps ge (path.(path_size)-(S idx)) f sp rs m pc = Some st -> + isteps ge (path.(psize)-(S idx)) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> - (S idx <= path.(path_size))%nat -> + (S idx <= path.(psize))%nat -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). Proof. - intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path.(path_size)-(S idx))); omega || eauto. + intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path.(psize)-(S idx))); omega || eauto. intros (i & pc' & Hi & Hpc & DUM). unfold nth_default_succ_inst in Hi. erewrite isteps_normal_exit in Hi; eauto. @@ -774,12 +790,12 @@ Proof. intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. - set (ps:=path.(path_size)). cutrewrite (ps - idx = S (ps - (S idx)))%nat; try omega. + set (ps:=path.(psize)). cutrewrite (ps - idx = S (ps - (S idx)))%nat; try omega. erewrite <- isteps_step_right; eauto. Qed. Lemma normal_exit path stack f sp rs m pc st t s1': - isteps ge path.(path_size) f sp rs m pc = Some st -> + isteps ge path.(psize) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> st.(continue) = true -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> @@ -801,9 +817,9 @@ Proof. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. - exists (path0.(path_size)); intuition eauto. + exists (path0.(psize)); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. * simpl; eauto. + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. - (* Icall *) @@ -824,18 +840,18 @@ Proof. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. - exists path0.(path_size); intuition eauto. + exists path0.(psize); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. * simpl; eauto. - (* Ijumptable *) intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. - exists path0.(path_size); intuition eauto. + exists path0.(psize); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(path_size)-path0.(path_size)=0)%nat; simpl; eauto || omega. + * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. * simpl; eauto. - (* Ireturn *) intros; exploit exec_Ireturn; eauto. @@ -844,17 +860,17 @@ Proof. Qed. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: - isteps ge (path.(path_size)-idx) f sp rs m pc = Some st -> + isteps ge (path.(psize)-idx) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> - (idx <= path.(path_size))%nat -> + (idx <= path.(psize))%nat -> RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> wf_stackframe stack -> exists idx' s2', - (path_step ge pge path.(path_size) stack f sp rs m pc t s2' + (path_step ge pge path.(psize) stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) - \/ (exists path', path_step ge pge path.(path_size) stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - /\ (fn_path f)!(the_pc st) = Some path' /\ path'.(path_size) = O - /\ path_step ge pge path'.(path_size) stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') + \/ (exists path', path_step ge pge path.(psize) stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + /\ (fn_path f)!(the_pc st) = Some path' /\ path'.(psize) = O + /\ path_step ge pge path'.(psize) stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') ) /\ match_states idx' s1' s2'. Proof. @@ -862,7 +878,7 @@ Proof. intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. + (* path_step on normal_exit *) - cutrewrite (path.(path_size)-0=path.(path_size))%nat in PSTEP; [|omega]. + cutrewrite (path.(psize)-0=path.(psize))%nat in PSTEP; [|omega]. exploit normal_exit; eauto. intros (s2' & LSTEP & (idx' & MATCH)). exists idx'; exists s2'; intuition eauto. @@ -871,7 +887,7 @@ Proof. unfold match_states; exists idx; exists (State stack f sp pc rs m); intuition. + (* one or two path_step on early_exit *) - exploit (isteps_resize ge (path.(path_size) - idx)%nat path.(path_size)); eauto; try omega. + exploit (isteps_resize ge (path.(psize) - idx)%nat path.(psize)); eauto; try omega. clear PSTEP; intros PSTEP. (* TODO for clarification: move the assert below into a separate lemma *) assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). @@ -890,7 +906,7 @@ Proof. destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. } destruct HPATH0 as (path1 & Hpath1). - destruct (path1.(path_size)) as [|ps] eqn:Hpath1size. + destruct (path1.(psize)) as [|ps] eqn:Hpath1size. * (* two step case *) exploit (normal_exit path1); try rewrite Hpath1size; simpl; eauto. simpl; intros (s2' & LSTEP & (idx' & MATCH)). @@ -926,9 +942,9 @@ Proof. * eapply exec_function_internal; eauto. * unfold match_states, match_inst_states; simpl. destruct (initialize_path (*fn_code f*) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. - exists path.(path_size). constructor; auto. + exists path.(psize). constructor; auto. econstructor; eauto. - - set (ps:=path.(path_size)). cutrewrite (ps-ps=O)%nat; simpl; eauto. + - set (ps:=path.(psize)). cutrewrite (ps-ps=O)%nat; simpl; eauto. omega. - simpl; auto. + (* exec_function_external *) @@ -945,9 +961,9 @@ Proof. rewrite wf_stackframe_cons in WFSTACK. destruct WFSTACK as (H0 & H1); simpl in H0. destruct (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto. - exists path.(path_size). constructor; auto. + exists path.(psize). constructor; auto. econstructor; eauto. - - set (ps:=path.(path_size)). cutrewrite (ps-ps=O)%nat; simpl; eauto. + - set (ps:=path.(psize)). cutrewrite (ps-ps=O)%nat; simpl; eauto. omega. - simpl; auto. Qed. diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v new file mode 100644 index 00000000..ab8a250e --- /dev/null +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -0,0 +1,479 @@ +(** Building a RTLpath program with liveness annotation. +*) + + +Require Import Coqlib. +Require Import Maps. +Require Import Lattice. +Require Import AST. +Require Import Op. +Require Import Registers. +Require Import Globalenvs Smallstep RTL RTLpath. +Require Import Bool Errors. +Require Import Program. + +Local Open Scope lazy_bool_scope. + +Local Notation ext alive := (fun r => Regset.In r alive). + +Local Open Scope option_monad_scope. + +Axiom build_path_map: RTL.function -> path_map. + +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. + +Definition exit_checker {A} (pm: path_map) (alive: Regset.t) (pc: node) (v:A): option A := + SOME path <- pm!pc IN + ASSERT Regset.subset path.(input_regs) alive IN + Some v. + +Lemma exit_checker_path_entry A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res: + exit_checker pm alive pc v = Some res -> path_entry pm pc. +Proof. + unfold exit_checker, path_entry. + inversion_SOME path; simpl; congruence. +Qed. + +Lemma exit_checker_res A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res: + exit_checker pm alive pc v = Some res -> v=res. +Proof. + unfold exit_checker, path_entry. + inversion_SOME path; try_simplify_someHyps. + inversion_ASSERT; try_simplify_someHyps. +Qed. + +Definition iinst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option (Regset.t * node) := + match i with + | Inop pc' => Some (alive, pc') + | Iop op args dst pc' => + ASSERT list_mem args alive IN + Some (Regset.add dst alive, pc') + | Iload chunk addr args dst pc' => + ASSERT list_mem args alive IN + Some (Regset.add dst alive, pc') + | Istore chunk addr args src pc' => + ASSERT Regset.mem src alive IN + ASSERT list_mem args alive IN + Some (alive, pc') + | Icond cond args ifso ifnot => + ASSERT list_mem args alive IN + exit_checker pm alive ifso (alive, ifnot) + | _ => None (* TODO jumptable ? *) + end. + + +Local Hint Resolve exit_checker_path_entry. + +Lemma iinst_checker_path_entry (pm: path_map) (alive: Regset.t) (i: instruction) res pc: + iinst_checker pm alive i = Some res -> + early_exit i = Some pc -> path_entry pm pc. +Proof. + destruct i; simpl; try_simplify_someHyps; subst. + inversion_ASSERT; try_simplify_someHyps. +Qed. + +Lemma iinst_checker_default_succ (pm: path_map) (alive: Regset.t) (i: instruction) res pc: + iinst_checker pm alive i = Some res -> + pc = snd res -> + default_succ i = Some pc. +Proof. + destruct i; simpl; try_simplify_someHyps; subst; + repeat (inversion_ASSERT); try_simplify_someHyps. + intros; exploit exit_checker_res; eauto. + intros; subst. simpl; auto. +Qed. + +Fixpoint ipath_checker (ps:nat) (f: RTL.function) (pm: path_map) (alive: Regset.t) (pc:node): option (Regset.t * node) := + match ps with + | O => Some (alive, pc) + | S p => + SOME i <- f.(fn_code)!pc IN + SOME res <- iinst_checker pm alive i IN + ipath_checker p f pm (fst res) (snd res) + end. + +Lemma ipath_checker_wellformed f pm ps: forall alive pc res, + ipath_checker ps f pm alive pc = Some res -> + wellformed_path f.(fn_code) pm 0 (snd res) -> + wellformed_path f.(fn_code) pm ps pc. +Proof. + induction ps; simpl; try_simplify_someHyps. + inversion_SOME i; inversion_SOME res'. + intros. eapply wf_internal_node; eauto. + * eapply iinst_checker_default_succ; eauto. + * intros; eapply iinst_checker_path_entry; eauto. +Qed. + +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. + +Fixpoint reg_list_add (rl: list reg) (alive: Regset.t) {struct rl} : Regset.t := + match rl with + | nil => alive + | r1 :: rs => reg_list_add rs (Regset.add r1 alive) + end. + +Fixpoint exit_list_checker (pm: path_map) (alive: Regset.t) (l: list node): bool := + match l with + | nil => true + | pc::l' => exit_checker pm alive pc tt &&& exit_list_checker pm alive 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 exit_list_checker_path_entry pm alive l pc: + exit_list_checker pm alive l = true -> List.In pc l -> path_entry pm pc. +Proof. + intros EXIT PC; induction l; intuition. + simpl in * |-. rewrite lazy_and_Some_true in EXIT. + firstorder (subst; eauto). +Qed. + +Local Hint Resolve exit_list_checker_path_entry. + +Definition inst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option unit := + match i with + | Icall sig ros args res pc' => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN + exit_checker pm (Regset.add res alive) pc' tt + | Itailcall sig ros args => + ASSERT list_mem args alive IN + ASSERT reg_sum_mem ros alive IN + Some tt + | Ibuiltin ef args res pc' => + ASSERT list_mem (params_of_builtin_args args) alive IN + exit_checker pm (reg_list_add (params_of_builtin_res res) alive) pc' tt + | Ijumptable arg tbl => + ASSERT Regset.mem arg alive IN + ASSERT exit_list_checker pm alive tbl IN + Some tt + | Ireturn optarg => + ASSERT (reg_option_mem optarg) alive IN + Some tt + | _ => + SOME res <- iinst_checker pm alive i IN + exit_checker pm (fst res) (snd res) tt + end. + +Lemma inst_checker_wellformed (c:code) pc (pm: path_map) (alive: Regset.t) (i: instruction): + inst_checker pm alive i = Some tt -> + c!pc = Some i -> wellformed_path c pm 0 pc. +Proof. + intros CHECK PC. eapply wf_last_node; eauto. + clear c pc PC. intros pc PC. + destruct i; simpl in * |- *; intuition (subst; eauto); + try (generalize CHECK; clear CHECK; try (inversion_SOME path); repeat inversion_ASSERT; try_simplify_someHyps). + intros X; exploit exit_checker_res; eauto. + clear X. intros; subst; eauto. +Qed. + +Definition path_checker (f: RTL.function) pm (pc: node) (path:path_info): option unit := + SOME res <- ipath_checker (path.(psize)) f pm (path.(input_regs)) pc IN + SOME i <- f.(fn_code)!(snd res) IN + inst_checker pm (fst res) i. + +Lemma path_checker_wellformed f pm pc path: + path_checker f pm pc path = Some tt -> wellformed_path (f.(fn_code)) pm (path.(psize)) pc. +Proof. + unfold path_checker. + inversion_SOME res. + inversion_SOME i. + intros; eapply ipath_checker_wellformed; eauto. + eapply inst_checker_wellformed; eauto. +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. + +Fixpoint list_path_checker f pm (l:list (node*path_info)): bool := + match l with + | nil => true + | (pc, path)::l' => + path_checker f pm pc path &&& list_path_checker f pm l' + end. + +Lemma list_path_checker_correct f pm l: + list_path_checker f pm l = true -> forall e, List.In e l -> path_checker f pm (fst e) (snd e) = Some tt. +Proof. + intros CHECKER e H; induction l as [|(pc & path) l]; intuition. + simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto). +Qed. + +Definition function_checker (f: RTL.function) pm: bool := + pm!(f.(fn_entrypoint)) &&& list_path_checker f pm (PTree.elements pm). + +Lemma function_checker_correct f pm pc path: + function_checker f pm = true -> + pm!pc = Some path -> + path_checker f pm pc path = Some tt. +Proof. + unfold function_checker; rewrite lazy_and_Some_true. + intros (ENTRY & PATH) PC. + exploit list_path_checker_correct; eauto. + - eapply PTree.elements_correct; eauto. + - simpl; auto. +Qed. + +Lemma function_checker_wellformed_path_map f pm: + function_checker f pm = true -> wellformed_path_map f.(fn_code) pm. +Proof. + unfold wellformed_path_map. + intros; eapply path_checker_wellformed; eauto. + intros; eapply function_checker_correct; eauto. +Qed. + +Lemma function_checker_path_entry f pm: + function_checker f pm = true -> path_entry pm (f.(fn_entrypoint)). +Proof. + unfold function_checker; rewrite lazy_and_Some_true; + unfold path_entry. firstorder congruence. +Qed. + +Definition liveness_ok_function (f: function): Prop := + forall pc path, f.(fn_path)!pc = Some path -> path_checker f f.(fn_path) pc path = Some tt. + +Program Definition transf_function (f: RTL.function): { r: res function | forall f', r = OK f' -> liveness_ok_function f' /\ f'.(fn_RTL) = f } := + let pm := build_path_map f in + match function_checker f pm with + | true => OK {| fn_RTL := f; fn_path := pm |} + | false => Error(msg "RTLpathGen: function_checker failed") + end. +Obligation 1. + apply function_checker_path_entry; auto. +Qed. +Obligation 2. + apply function_checker_wellformed_path_map; auto. +Qed. +Obligation 3. + unfold liveness_ok_function; simpl; intros; intuition. + apply function_checker_correct; auto. +Qed. + +Definition transf_fundef (f: RTL.fundef) : res fundef := + transf_partial_fundef (fun f => ` (transf_function f)) f. + +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). + +Lemma transf_fundef_correct f f': + transf_fundef f = OK f' -> (liveness_ok_fundef f') /\ fundef_RTL f' = f. +Proof. + intros TRANSF; destruct f; simpl; monadInv TRANSF. + - destruct (transf_function f) as [res H]; simpl in * |- *; auto. + destruct (H _ EQ). + intuition subst; auto. apply liveness_ok_Internal; auto. + - intuition. apply liveness_ok_External; auto. +Qed. + +Definition transf_program (p: RTL.program) : res program := + transform_partial_program transf_fundef p. + +(* TODO: suite à déplacer dans un autre fichier ? *) + +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. + +(* +Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). +Proof. + intuition. + - eapply Regset.remove_3; eauto. + - eapply Regset.remove_1; [| eauto]. auto. + - eapply Regset.remove_2; eauto. +Qed. + +Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). +Proof. + induction args; simpl. + + intuition. + + intros; rewrite IHargs, regset_add_spec. intuition. +Qed. +*) + + +Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := + forall r, (alive r) -> rs1#r = rs2#r. + +Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs. +Proof. + unfold eqlive_reg; auto. +Qed. + +Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1. +Proof. + unfold eqlive_reg; intros; symmetry; auto. +Qed. + +Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3. +Proof. + unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. +Qed. + +Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). +Proof. + unfold eqlive_reg; intros NOTALIVE r0 ALIVE. + destruct (Pos.eq_dec r r0) as [H|H]. + - subst; tauto. + - rewrite Regmap.gso; auto. +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. + +Local Hint Resolve Regset.mem_2 Regset.subset_2. + +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. + + +(* TODO: suite à revoir... *) + +Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := + | eqlive_stackframe_intro path res f sp pc rs1 rs2 + (PATH: f.(fn_path)!pc = Some path) + (LIVE: path_checker f pc path <> None) + (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): + eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + +Inductive eqlive_states: state -> state -> Prop := + | eqlive_states_intro + path st1 st2 f sp pc rs1 rs2 m + (STACKS: list_forall2 eqlive_stackframes st1 st2) + (PATH: f.(fn_path)!pc = Some path) + (LIVE: path_checker f pc path <> None) + (EQUIV: eqlive_reg (ext path.(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 + (* TODO: add info on 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 + (* TODO: add info on f *) + (STACKS: list_forall2 eqlive_stackframes st1 st2): + eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). + +(* ESSAIS *) + +Section PRESERVATION. + +Variable prog: program. + +Let pge := Genv.globalenv prog. +Let ge := Genv.globalenv (program_RTL prog). + +Theorem step_eqlive t s1 s1' s2: + step ge pge s1 t s1' -> + (eqlive_states s1 s2) -> + exists s2', step ge pge s2 t s2' /\ (eqlive_states s1' s2'). +Proof. + destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]. + - intros EQLIVE; inv EQLIVE; simplify_someHyps. + intro PATH. + admit. + - intros EQLIVE; inv EQLIVE. + eexists; split. + * eapply exec_function_internal; eauto. + * eapply eqlive_states_intro; eauto. +Qed. + +(* +path : path_info +stack : list stackframe +f : function +sp : Values.val +rs : regset +m : Memory.Mem.mem +pc : positive +t : Events.trace +s : state +STEP : path_step ge pge (psize path) stack f sp rs m pc t s +st2 : list stackframe +rs2 : regset +STACKS : list_forall2 eqlive_stackframes stack st2 +EQUIV : eqlive_reg (ext (input_regs path)) rs rs2 +LIVE : path_checker f pc path <> None +PATH : (fn_path f) ! pc = Some path +______________________________________(1/1) +exists s2' : state, step ge pge (State st2 f sp pc rs2 m) t s2' /\ eqlive_states s s2' +*) + +End PRESERVATION. + +(* + +Definition is_live (live: PMap.t Regset.t) (f: RTL.function): Prop + := False. + + +Lemma is_live_eqlive_reg pc' f live pc i rs1 rs2: + f.(fn_code)!pc = Some i -> + is_live live f -> + In pc' (successors_instr i) -> + eqlive_reg (ext live#pc) rs1 rs2 -> + eqlive_reg (ext (transfer f pc' live#pc')) rs1 rs2. +Proof. + unfold is_live; intro PC; intros. + eapply eqlive_reg_monotonic; eauto. + eapply Liveness.analyze_solution; eauto. +Qed. + +Ltac exploit_liveness pc' := refine (modusponens _ _ (is_live_eqlive_reg pc' _ _ _ _ _ _ _ _ _ _) _); [ eauto | eauto | try (simpl; intuition) | eauto | eauto]. + +(* Y aura-t-il besoin de chose comme ça ? *) + +*) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathLivenessChecker.v b/mppa_k1c/lib/RTLpathLivenessChecker.v deleted file mode 100644 index 2b20fb82..00000000 --- a/mppa_k1c/lib/RTLpathLivenessChecker.v +++ /dev/null @@ -1,291 +0,0 @@ -(** A Liveness checker on RTLpath *) - - -Require Import Coqlib. -Require Import Maps. -Require Import Lattice. -Require Import AST. -Require Import Op. -Require Import Registers. -Require Import Globalenvs Smallstep RTL RTLpath. -Require Import Bool. - -Local Open Scope lazy_bool_scope. - -Local Notation ext alive := (fun r => Regset.In r alive). - -Notation "'ASSERT' A 'IN' B" := (if A then B else None) - (at level 200, A at level 100, B at level 200) - : option_monad_scope. - -Local Open Scope option_monad_scope. - -Local Hint Resolve Regset.mem_2 Regset.subset_2. - -Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. -Proof. - destruct b1; simpl; intuition. -Qed. - -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. - -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 exit_checker {A} (f: function) (alive: Regset.t) (pc: node) (v:A): option A := - SOME path <- f.(fn_path)!pc IN - ASSERT Regset.subset path.(path_uses) alive IN - Some v. - -Definition iinst_checker (f: function) (alive: Regset.t) (i: instruction): option (Regset.t * node) := - match i with - | Inop pc' => Some (alive, pc') - | Iop op args dst pc' => - ASSERT list_mem args alive IN - Some (Regset.add dst alive, pc') - | Iload chunk addr args dst pc' => - ASSERT list_mem args alive IN - Some (Regset.add dst alive, pc') - | Istore chunk addr args src pc' => - ASSERT Regset.mem src alive IN - ASSERT list_mem args alive IN - Some (alive, pc') - | Icond cond args ifso ifnot => - ASSERT list_mem args alive IN - exit_checker f alive ifso (alive, ifnot) - | _ => None (* TODO jumptable ? *) - end. - -Fixpoint ipath_checker (path:nat) (f: function) (alive: Regset.t) (pc:node): option (Regset.t * node) := - match path with - | O => Some (alive, pc) - | S p => - SOME i <- f.(fn_code)!pc IN - SOME res <- iinst_checker f alive i IN - ipath_checker p f (fst res) (snd res) - end. - -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. - -Fixpoint reg_list_add (rl: list reg) (alive: Regset.t) {struct rl} : Regset.t := - match rl with - | nil => alive - | r1 :: rs => reg_list_add rs (Regset.add r1 alive) - end. - -Fixpoint exit_list_checker (f: function) (alive: Regset.t) (l: list node): bool := - match l with - | nil => true - | pc::l' => exit_checker f alive pc tt &&& exit_list_checker f alive l' - end. - -Definition inst_checker (f: function) (alive: Regset.t) (i: instruction): option unit := - match i with - | Icall sig ros args res pc' => - ASSERT list_mem args alive IN - ASSERT reg_sum_mem ros alive IN - exit_checker f (Regset.add res alive) pc' tt - | Itailcall sig ros args => - ASSERT list_mem args alive IN - ASSERT reg_sum_mem ros alive IN - Some tt - | Ibuiltin ef args res pc' => - ASSERT list_mem (params_of_builtin_args args) alive IN - exit_checker f (reg_list_add (params_of_builtin_res res) alive) pc' tt - | Ijumptable arg tbl => - ASSERT Regset.mem arg alive IN - ASSERT exit_list_checker f alive tbl IN - Some tt - | Ireturn optarg => - ASSERT (reg_option_mem optarg) alive IN - Some tt - | _ => - SOME res <- iinst_checker f alive i IN - exit_checker f (fst res) (snd res) tt - end. - -Definition path_checker (f: function) (pc: node) (path:path_info): bool := - (SOME res <- ipath_checker (path.(path_size)) f (path.(path_uses)) pc IN - SOME i <- f.(fn_code)!(snd res) IN - inst_checker f (fst res) i) &&& true. - - - - - -(* AUTRES LEMMES *) - - -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. - -(* -Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). -Proof. - intuition. - - eapply Regset.remove_3; eauto. - - eapply Regset.remove_1; [| eauto]. auto. - - eapply Regset.remove_2; eauto. -Qed. - -Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). -Proof. - induction args; simpl. - + intuition. - + intros; rewrite IHargs, regset_add_spec. intuition. -Qed. -*) - - -Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := - forall r, (alive r) -> rs1#r = rs2#r. - -Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs. -Proof. - unfold eqlive_reg; auto. -Qed. - -Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1. -Proof. - unfold eqlive_reg; intros; symmetry; auto. -Qed. - -Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3. -Proof. - unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. -Qed. - -Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). -Proof. - unfold eqlive_reg; intros NOTALIVE r0 ALIVE. - destruct (Pos.eq_dec r r0) as [H|H]. - - subst; tauto. - - rewrite Regmap.gso; auto. -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. - - - -(* ESSAIS - - -Section PRESERVATION. - -Variable prog: program. - -Let ge := Genv.globalenv prog. - -(** ESSAI pour voir: - -Lemme faux, juste pour se faire la main sur transfer et step... - -REM: peux-être faut-il un transfer plus faible qui rend ce lemme vrai ? -Sinon la preuve de simulation au niveau RTLpath risque d'être infaisable ? - -*) -Lemma step_transfer_wrong i stk f sp pc rs1 m m' t rs1' rs2 (after: Regset.t): - f.(fn_code)!pc = Some i -> - step ge (State stk f sp pc rs1 m) t (State stk f sp pc rs1' m') -> - (eqlive_reg (ext (transfer f pc after)) rs1 rs2) -> - exists rs2', step ge (State stk f sp pc rs2 m) t (State stk f sp pc rs2' m') /\ (eqlive_reg (ext after) rs1' rs2'). -Proof. - intros PC STEP. unfold transfer; rewrite PC. - inv STEP. - + (* Inop *) rewrite H6 in PC. inv PC. intros. eexists; split. - - eapply exec_Inop; eauto. - - auto. - + (* Iop *) rewrite H6 in PC. inv PC. destruct (Regset.mem res after) eqn:RES. - - eexists; split. - * eapply exec_Iop; eauto. - erewrite eqlive_reg_list; eauto. - { eapply eqlive_reg_symmetry; eauto. } - intros r; rewrite reg_list_live_spec. intuition. - * eapply eqlive_reg_update. - eapply eqlive_reg_monotonic; eauto. - intros; rewrite reg_list_live_spec, regset_remove_spec. - intuition. - - (* cas faux: ici, on voudrait juste dire que le res est inutile pour la suite du calcul *) -Abort. - - - -End PRESERVATION. - - -Definition is_live (live: PMap.t Regset.t) (f: RTL.function): Prop - := False. - - -Lemma is_live_eqlive_reg pc' f live pc i rs1 rs2: - f.(fn_code)!pc = Some i -> - is_live live f -> - In pc' (successors_instr i) -> - eqlive_reg (ext live#pc) rs1 rs2 -> - eqlive_reg (ext (transfer f pc' live#pc')) rs1 rs2. -Proof. - unfold is_live; intro PC; intros. - eapply eqlive_reg_monotonic; eauto. - eapply Liveness.analyze_solution; eauto. -Qed. - -Ltac exploit_liveness pc' := refine (modusponens _ _ (is_live_eqlive_reg pc' _ _ _ _ _ _ _ _ _ _) _); [ eauto | eauto | try (simpl; intuition) | eauto | eauto]. - -(* Y aura-t-il besoin de chose comme ça ? *) -Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := - | eqlive_stackframe_intro live res f sp pc rs1 rs2 - (LIVE: is_live live f) - (EQUIV: eqlive_reg (ext (live#pc)) rs1 rs2): - eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). - -Inductive eqlive_states: state -> state -> Prop := - | eqlive_states_intro - live st1 st2 f sp pc rs1 rs2 m - (STACKS: list_forall2 eqlive_stackframes st1 st2) - (LIVE: is_live live f) - (EQUIV: eqlive_reg (ext (live#pc)) 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 - (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). - -*) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index f4958082..a723f7b2 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -744,7 +744,7 @@ Qed. Definition symb_step (f: function) (pc:node): option s_state := SOME path <- (fn_path f)!pc IN - SOME st <- symb_isteps path.(path_size) f (init_istate pc) IN + SOME st <- symb_isteps path.(psize) f (init_istate pc) IN SOME i <- (fn_code f)!(st.(the_pc)) IN Some (match symb_istep i st with | Some st' => {| internal := st'; final := Snone |} @@ -752,7 +752,7 @@ Definition symb_step (f: function) (pc:node): option s_state := end). Lemma final_node_path_simpl f path pc: - (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(path_size) pc <> None. + (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). @@ -774,15 +774,15 @@ Qed. *) Theorem symb_step_correct f pc pge ge sp path stack rs m t s: (fn_path f)!pc = Some path -> - path_step ge pge path.(path_size) stack f sp rs m pc t s -> + path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, symb_step f pc = Some st /\ sem_state pge ge sp st stack f rs m t s. Proof. Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. intros PATH STEP; unfold symb_step; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path.(path_size) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist 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.(path_size) pc) as [i|] eqn: Hi; [clear WF|congruence]). + (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 sem_option2_istate; destruct (symb_isteps _ _ _) as [st|] eqn: Hst; try congruence); @@ -910,17 +910,17 @@ Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> symb_step f pc = Some st -> sem_state pge ge sp st stack f rs m t s1 -> - exists s2, path_step ge pge path.(path_size) stack f sp rs m pc t s2 /\ + 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_sem_istate. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path.(path_size) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist 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.(path_size) pc) as [i|] eqn: Hi; [clear WF|congruence]). + (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 (symb_isteps path.(path_size) f (init_istate pc)) as [st0|] eqn: Hst0; simpl. + destruct (symb_isteps path.(psize) f (init_istate pc)) as [st0|] eqn: Hst0; simpl. 2:{ (* absurd case *) exploit symb_isteps_WF; eauto. simpl; intros NDS; rewrite NDS in Hi; congruence. } diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 102e8e4c..9c1ba48d 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -325,11 +325,11 @@ Qed. Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: (fn_path f)!pc = Some path -> - path_step ge pge path.(path_size) stk f sp rs m pc t s -> + path_step ge pge path.(psize) stk f sp rs m pc t s -> list_forall2 match_stackframes stk stk' -> dupmap ! pc' = Some pc -> match_function dupmap f f' -> - exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(path_size) stk' f' sp rs m pc' t s' /\ match_states s s'. + exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(psize) stk' f' sp rs m pc' t s' /\ match_states s s'. Proof. intros PATH STEP STACKS PC FUNCT. exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } -- cgit From 7bfaa1011cd3c6e74c8bd2a41fdbc67642eb1ac5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 24 Dec 2019 09:02:36 +0100 Subject: RTLpathLivegen: proof of preservation --- configure | 2 +- mppa_k1c/lib/RTLpath.v | 1 - mppa_k1c/lib/RTLpathLivegen.v | 274 +++++++++++++++------------------------ mppa_k1c/lib/RTLpathLiveproofs.v | 186 ++++++++++++++++++++++++++ 4 files changed, 294 insertions(+), 169 deletions(-) create mode 100644 mppa_k1c/lib/RTLpathLiveproofs.v diff --git a/configure b/configure index da20d0dc..ddcc6619 100755 --- a/configure +++ b/configure @@ -849,7 +849,7 @@ ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure EXECUTE=k1-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __K1C_COS__ SIMU=k1-cluster -- -BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathSE_theory.v\\ +BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathLivegen.v RTLpathSE_theory.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index eced2635..37ee5e70 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -981,7 +981,6 @@ Proof. intros (idx' & s2' & H0 & H1). eexists; eexists; eauto. destruct H0 as [H0|[H0|(path'&H0)]]; intuition subst; eauto. - - intros; exploit step_noninst_complete; eauto. intros (s2' & STEP & (idx0 & MATCH)). exists idx0; exists s2'; intuition auto. diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index ab8a250e..d4ae9290 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -9,13 +9,11 @@ Require Import AST. Require Import Op. Require Import Registers. Require Import Globalenvs Smallstep RTL RTLpath. -Require Import Bool Errors. +Require Import Bool Errors Linking Values Events. Require Import Program. Local Open Scope lazy_bool_scope. -Local Notation ext alive := (fun r => Regset.In r alive). - Local Open Scope option_monad_scope. Axiom build_path_map: RTL.function -> path_map. @@ -285,195 +283,137 @@ Qed. Definition transf_program (p: RTL.program) : res program := transform_partial_program transf_fundef p. -(* TODO: suite à déplacer dans un autre fichier ? *) - -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. - -(* -Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). -Proof. - intuition. - - eapply Regset.remove_3; eauto. - - eapply Regset.remove_1; [| eauto]. auto. - - eapply Regset.remove_2; eauto. -Qed. +Definition match_prog (p: RTL.program) (tp: program) := + match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp. -Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). +Lemma transf_program_match: + forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog. Proof. - induction args; simpl. - + intuition. - + intros; rewrite IHargs, regset_add_spec. intuition. + intros. eapply match_transform_partial_program_contextual; eauto. Qed. -*) - - -Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := - forall r, (alive r) -> rs1#r = rs2#r. -Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs. -Proof. - unfold eqlive_reg; auto. -Qed. +Section PRESERVATION. -Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1. -Proof. - unfold eqlive_reg; intros; symmetry; auto. -Qed. +Variables prog: RTL.program. +Variables tprog: program. +Hypothesis TRANSL: match_prog prog tprog. +Let ge := Genv.globalenv prog. +Let tpge := Genv.globalenv tprog. +Let tge := Genv.globalenv (program_RTL tprog). -Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3. +Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. - unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. + rewrite <- (Genv.find_symbol_match TRANSL). + apply (Genv.find_symbol_match (match_prog_RTL tprog)). Qed. -Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). +Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z. Proof. - unfold eqlive_reg; intros NOTALIVE r0 ALIVE. - destruct (Pos.eq_dec r r0) as [H|H]. - - subst; tauto. - - rewrite Regmap.gso; auto. + unfold Senv.equiv. intuition congruence. 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). +Lemma senv_preserved: Senv.equiv ge tge. 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. + eapply senv_transitivity. { eapply (Genv.senv_match TRANSL). } + eapply RTLpath.senv_preserved. 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. +Lemma function_ptr_preserved v f: Genv.find_funct_ptr ge v = Some f -> + exists tf, Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf. Proof. - unfold eqlive_reg; intuition. + intros; apply (Genv.find_funct_ptr_transf_partial TRANSL); eauto. Qed. -Local Hint Resolve Regset.mem_2 Regset.subset_2. -Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. +Lemma function_ptr_RTL_preserved v f: Genv.find_funct_ptr ge v = Some f -> Genv.find_funct_ptr tge v = Some f. Proof. - destruct b1; simpl; intuition. + intros; exploit function_ptr_preserved; eauto. + intros (tf & Htf & TRANS). + exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto. + intros (cunit & tf0 & X & Y & DUM); subst. + unfold tge. rewrite X. + exploit transf_fundef_correct; eauto. + intuition subst; auto. 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. +Lemma find_function_preserved ros rs fd: + RTL.find_function ge ros rs = Some fd -> RTL.find_function tge ros rs = Some fd. 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. + intros H; assert (X: exists tfd, find_function tpge ros rs = Some tfd /\ fd = fundef_RTL tfd). + * destruct ros; simpl in * |- *. + + intros; exploit (Genv.find_funct_match TRANSL); eauto. + intros (cuint & tf & H1 & H2 & H3); subst; repeat econstructor; eauto. + exploit transf_fundef_correct; eauto. + intuition auto. + + rewrite <- (Genv.find_symbol_match TRANSL) in H. + unfold tpge. destruct (Genv.find_symbol _ i); simpl; try congruence. + exploit function_ptr_preserved; eauto. + intros (tf & H1 & H2); subst; repeat econstructor; eauto. + exploit transf_fundef_correct; eauto. + intuition auto. + * destruct X as (tf & X1 & X2); subst. + eapply find_function_RTL_match; eauto. +Qed. + + +Local Hint Resolve symbols_preserved senv_preserved. + +Theorem transf_program_RTL_correct: + forward_simulation (RTL.semantics prog) (RTL.semantics (program_RTL tprog)). 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. + eapply forward_simulation_step with (match_states:=fun (s1 s2:RTL.state) => s1=s2); simpl; eauto. + - eapply senv_preserved. + - (* initial states *) + intros s1 INIT. destruct INIT as [b f m0 ge0 INIT SYMB PTR SIG]. eexists; intuition eauto. + econstructor; eauto. + + intros; eapply (Genv.init_mem_match (match_prog_RTL tprog)). apply (Genv.init_mem_match TRANSL); auto. + + rewrite symbols_preserved. + replace (prog_main (program_RTL tprog)) with (prog_main prog). + * eapply SYMB. + * erewrite (match_program_main (match_prog_RTL tprog)). erewrite (match_program_main TRANSL); auto. + + exploit function_ptr_RTL_preserved; eauto. + - intros; subst; auto. + - intros s t s2 STEP s1 H; subst. + eexists; intuition. + destruct STEP. + + (* Inop *) eapply exec_Inop; eauto. + + (* Iop *) eapply exec_Iop; eauto. + erewrite eval_operation_preserved; eauto. + + (* Iload *) eapply exec_Iload; eauto. + erewrite eval_addressing_preserved; eauto. + + (* Istore *) eapply exec_Istore; eauto. + erewrite eval_addressing_preserved; eauto. + + (* Icall *) + eapply RTL.exec_Icall; eauto. + eapply find_function_preserved; eauto. + + (* Itailcall *) + eapply RTL.exec_Itailcall; eauto. + eapply find_function_preserved; eauto. + + (* Ibuiltin *) + eapply RTL.exec_Ibuiltin; eauto. + * eapply eval_builtin_args_preserved; eauto. + * eapply external_call_symbols_preserved; eauto. + + (* Icond *) + eapply exec_Icond; eauto. + + (* Ijumptable *) + eapply RTL.exec_Ijumptable; eauto. + + (* Ireturn *) + eapply RTL.exec_Ireturn; eauto. + + (* exec_function_internal *) + eapply RTL.exec_function_internal; eauto. + + (* exec_function_external *) + eapply RTL.exec_function_external; eauto. + eapply external_call_symbols_preserved; eauto. + + (* exec_return *) + eapply RTL.exec_return; eauto. +Qed. + +Theorem transf_program_correct: + forward_simulation (RTL.semantics prog) (RTLpath.semantics tprog). Proof. - intros; eapply eqlive_reg_list; eauto. - intros; eapply list_mem_correct; eauto. + eapply compose_forward_simulations. + + eapply transf_program_RTL_correct. + + eapply RTLpath_complete. Qed. - -(* TODO: suite à revoir... *) - -Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := - | eqlive_stackframe_intro path res f sp pc rs1 rs2 - (PATH: f.(fn_path)!pc = Some path) - (LIVE: path_checker f pc path <> None) - (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): - eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). - -Inductive eqlive_states: state -> state -> Prop := - | eqlive_states_intro - path st1 st2 f sp pc rs1 rs2 m - (STACKS: list_forall2 eqlive_stackframes st1 st2) - (PATH: f.(fn_path)!pc = Some path) - (LIVE: path_checker f pc path <> None) - (EQUIV: eqlive_reg (ext path.(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 - (* TODO: add info on 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 - (* TODO: add info on f *) - (STACKS: list_forall2 eqlive_stackframes st1 st2): - eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). - -(* ESSAIS *) - -Section PRESERVATION. - -Variable prog: program. - -Let pge := Genv.globalenv prog. -Let ge := Genv.globalenv (program_RTL prog). - -Theorem step_eqlive t s1 s1' s2: - step ge pge s1 t s1' -> - (eqlive_states s1 s2) -> - exists s2', step ge pge s2 t s2' /\ (eqlive_states s1' s2'). -Proof. - destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]. - - intros EQLIVE; inv EQLIVE; simplify_someHyps. - intro PATH. - admit. - - intros EQLIVE; inv EQLIVE. - eexists; split. - * eapply exec_function_internal; eauto. - * eapply eqlive_states_intro; eauto. -Qed. - -(* -path : path_info -stack : list stackframe -f : function -sp : Values.val -rs : regset -m : Memory.Mem.mem -pc : positive -t : Events.trace -s : state -STEP : path_step ge pge (psize path) stack f sp rs m pc t s -st2 : list stackframe -rs2 : regset -STACKS : list_forall2 eqlive_stackframes stack st2 -EQUIV : eqlive_reg (ext (input_regs path)) rs rs2 -LIVE : path_checker f pc path <> None -PATH : (fn_path f) ! pc = Some path -______________________________________(1/1) -exists s2' : state, step ge pge (State st2 f sp pc rs2 m) t s2' /\ eqlive_states s s2' -*) - End PRESERVATION. - -(* - -Definition is_live (live: PMap.t Regset.t) (f: RTL.function): Prop - := False. - - -Lemma is_live_eqlive_reg pc' f live pc i rs1 rs2: - f.(fn_code)!pc = Some i -> - is_live live f -> - In pc' (successors_instr i) -> - eqlive_reg (ext live#pc) rs1 rs2 -> - eqlive_reg (ext (transfer f pc' live#pc')) rs1 rs2. -Proof. - unfold is_live; intro PC; intros. - eapply eqlive_reg_monotonic; eauto. - eapply Liveness.analyze_solution; eauto. -Qed. - -Ltac exploit_liveness pc' := refine (modusponens _ _ (is_live_eqlive_reg pc' _ _ _ _ _ _ _ _ _ _) _); [ eauto | eauto | try (simpl; intuition) | eauto | eauto]. - -(* Y aura-t-il besoin de chose comme ça ? *) - -*) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v new file mode 100644 index 00000000..0701c598 --- /dev/null +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -0,0 +1,186 @@ +(** Building a RTLpath program with liveness annotation. +*) + + +Require Import Coqlib. +Require Import Maps. +Require Import Lattice. +Require Import AST. +Require Import Op. +Require Import Registers. +Require Import Globalenvs Smallstep RTL RTLpath RTLpathLivegen. +Require Import Bool Errors. +Require Import Program. + + +Local Open Scope lazy_bool_scope. + +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. + +(* +Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). +Proof. + intuition. + - eapply Regset.remove_3; eauto. + - eapply Regset.remove_1; [| eauto]. auto. + - eapply Regset.remove_2; eauto. +Qed. + +Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). +Proof. + induction args; simpl. + + intuition. + + intros; rewrite IHargs, regset_add_spec. intuition. +Qed. +*) + + +Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := + forall r, (alive r) -> rs1#r = rs2#r. + +Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs. +Proof. + unfold eqlive_reg; auto. +Qed. + +Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1. +Proof. + unfold eqlive_reg; intros; symmetry; auto. +Qed. + +Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3. +Proof. + unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. +Qed. + +Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). +Proof. + unfold eqlive_reg; intros NOTALIVE r0 ALIVE. + destruct (Pos.eq_dec r r0) as [H|H]. + - subst; tauto. + - rewrite Regmap.gso; auto. +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. + +Local Hint Resolve Regset.mem_2 Regset.subset_2. + +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. + + +(* TODO: suite à revoir... + +Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := + | eqlive_stackframe_intro path res f sp pc rs1 rs2 + (PATH: f.(fn_path)!pc = Some path) + (LIVE: path_checker f pc path <> None) + (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): + eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + +Inductive eqlive_states: state -> state -> Prop := + | eqlive_states_intro + path st1 st2 f sp pc rs1 rs2 m + (STACKS: list_forall2 eqlive_stackframes st1 st2) + (PATH: f.(fn_path)!pc = Some path) + (LIVE: path_checker f pc path <> None) + (EQUIV: eqlive_reg (ext path.(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 + (* TODO: add info on 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 + (* TODO: add info on f *) + (STACKS: list_forall2 eqlive_stackframes st1 st2): + eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). + +(* ESSAIS *) + +Section PRESERVATION. + +Variable prog: program. + +Let pge := Genv.globalenv prog. +Let ge := Genv.globalenv (program_RTL prog). + +Theorem step_eqlive t s1 s1' s2: + step ge pge s1 t s1' -> + (eqlive_states s1 s2) -> + exists s2', step ge pge s2 t s2' /\ (eqlive_states s1' s2'). +Proof. + destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]. + - intros EQLIVE; inv EQLIVE; simplify_someHyps. + intro PATH. + admit. + - intros EQLIVE; inv EQLIVE. + eexists; split. + * eapply exec_function_internal; eauto. + * eapply eqlive_states_intro; eauto. +Qed. + +(* +path : path_info +stack : list stackframe +f : function +sp : Values.val +rs : regset +m : Memory.Mem.mem +pc : positive +t : Events.trace +s : state +STEP : path_step ge pge (psize path) stack f sp rs m pc t s +st2 : list stackframe +rs2 : regset +STACKS : list_forall2 eqlive_stackframes stack st2 +EQUIV : eqlive_reg (ext (input_regs path)) rs rs2 +LIVE : path_checker f pc path <> None +PATH : (fn_path f) ! pc = Some path +______________________________________(1/1) +exists s2' : state, step ge pge (State st2 f sp pc rs2 m) t s2' /\ eqlive_states s s2' +*) + +End PRESERVATION. +*) \ No newline at end of file -- cgit From b4b3cc4eeb3e2576cb473abe58cf2dcdc200b8d6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 25 Dec 2019 09:15:13 +0100 Subject: RTLpath: progress on liveness properties --- mppa_k1c/lib/RTLpathLiveproofs.v | 195 +++++++++++++++++++++++++++++++-------- 1 file changed, 159 insertions(+), 36 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index 0701c598..dc915345 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -1,4 +1,4 @@ -(** Building a RTLpath program with liveness annotation. +(** Proofs of the liveness properties from the liveness checker of RTLpathLivengen. *) @@ -14,6 +14,7 @@ Require Import Program. Local Open Scope lazy_bool_scope. +Local Open Scope option_monad_scope. Local Notation ext alive := (fun r => Regset.In r alive). @@ -110,77 +111,199 @@ Proof. intros; eapply list_mem_correct; eauto. Qed. +Record eqlive_istate alive (st1 st2: istate): Prop := + { eqlive_continue: continue st1 = continue st2; + eqlive_the_pc: the_pc st1 = the_pc st2; + eqlive_the_rs: eqlive_reg alive (the_rs st1) (the_rs st2); + eqlive_the_mem: (the_mem st1) = (the_mem st2) }. + +Lemma iinst_checker_eqlive ge sp pm alive i res rs1 rs2 m st1 st2: + eqlive_reg (ext alive) rs1 rs2 -> + iinst_checker pm alive i = Some res -> + istep ge i sp rs1 m = Some st1 -> + istep ge i sp rs2 m = Some st2 -> + eqlive_istate (ext (fst res)) st1 st2. +Proof. + intros EQLIVE. + destruct i; simpl; try_simplify_someHyps. + - (* Inop *) + repeat (econstructor; eauto). + - (* Iop *) + inversion_ASSERT; try_simplify_someHyps. + inversion_SOME v. intros EVAL. + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + try_simplify_someHyps. + repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + - (* Iload *) + inversion_ASSERT; try_simplify_someHyps. + inversion_SOME a0. intros EVAL. + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + inversion_SOME v; try_simplify_someHyps. + try_simplify_someHyps. + repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + - (* Istore *) + (repeat inversion_ASSERT); try_simplify_someHyps. + inversion_SOME a0. intros EVAL. + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + rewrite <- (EQLIVE r); auto. + inversion_SOME v; try_simplify_someHyps. + try_simplify_someHyps. + repeat (econstructor; simpl; eauto). + - (* Icond *) + inversion_ASSERT. + inversion_SOME b. intros EVAL. + intros ARGS; erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + try_simplify_someHyps. + repeat (econstructor; simpl; eauto). + exploit exit_checker_res; eauto. + intro; subst; simpl. auto. +Qed. + +Lemma iinst_checker_istep_continue ge sp pm alive i res rs m st: + iinst_checker pm alive i = Some res -> + istep ge i sp rs m = Some st -> + continue st = true -> + (snd res)=(the_pc st). +Proof. + intros; exploit iinst_checker_default_succ; eauto. + erewrite istep_normal_exit; eauto. + congruence. +Qed. + +Lemma ipath_checker_eqlive ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1 st2, + eqlive_reg (ext alive) rs1 rs2 -> + ipath_checker ps f pm alive pc = Some res -> + isteps ge ps f sp rs1 m pc = Some st1 -> + isteps ge ps f sp rs2 m pc = Some st2 -> eqlive_istate (ext (fst res)) st1 st2. +Proof. + induction ps as [|ps]; simpl; try_simplify_someHyps. + - repeat (econstructor; simpl; eauto). + - inversion_SOME i; try_simplify_someHyps. + inversion_SOME res0. + inversion_SOME st0. + inversion_SOME st0'. + intros. + exploit iinst_checker_eqlive; eauto. + destruct 1 as [CONT PC RS MEM]. + rewrite <- CONT, <- MEM, <- PC in * |-. + destruct (continue st0) eqn:CONT'. + * exploit iinst_checker_istep_continue; eauto. + intros X; rewrite X in * |-; clear X. + eapply IHps; eauto. + * try_simplify_someHyps. clear IHps. + repeat (econstructor; simpl; eauto). + admit. +Admitted. -(* TODO: suite à revoir... Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := | eqlive_stackframe_intro path res f sp pc rs1 rs2 + (LIVE: liveness_ok_function f) (PATH: f.(fn_path)!pc = Some path) - (LIVE: path_checker f pc path <> None) (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): - eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). + eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2). Inductive eqlive_states: state -> state -> Prop := | eqlive_states_intro path st1 st2 f sp pc rs1 rs2 m (STACKS: list_forall2 eqlive_stackframes st1 st2) + (LIVE: liveness_ok_function f) (PATH: f.(fn_path)!pc = Some path) - (LIVE: path_checker f pc path <> None) (EQUIV: eqlive_reg (ext path.(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 - (* TODO: add info on f *) + (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 - (* TODO: add info on f *) (STACKS: list_forall2 eqlive_stackframes st1 st2): eqlive_states (Returnstate st1 v m) (Returnstate st2 v m). -(* ESSAIS *) -Section PRESERVATION. +Section LivenessProperties. Variable prog: program. Let pge := Genv.globalenv prog. Let ge := Genv.globalenv (program_RTL prog). +Hypothesis all_fundef_are_checked: forall b f, + Genv.find_funct_ptr pge b = Some f -> + liveness_ok_fundef f. + +Lemma find_function_liveness ros rs f: + find_function pge ros rs = Some f -> liveness_ok_fundef f. +Proof. + destruct ros as [r|i]; simpl. + - unfold Genv.find_funct. + destruct (rs#r); try congruence. + destruct (Integers.Ptrofs.eq_dec _ _); try congruence. + eapply all_fundef_are_checked; eauto. + - destruct (Genv.find_symbol pge i); try congruence. + eapply all_fundef_are_checked; eauto. +Qed. + +Lemma path_step_eqlive path st1 f sp rs1 m pc t s1 st2 rs2: + path_step ge pge (psize path) st1 f sp rs1 m pc t s1 -> + list_forall2 eqlive_stackframes st1 st2 -> + eqlive_reg (ext (input_regs path)) rs1 rs2 -> + liveness_ok_function f -> + (fn_path f) ! pc = Some path -> + exists s2, path_step ge pge (psize path) st2 f sp rs2 m pc t s2 /\ eqlive_states s1 s2. +Proof. + intros STEP STACKS EQLIVE LIVE PC. + unfold liveness_ok_function in LIVE. + exploit LIVE; eauto. + clear LIVE; unfold path_checker. + inversion_SOME res; (* destruct res as [alive pc']. *) intros ICHECK. (* simpl. *) + inversion_SOME i; intros PC'. + destruct STEP as [st ISTEPS CONT|]. +Admitted. + Theorem step_eqlive t s1 s1' s2: step ge pge s1 t s1' -> - (eqlive_states s1 s2) -> - exists s2', step ge pge s2 t s2' /\ (eqlive_states s1' s2'). + eqlive_states s1 s2 -> + exists s2', step ge pge s2 t s2' /\ eqlive_states s1' s2'. Proof. destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]. - intros EQLIVE; inv EQLIVE; simplify_someHyps. intro PATH. - admit. - - intros EQLIVE; inv EQLIVE. + exploit path_step_eqlive; eauto. + intros (s2 & STEP2 & EQUIV2). + eexists; split; eauto. + eapply exec_path; eauto. + - intros EQLIVE; inv EQLIVE; inv LIVE. + exploit initialize_path. { eapply fn_entry_point_wf. } + intros (path & Hpath). eexists; split. * eapply exec_function_internal; eauto. * eapply eqlive_states_intro; eauto. + eapply eqlive_reg_refl. + - intros EQLIVE; inv EQLIVE. + eexists; split. + * eapply exec_function_external; eauto. + * eapply eqlive_states_return; eauto. + - intros EQLIVE; inv EQLIVE. + 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. + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intuition. Qed. -(* -path : path_info -stack : list stackframe -f : function -sp : Values.val -rs : regset -m : Memory.Mem.mem -pc : positive -t : Events.trace -s : state -STEP : path_step ge pge (psize path) stack f sp rs m pc t s -st2 : list stackframe -rs2 : regset -STACKS : list_forall2 eqlive_stackframes stack st2 -EQUIV : eqlive_reg (ext (input_regs path)) rs rs2 -LIVE : path_checker f pc path <> None -PATH : (fn_path f) ! pc = Some path -______________________________________(1/1) -exists s2' : state, step ge pge (State st2 f sp pc rs2 m) t s2' /\ eqlive_states s s2' -*) - -End PRESERVATION. -*) \ No newline at end of file +End LivenessProperties. \ No newline at end of file -- cgit From 6e11324a87ab757dd5a61bfe2110f3374b9990ff Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 26 Dec 2019 09:37:12 +0100 Subject: RTLpath: progress on liveness properties --- mppa_k1c/lib/RTLpathLivegen.v | 21 ++-- mppa_k1c/lib/RTLpathLiveproofs.v | 232 ++++++++++++++++++++++++++++++--------- 2 files changed, 193 insertions(+), 60 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index d4ae9290..58393f78 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -131,15 +131,22 @@ Proof. - firstorder. try_simplify_someHyps. Qed. -Lemma exit_list_checker_path_entry pm alive l pc: - exit_list_checker pm alive l = true -> List.In pc l -> path_entry pm pc. +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 exit_list_checker_correct pm alive l pc: + exit_list_checker pm alive l = true -> List.In pc l -> exit_checker pm alive pc tt = Some tt. Proof. intros EXIT PC; induction l; intuition. - simpl in * |-. rewrite lazy_and_Some_true in EXIT. + simpl in * |-. rewrite lazy_and_Some_tt_true in EXIT. firstorder (subst; eauto). Qed. -Local Hint Resolve exit_list_checker_path_entry. +Local Hint Resolve exit_list_checker_correct. Definition inst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option unit := match i with @@ -193,12 +200,6 @@ Proof. eapply inst_checker_wellformed; eauto. 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. - Fixpoint list_path_checker f pm (l:list (node*path_info)): bool := match l with | nil => true diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index dc915345..c022e950 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -27,24 +27,6 @@ Proof. * eapply Regset.add_2; auto. Qed. -(* -Lemma regset_remove_spec live r1 r2: Regset.In r1 (Regset.remove r2 live) <-> (Regset.In r1 live /\ r1 <> r2). -Proof. - intuition. - - eapply Regset.remove_3; eauto. - - eapply Regset.remove_1; [| eauto]. auto. - - eapply Regset.remove_2; eauto. -Qed. - -Lemma reg_list_live_spec args r: forall live, Regset.In r (reg_list_live args live) <-> List.In r args \/ (Regset.In r live). -Proof. - induction args; simpl. - + intuition. - + intros; rewrite IHargs, regset_add_spec. intuition. -Qed. -*) - - Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop := forall r, (alive r) -> rs1#r = rs2#r. @@ -117,12 +99,11 @@ Record eqlive_istate alive (st1 st2: istate): Prop := eqlive_the_rs: eqlive_reg alive (the_rs st1) (the_rs st2); eqlive_the_mem: (the_mem st1) = (the_mem st2) }. -Lemma iinst_checker_eqlive ge sp pm alive i res rs1 rs2 m st1 st2: +Lemma iinst_checker_eqlive ge sp pm alive i res rs1 rs2 m st1: eqlive_reg (ext alive) rs1 rs2 -> iinst_checker pm alive i = Some res -> istep ge i sp rs1 m = Some st1 -> - istep ge i sp rs2 m = Some st2 -> - eqlive_istate (ext (fst res)) st1 st2. + exists st2, istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext (fst res)) st1 st2. Proof. intros EQLIVE. destruct i; simpl; try_simplify_someHyps. @@ -133,7 +114,6 @@ Proof. inversion_SOME v. intros EVAL. erewrite <- eqlive_reg_listmem; eauto. try_simplify_someHyps. - try_simplify_someHyps. repeat (econstructor; simpl; eauto). eapply eqlive_reg_update. eapply eqlive_reg_monotonic; eauto. @@ -145,7 +125,6 @@ Proof. erewrite <- eqlive_reg_listmem; eauto. try_simplify_someHyps. inversion_SOME v; try_simplify_someHyps. - try_simplify_someHyps. repeat (econstructor; simpl; eauto). eapply eqlive_reg_update. eapply eqlive_reg_monotonic; eauto. @@ -155,7 +134,6 @@ Proof. (repeat inversion_ASSERT); try_simplify_someHyps. inversion_SOME a0. intros EVAL. erewrite <- eqlive_reg_listmem; eauto. - try_simplify_someHyps. rewrite <- (EQLIVE r); auto. inversion_SOME v; try_simplify_someHyps. try_simplify_someHyps. @@ -165,15 +143,14 @@ Proof. inversion_SOME b. intros EVAL. intros ARGS; erewrite <- eqlive_reg_listmem; eauto. try_simplify_someHyps. - try_simplify_someHyps. repeat (econstructor; simpl; eauto). exploit exit_checker_res; eauto. intro; subst; simpl. auto. Qed. Lemma iinst_checker_istep_continue ge sp pm alive i res rs m st: - iinst_checker pm alive i = Some res -> - istep ge i sp rs m = Some st -> + iinst_checker pm alive i = Some res -> + istep ge i sp rs m = Some st -> continue st = true -> (snd res)=(the_pc st). Proof. @@ -182,37 +159,111 @@ Proof. congruence. Qed. -Lemma ipath_checker_eqlive ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1 st2, +Lemma exit_checker_eqlive A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res rs1 rs2: + exit_checker pm alive pc v = Some res -> eqlive_reg (ext alive) rs1 rs2 -> - ipath_checker ps f pm alive pc = Some res -> - isteps ge ps f sp rs1 m pc = Some st1 -> - isteps ge ps f sp rs2 m pc = Some st2 -> eqlive_istate (ext (fst res)) st1 st2. + exists path, pm!pc = Some path /\ eqlive_reg (ext path.(input_regs)) rs1 rs2. +Proof. + unfold exit_checker. + inversion_SOME path. + inversion_ASSERT. try_simplify_someHyps. + repeat (econstructor; eauto). + intros; eapply eqlive_reg_monotonic; eauto. + intros; exploit Regset.subset_2; eauto. +Qed. + +Lemma iinst_checker_eqlive_stopped ge sp pm alive i res rs1 rs2 m st1: + eqlive_reg (ext alive) rs1 rs2 -> + istep ge i sp rs1 m = Some st1 -> + iinst_checker pm alive i = Some res -> + continue st1 = false -> + exists path st2, pm!(the_pc st1) = Some path /\ istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. +Proof. + intros EQLIVE. + set (tmp := istep ge i sp rs2). + destruct i; simpl; try_simplify_someHyps; repeat (inversion_ASSERT || inversion_SOME b); try_simplify_someHyps; try congruence. + (* Icond *) + unfold tmp; clear tmp; simpl. + intros EVAL; erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + destruct b eqn:EQb; simpl in * |-; try congruence. + intros; exploit exit_checker_eqlive; eauto. + intros (path & PATH & EQLIVE2). + repeat (econstructor; simpl; eauto). +Qed. + +Lemma ipath_checker_eqlive_normal ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1, + eqlive_reg (ext alive) rs1 rs2 -> + ipath_checker ps f pm alive pc = Some res -> + isteps ge ps f sp rs1 m pc = Some st1 -> + continue st1 = true -> + exists st2, isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext (fst res)) st1 st2. Proof. induction ps as [|ps]; simpl; try_simplify_someHyps. - repeat (econstructor; simpl; eauto). - inversion_SOME i; try_simplify_someHyps. inversion_SOME res0. inversion_SOME st0. - inversion_SOME st0'. intros. exploit iinst_checker_eqlive; eauto. - destruct 1 as [CONT PC RS MEM]. - rewrite <- CONT, <- MEM, <- PC in * |-. + destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]). + try_simplify_someHyps. + rewrite <- CONT, <- MEM, <- PC. destruct (continue st0) eqn:CONT'. - * exploit iinst_checker_istep_continue; eauto. - intros X; rewrite X in * |-; clear X. - eapply IHps; eauto. - * try_simplify_someHyps. clear IHps. - repeat (econstructor; simpl; eauto). - admit. -Admitted. + * intros; exploit iinst_checker_istep_continue; eauto. + rewrite <- PC; intros X; rewrite X in * |-. eauto. + * try_simplify_someHyps. + congruence. +Qed. +Lemma ipath_checker_isteps_continue ge ps (f:function) sp pm: forall alive pc res rs m st, + ipath_checker ps f pm alive pc = Some res -> + isteps ge ps f sp rs m pc = Some st -> + continue st = true -> + (snd res)=(the_pc st). +Proof. + induction ps as [|ps]; simpl; try_simplify_someHyps. + inversion_SOME i; try_simplify_someHyps. + inversion_SOME res0. + inversion_SOME st0. + destruct (continue st0) eqn:CONT'. + - intros; exploit iinst_checker_istep_continue; eauto. + intros EQ; rewrite EQ in * |-; clear EQ; eauto. + - try_simplify_someHyps; congruence. +Qed. + +Lemma ipath_checker_eqlive_stopped ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1, + eqlive_reg (ext alive) rs1 rs2 -> + ipath_checker ps f pm alive pc = Some res -> + isteps ge ps f sp rs1 m pc = Some st1 -> + continue st1 = false -> + exists path st2, pm!(the_pc st1) = Some path /\ isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. +Proof. + induction ps as [|ps]; simpl; try_simplify_someHyps; try congruence. + inversion_SOME i; try_simplify_someHyps. + inversion_SOME res0. + inversion_SOME st0. + intros. + destruct (continue st0) eqn:CONT'; try_simplify_someHyps; intros. + * intros; exploit iinst_checker_eqlive; eauto. + destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]). + exploit iinst_checker_istep_continue; eauto. + intros PC'. + try_simplify_someHyps. + rewrite PC', <- CONT, <- MEM, <- PC, CONT'. + eauto. + * intros; exploit iinst_checker_eqlive_stopped; eauto. + intros EQLIVE; generalize EQLIVE; destruct 1 as (path & st2 & PATH & ISTEP & [CONT PC RS MEM]). + try_simplify_someHyps. + rewrite <- CONT, <- MEM, <- PC, CONT'. + try_simplify_someHyps. +Qed. Inductive eqlive_stackframes: stackframe -> stackframe -> Prop := - | eqlive_stackframe_intro path res f sp pc rs1 rs2 + | eqlive_stackframes_intro path res f sp pc rs1 rs2 (LIVE: liveness_ok_function f) (PATH: f.(fn_path)!pc = Some path) - (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): + (EQUIV: forall v, eqlive_reg (ext path.(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 := @@ -255,13 +306,97 @@ Proof. eapply all_fundef_are_checked; eauto. Qed. -Lemma path_step_eqlive path st1 f sp rs1 m pc t s1 st2 rs2: - path_step ge pge (psize path) st1 f sp rs1 m pc t s1 -> - list_forall2 eqlive_stackframes st1 st2 -> +Lemma inst_checker_from_iinst_checker i sp rs m st pm alive: + istep ge i sp rs m = Some st -> + inst_checker pm alive i = (SOME res <- iinst_checker pm alive i IN exit_checker pm (fst res) (snd res) tt). +Proof. + destruct i; simpl; try congruence. +Qed. + +Lemma exit_checker_eqlive_ext1 (pm: path_map) (alive: Regset.t) (pc: node) r rs1 rs2: + exit_checker pm (Regset.add r alive) pc tt = Some tt -> + eqlive_reg (ext alive) rs1 rs2 -> + exists path, pm!pc = Some path /\ (forall v, eqlive_reg (ext path.(input_regs)) (rs1 # r <- v) (rs2 # r <- v)). +Proof. + unfold exit_checker. + inversion_SOME path. + 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 inst_checker_eqlive f sp alive pc i rs1 rs2 m stk1 stk2 t s1: + list_forall2 eqlive_stackframes stk1 stk2 -> + eqlive_reg (ext alive) rs1 rs2 -> + liveness_ok_function f -> + (fn_code f) ! pc = Some i -> + path_last_step ge pge stk1 f sp pc rs1 m t s1 -> + inst_checker (fn_path f) alive i = Some tt -> + exists s2, path_last_step ge pge stk2 f sp pc rs2 m t s2 /\ eqlive_states s1 s2. +Proof. + intros STACKS EQLIVE LIVENESS PC; + destruct 1 as [i' sp pc rs1 m st1| + sp pc rs1 m sig ros args res pc' fd| | | | + st1 pc rs1 m optr m']; + try_simplify_someHyps. + + (* istate *) + intros PC ISTEP. erewrite inst_checker_from_iinst_checker; eauto. + inversion_SOME res. + intros. + destruct (continue st1) eqn: CONT. + - (* CONT => true *) + exploit iinst_checker_eqlive; eauto. + destruct 1 as (st2 & ISTEP2 & [CONT' PC2 RS MEM]). + repeat (econstructor; simpl; eauto). + rewrite <- MEM, <- PC2. + exploit exit_checker_eqlive; eauto. + intros (path & PATH & EQLIVE2). + eapply eqlive_states_intro; eauto. + erewrite <- iinst_checker_istep_continue; eauto. + - (* CONT => false *) + intros; exploit iinst_checker_eqlive_stopped; eauto. + destruct 1 as (path & st2 & PATH & ISTEP2 & [CONT2 PC2 RS MEM]). + repeat (econstructor; simpl; eauto). + rewrite <- MEM, <- PC2. + eapply eqlive_states_intro; eauto. + + (* Icall *) + repeat inversion_ASSERT. intros. + exploit exit_checker_eqlive_ext1; eauto. + intros (path & PATH & EQLIVE2). + eexists; split. + - eapply exec_Icall; eauto. + admit. + - erewrite eqlive_reg_listmem; eauto. + eapply eqlive_states_call; eauto. + eapply find_function_liveness; eauto. + repeat (econstructor; eauto). + + (* Itailcall *) + admit. + + (* Ibuiltin *) + admit. + + (* Ijumptable *) + admit. + + (* Ireturn *) + repeat inversion_ASSERT. intros. + eexists; split. + - eapply exec_Ireturn; eauto. + - destruct optr; simpl in * |- *. + * erewrite (EQLIVE r); eauto. + eapply eqlive_states_return; eauto. + * eapply eqlive_states_return; eauto. +Admitted. + + +Lemma path_step_eqlive path stk1 f sp rs1 m pc t s1 stk2 rs2: + path_step ge pge (psize path) stk1 f sp rs1 m pc t s1 -> + list_forall2 eqlive_stackframes stk1 stk2 -> eqlive_reg (ext (input_regs path)) rs1 rs2 -> liveness_ok_function f -> (fn_path f) ! pc = Some path -> - exists s2, path_step ge pge (psize path) st2 f sp rs2 m pc t s2 /\ eqlive_states s1 s2. + exists s2, path_step ge pge (psize path) stk2 f sp rs2 m pc t s2 /\ eqlive_states s1 s2. Proof. intros STEP STACKS EQLIVE LIVE PC. unfold liveness_ok_function in LIVE. @@ -301,9 +436,6 @@ Proof. exists (State s2 f sp pc (rs2 # res <- vres) m); split. * apply exec_return. * eapply eqlive_states_intro; eauto. - eapply eqlive_reg_update. - eapply eqlive_reg_monotonic; eauto. - intuition. Qed. End LivenessProperties. \ No newline at end of file -- cgit From 9627ebe08f8004a2380839aeae39db16a4ab8bb9 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 26 Dec 2019 15:06:55 +0100 Subject: RTLpath: semantical properties of liveness (with most of proofs) --- configure | 2 +- mppa_k1c/lib/RTLpathLivegen.v | 19 +++++++++++- mppa_k1c/lib/RTLpathLiveproofs.v | 66 ++++++++++++++++++++++++++++++---------- 3 files changed, 69 insertions(+), 18 deletions(-) diff --git a/configure b/configure index ddcc6619..5065782e 100755 --- a/configure +++ b/configure @@ -849,7 +849,7 @@ ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure EXECUTE=k1-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __K1C_COS__ SIMU=k1-cluster -- -BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathLivegen.v RTLpathSE_theory.v\\ +BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathLivegen.v RTLpathLiveproofs.v RTLpathSE_theory.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 58393f78..2c886774 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -359,7 +359,7 @@ Qed. Local Hint Resolve symbols_preserved senv_preserved. -Theorem transf_program_RTL_correct: +Lemma transf_program_RTL_correct: forward_simulation (RTL.semantics prog) (RTL.semantics (program_RTL tprog)). Proof. eapply forward_simulation_step with (match_states:=fun (s1 s2:RTL.state) => s1=s2); simpl; eauto. @@ -417,4 +417,21 @@ Proof. + eapply RTLpath_complete. Qed. + +(* Properties used in hypothesis of [RTLpathLiveproofs.step_eqlive] theorem *) +Theorem all_fundef_are_checked b f: + Genv.find_funct_ptr tpge b = Some f -> liveness_ok_fundef f. +Proof. + unfold match_prog, match_program in TRANSL. + unfold Genv.find_funct_ptr, tpge; 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. + intuition. +Qed. + End PRESERVATION. diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index c022e950..ea182d84 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -306,6 +306,16 @@ Proof. eapply all_fundef_are_checked; eauto. Qed. +Lemma find_function_eqlive alive ros rs1 rs2: + eqlive_reg (ext alive) rs1 rs2 -> + reg_sum_mem ros alive = true -> + find_function pge ros rs1 = find_function pge ros rs2. +Proof. + intros EQLIVE. + destruct ros; simpl; auto. + intros H; erewrite (EQLIVE r); eauto. +Qed. + Lemma inst_checker_from_iinst_checker i sp rs m st pm alive: istep ge i sp rs m = Some st -> inst_checker pm alive i = (SOME res <- iinst_checker pm alive i IN exit_checker pm (fst res) (snd res) tt). @@ -328,7 +338,7 @@ Proof. rewrite regset_add_spec. intuition subst. Qed. -Lemma inst_checker_eqlive f sp alive pc i rs1 rs2 m stk1 stk2 t s1: +Lemma inst_checker_eqlive (f: function) sp alive pc i rs1 rs2 m stk1 stk2 t s1: list_forall2 eqlive_stackframes stk1 stk2 -> eqlive_reg (ext alive) rs1 rs2 -> liveness_ok_function f -> @@ -339,7 +349,10 @@ Lemma inst_checker_eqlive f sp alive pc i rs1 rs2 m stk1 stk2 t s1: Proof. intros STACKS EQLIVE LIVENESS PC; destruct 1 as [i' sp pc rs1 m st1| - sp pc rs1 m sig ros args res pc' fd| | | | + sp pc rs1 m sig ros args res pc' fd| + st1 pc rs1 m sig ros args fd m'|  + sp pc rs1 m ef args res pc' vargs t vres m'| + sp pc rs1 m arg tbl n pc' | st1 pc rs1 m optr m']; try_simplify_someHyps. + (* istate *) @@ -363,18 +376,24 @@ Proof. rewrite <- MEM, <- PC2. eapply eqlive_states_intro; eauto. + (* Icall *) - repeat inversion_ASSERT. intros. - exploit exit_checker_eqlive_ext1; eauto. - intros (path & PATH & EQLIVE2). - eexists; split. - - eapply exec_Icall; eauto. - admit. - - erewrite eqlive_reg_listmem; eauto. - eapply eqlive_states_call; eauto. - eapply find_function_liveness; eauto. - repeat (econstructor; eauto). + repeat inversion_ASSERT. intros. + exploit exit_checker_eqlive_ext1; eauto. + intros (path & PATH & EQLIVE2). + eexists; split. + - eapply exec_Icall; eauto. + erewrite <- find_function_eqlive; eauto. + - erewrite eqlive_reg_listmem; eauto. + eapply eqlive_states_call; eauto. + eapply find_function_liveness; eauto. + repeat (econstructor; eauto). + (* Itailcall *) - admit. + repeat inversion_ASSERT. intros. + eexists; split. + - eapply exec_Itailcall; eauto. + erewrite <- find_function_eqlive; eauto. + - erewrite eqlive_reg_listmem; eauto. + eapply eqlive_states_call; eauto. + eapply find_function_liveness; eauto. + (* Ibuiltin *) admit. + (* Ijumptable *) @@ -389,7 +408,6 @@ Proof. * eapply eqlive_states_return; eauto. Admitted. - Lemma path_step_eqlive path stk1 f sp rs1 m pc t s1 stk2 rs2: path_step ge pge (psize path) stk1 f sp rs1 m pc t s1 -> list_forall2 eqlive_stackframes stk1 stk2 -> @@ -401,11 +419,27 @@ Proof. intros STEP STACKS EQLIVE LIVE PC. unfold liveness_ok_function in LIVE. exploit LIVE; eauto. - clear LIVE; unfold path_checker. + unfold path_checker. inversion_SOME res; (* destruct res as [alive pc']. *) intros ICHECK. (* simpl. *) inversion_SOME i; intros PC'. destruct STEP as [st ISTEPS CONT|]. -Admitted. + - (* early_exit *) + intros; exploit ipath_checker_eqlive_stopped; eauto. + destruct 1 as (path2 & st2 & PATH & ISTEP2 & [CONT2 PC2 RS MEM]). + repeat (econstructor; simpl; eauto). + rewrite <- MEM, <- PC2. + eapply eqlive_states_intro; eauto. + - (* normal_exit *) + intros; exploit ipath_checker_eqlive_normal; eauto. + destruct 1 as (st2 & ISTEP2 & [CONT' PC2 RS MEM]). + exploit ipath_checker_isteps_continue; eauto. + intros PC3; rewrite <- PC3, <- PC2 in * |-. + exploit inst_checker_eqlive; eauto. + intros (s2 & LAST_STEP & EQLIVE2). + eexists; split; eauto. + eapply exec_normal_exit; eauto. + rewrite <- PC3, <- MEM; auto. +Qed. Theorem step_eqlive t s1 s1' s2: step ge pge s1 t s1' -> -- cgit From 443dd0e3bd31d23a494895e3f921d63777b00c13 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 27 Dec 2019 10:20:02 +0100 Subject: RTLpath: starting proofs of scheduling modulo liveness --- mppa_k1c/lib/RTLpathLiveproofs.v | 14 +++++ mppa_k1c/lib/RTLpathSE_theory.v | 10 +--- mppa_k1c/lib/RTLpathScheduler.v | 120 ++++++++++++++++++++++++++++----------- 3 files changed, 103 insertions(+), 41 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index ea182d84..deb35ca1 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -45,6 +45,7 @@ Proof. unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. Qed. +(* Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). Proof. unfold eqlive_reg; intros NOTALIVE r0 ALIVE. @@ -52,6 +53,7 @@ Proof. - subst; tauto. - rewrite Regmap.gso; auto. 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. @@ -66,6 +68,18 @@ Proof. unfold eqlive_reg; intuition. Qed. +Lemma eqlive_reg_triv rs1 rs2: (forall r, rs1#r = rs2#r) <-> eqlive_reg (fun _ => True) rs1 rs2. +Proof. + unfold eqlive_reg; intuition. +Qed. + +Lemma eqlive_reg_triv_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> (forall r, rs2#r = rs3#r) -> eqlive_reg alive rs1 rs3. +Proof. + rewrite eqlive_reg_triv; intros; eapply eqlive_reg_trans; eauto. + eapply eqlive_reg_monotonic; eauto. + simpl; eauto. +Qed. + Local Hint Resolve Regset.mem_2 Regset.subset_2. Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true. diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index a723f7b2..c1b1b0f0 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1046,12 +1046,4 @@ Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := (* IDEA: we will provide an efficient test for checking "symb_step_simu" property, by hash-consing. See usage of [symb_step_simu] in [RTLpathScheduler]. -*) - -(* TODO: à déplacer dans RTLpath ? *) -Lemma path_step_equiv pge ge path stk f sp rs1 m pc t s1 rs2: - path_step ge pge path stk f sp rs1 m pc t s1 -> - (forall r : positive, rs1 !! r = rs2 !! r) -> - exists s2, path_step ge pge path stk f sp rs2 m pc t s2 /\ equiv_state s1 s2. -Admitted. - +*) \ No newline at end of file diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 9c1ba48d..e569803f 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -6,7 +6,7 @@ This module is inspired from [Duplicate] and [Duplicateproof] Require Import AST Linking Values Maps Globalenvs Smallstep Registers. Require Import Coqlib Maps Events Errors Op. -Require Import RTL RTLpath RTLpathSE_theory. +Require Import RTL RTLpath RTLpathLivegen RTLpathLiveproofs RTLpathSE_theory. Local Open Scope error_monad_scope. @@ -60,39 +60,49 @@ Definition transf_program (p: program) : res program := (** * Preservation proof *) +Local Notation ext alive := (fun r => Regset.In r alive). + Inductive match_fundef: RTLpath.fundef -> RTLpath.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 -> stackframe -> Prop := - | match_stackframe_intro dupmap res f sp pc rs1 rs2 f' pc' + | match_stackframe_intro dupmap res f sp pc rs1 rs2 f' pc' path (TRANSF: match_function dupmap f f') (DUPLIC: dupmap!pc' = Some pc) - (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + (LIVE: liveness_ok_function f) + (PATH: f.(fn_path)!pc = Some path) + (EQUIV: forall v, eqlive_reg (ext path.(input_regs)) (rs1 # res <- v) (rs2 # res <- v)): match_stackframes (Stackframe res f sp pc rs1) (Stackframe res f' sp pc' rs2). Inductive match_states: state -> state -> Prop := - | match_states_intro dupmap st f sp pc rs1 rs2 m st' f' pc' + | match_states_intro dupmap st f sp pc rs1 rs2 m st' f' pc' path (STACKS: list_forall2 match_stackframes st st') (TRANSF: match_function dupmap f f') (DUPLIC: dupmap!pc' = Some pc) - (EQUIV: forall r : positive, rs1 !! r = rs2 !! r): + (LIVE: liveness_ok_function f) + (PATH: f.(fn_path)!pc = Some path) + (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2): match_states (State st f sp pc rs1 m) (State st' f' sp pc' rs2 m) - | match_states_call: - forall st st' f f' args m + | match_states_call st st' f f' args m (STACKS: list_forall2 match_stackframes st st') - (TRANSF: match_fundef f f'), + (TRANSF: match_fundef f f') + (LIVE: liveness_ok_fundef f): match_states (Callstate st f args m) (Callstate st' f' args m) - | match_states_return: - forall st st' v m - (STACKS: list_forall2 match_stackframes st st'), + | 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_stackframes_equiv stf1 stf2 stf3: match_stackframes stf1 stf2 -> equiv_stackframe stf2 stf3 -> match_stackframes stf1 stf3. Proof. destruct 1; intros EQ; inv EQ; try econstructor; eauto. - intros; eapply eq_trans; eauto. + intros; eapply eqlive_reg_trans; eauto. + rewrite eqlive_reg_triv in * |-. + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + simpl; auto. Qed. Lemma match_stack_equiv stk1 stk2: @@ -108,7 +118,44 @@ Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> ma Proof. Local Hint Resolve match_stack_equiv. destruct 1; intros EQ; inv EQ; econstructor; eauto. - intros; eapply eq_trans; eauto. + intros; eapply eqlive_reg_triv_trans; eauto. +Qed. +*) + +Lemma eqlive_match_stackframes stf1 stf2 stf3: + eqlive_stackframes stf1 stf2 -> match_stackframes stf2 stf3 -> match_stackframes stf1 stf3. +Proof. + destruct 1; intros MS; inv MS; try econstructor; eauto. + try_simplify_someHyps. intros; eapply eqlive_reg_trans; eauto. +Qed. + +Lemma eqlive_match_stack stk1 stk2: + list_forall2 eqlive_stackframes stk1 stk2 -> + forall stk3, list_forall2 match_stackframes stk2 stk3 -> + list_forall2 match_stackframes stk1 stk3. +Proof. + induction 1; intros stk3 MS; inv MS; econstructor; eauto. + eapply eqlive_match_stackframes; eauto. +Qed. + +Lemma eqlive_match_states s1 s2 s3: eqlive_states s1 s2 -> match_states s2 s3 -> match_states s1 s3. +Proof. + Local Hint Resolve eqlive_match_stack. + destruct 1; intros MS; inv MS; try_simplify_someHyps; econstructor; eauto. + eapply eqlive_reg_trans; eauto. +Qed. + +Lemma eqlive_stackframes_refl stf1 stf2: match_stackframes stf1 stf2 -> eqlive_stackframes stf1 stf1. +Proof. + destruct 1; econstructor; eauto. + intros; eapply eqlive_reg_refl; eauto. +Qed. + +Lemma eqlive_stacks_refl stk1 stk2: + list_forall2 match_stackframes stk1 stk2 -> list_forall2 eqlive_stackframes stk1 stk1. +Proof. + induction 1; simpl; econstructor; eauto. + eapply eqlive_stackframes_refl; eauto. Qed. Lemma transf_fundef_correct f f': @@ -140,6 +187,8 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. +Hypothesis all_fundef_are_checked: forall b f, + Genv.find_funct_ptr pge b = Some f -> liveness_ok_fundef f. Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s. Proof. @@ -269,6 +318,7 @@ Proof. intros SEM X; eapply X; eauto. Qed. +(* Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> list_forall2 match_stackframes stk stk' -> @@ -322,6 +372,7 @@ Proof. + eapply sem_normal; eauto. + eapply match_states_equiv; eauto. Qed. +*) Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: (fn_path f)!pc = Some path -> @@ -329,9 +380,11 @@ Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: list_forall2 match_stackframes stk stk' -> dupmap ! pc' = Some pc -> match_function dupmap f f' -> + liveness_ok_function f -> exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(psize) stk' f' sp rs m pc' t s' /\ match_states s s'. -Proof. - intros PATH STEP STACKS PC FUNCT. +Admitted. +(* + intros PATH STEP STACKS DUPPC MATCHF LIVE. exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } unfold symb_step_simu; intros (PATH' & SIMU). exploit initialize_path; eauto. @@ -348,30 +401,36 @@ Proof. exists s'; intuition. eapply match_states_equiv; eauto. Qed. +*) -Lemma step_simulation: - forall s1 t s1', step ge pge s1 t s1' -> - forall s2 (MS: match_states s1 s2), +Lemma step_simulation s1 t s1' s2: + step ge pge s1 t s1' -> + match_states s1 s2 -> exists s2', step tge tpge s2 t s2' /\ match_states s1' s2'. Proof. - Local Hint Resolve transf_fundef_correct. - induction 1; intros; inv MS. + Local Hint Resolve eqlive_stacks_refl transf_fundef_correct. + destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS. (* exec_path *) - - intros. + - try_simplify_someHyps. intros. + exploit path_step_eqlive; eauto. + clear STEP EQUIV rs; intros (s2 & STEP & EQLIVE). exploit exec_path_simulation; eauto. - intros (path' & s' & STEP'). clear H0. intuition. - exploit path_step_equiv; eauto. - intros (s2 & STEP & EQ). - exists s2; split. + clear STEP; intros (path' & s' & PATH' & STEP' & MATCH'). + exists s'; split. + eapply exec_path; eauto. - + eapply match_states_equiv; eauto. + + eapply eqlive_match_states; eauto. (* exec_function_internal *) - - inversion TRANSF as [f0 xf MATCHF|]; subst. eexists. split. + - inv LIVE. + exploit initialize_path. { eapply (fn_entry_point_wf f). } + destruct 1 as (path & PATH). + inversion TRANSF as [f0 xf tf MATCHF|]; subst. eexists. split. + eapply exec_function_internal. erewrite <- preserv_fnstacksize; eauto. + erewrite preserv_fnparams; eauto. - econstructor; eauto. apply preserv_entrypoint. assumption. + econstructor; eauto. + { apply preserv_entrypoint; auto. } + { apply eqlive_reg_refl. } (* exec_function_external *) - inversion TRANSF as [|]; subst. eexists. split. + econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved_RTL. @@ -380,9 +439,6 @@ Proof. - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split. + constructor. + inv H1. econstructor; eauto. - intros r. destruct (Pos.eq_dec res' r). - * subst. rewrite ! Registers.Regmap.gss; auto. - * rewrite ! Registers.Regmap.gso; auto. Qed. Theorem transf_program_correct: @@ -392,7 +448,7 @@ Proof. - eapply senv_preserved. - eapply transf_initial_states. - intros; eapply transf_final_states; eauto. - - eapply step_simulation. + - intros; eapply step_simulation; eauto. Qed. End PRESERVATION. -- cgit From fc92bddeafd13d5aa93b44a2001bb6b6f419f973 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 28 Dec 2019 10:04:05 +0100 Subject: RTLpath: specify comparison of symbolic executions modulo liveness --- mppa_k1c/lib/RTLpathLivegen.v | 2 +- mppa_k1c/lib/RTLpathLiveproofs.v | 24 ++++++++++------ mppa_k1c/lib/RTLpathSE_theory.v | 36 ++++++++++++++++-------- mppa_k1c/lib/RTLpathScheduler.v | 61 +++++++++++++++++++++++----------------- 4 files changed, 75 insertions(+), 48 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 2c886774..4ec854d3 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -419,7 +419,7 @@ Qed. (* Properties used in hypothesis of [RTLpathLiveproofs.step_eqlive] theorem *) -Theorem all_fundef_are_checked b f: +Theorem all_fundef_liveness_ok b f: Genv.find_funct_ptr tpge b = Some f -> liveness_ok_fundef f. Proof. unfold match_prog, match_program in TRANSL. diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index deb35ca1..899f023b 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -304,20 +304,26 @@ Variable prog: program. Let pge := Genv.globalenv prog. Let ge := Genv.globalenv (program_RTL prog). -Hypothesis all_fundef_are_checked: forall b f, +Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr pge b = Some f -> liveness_ok_fundef f. -Lemma find_function_liveness ros rs f: +Lemma find_funct_liveness_ok v fd: + Genv.find_funct pge 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 pge ros rs = Some f -> liveness_ok_fundef f. Proof. destruct ros as [r|i]; simpl. - - unfold Genv.find_funct. - destruct (rs#r); try congruence. - destruct (Integers.Ptrofs.eq_dec _ _); try congruence. - eapply all_fundef_are_checked; eauto. + - intros; eapply find_funct_liveness_ok; eauto. - destruct (Genv.find_symbol pge i); try congruence. - eapply all_fundef_are_checked; eauto. + eapply all_fundef_liveness_ok; eauto. Qed. Lemma find_function_eqlive alive ros rs1 rs2: @@ -398,7 +404,7 @@ Proof. erewrite <- find_function_eqlive; eauto. - erewrite eqlive_reg_listmem; eauto. eapply eqlive_states_call; eauto. - eapply find_function_liveness; eauto. + eapply find_function_liveness_ok; eauto. repeat (econstructor; eauto). + (* Itailcall *) repeat inversion_ASSERT. intros. @@ -407,7 +413,7 @@ Proof. erewrite <- find_function_eqlive; eauto. - erewrite eqlive_reg_listmem; eauto. eapply eqlive_states_call; eauto. - eapply find_function_liveness; eauto. + eapply find_function_liveness_ok; eauto. + (* Ibuiltin *) admit. + (* Ijumptable *) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index c1b1b0f0..aefe43ae 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1015,17 +1015,28 @@ Qed. End SymbValPreserved. -Definition istate_simu (srce: PTree.t node) ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). +Require Import RTLpathLivegen RTLpathLiveproofs. + +Record istate_simulive alive (srce: PTree.t node) (ist1 ist2: istate): Prop := + { simu_continue: ist1.(continue) = ist2.(continue); + simu_the_pc: srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc); + simu_the_rs: eqlive_reg alive ist1.(the_rs) ist2.(the_rs); + simu_the_mem: ist1.(the_mem) = ist2.(the_mem) }. + +Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := + if is1.(continue) then + istate_simulive (fun _ => True) srce is1 is2 + else + exists path, f.(fn_path)!(is1.(RTLpath.the_pc)) = Some path + /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2. (* NOTE: a pure semantic definition on [s_istate], for a total freedom in refinements *) -Definition s_istate_simu (srce: PTree.t node) (st1 st2: s_istate): Prop := - forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - forall rs m is1, sem_istate ge1 sp st1 rs m is1 -> - exists is2, sem_istate ge2 sp st2 rs m is2 /\ istate_simu srce is1 is2. +Definition s_istate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: s_istate): Prop := + forall sp ge1 ge2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + liveness_ok_function f -> + forall rs m is1, sem_istate ge1 sp st1 rs m is1 -> + exists is2, sem_istate ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := @@ -1037,12 +1048,13 @@ Inductive sfval_simu (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> P | Sreturn_simu os: sfval_simu srce opc1 opc2 (Sreturn os) (Sreturn os). -Definition s_state_simu srce (s1 s2: s_state): Prop := - s_istate_simu srce s1.(internal) s2.(internal) +Definition s_state_simu f srce (s1 s2: s_state): Prop := + s_istate_simu f srce s1.(internal) s2.(internal) /\ sfval_simu srce s1.(the_pc) s2.(the_pc) s1.(final) s2.(final). Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, symb_step f1 pc1 = Some st1 -> exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu srce st1 st2. + forall st1, symb_step f1 pc1 = Some st1 -> + exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu f1 srce st1 st2. (* IDEA: we will provide an efficient test for checking "symb_step_simu" property, by hash-consing. See usage of [symb_step_simu] in [RTLpathScheduler]. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index e569803f..15eb1398 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -31,8 +31,11 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := 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 -> path_entry (*fn_code f1*) (fn_path f1) pc1 -> path_entry (*fn_code f2*) (fn_path f2) pc2 /\ symb_step_simu dupmap f1 f2 pc1 pc2 + dupmap_correct: + forall pc1 pc2, + dupmap!pc2 = Some pc1 -> + path_entry (fn_path f1) pc1 -> + path_entry (fn_path f2) pc2 /\ symb_step_simu dupmap f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) @@ -93,7 +96,6 @@ Inductive match_states: state -> state -> Prop := (STACKS: list_forall2 match_stackframes st st'): match_states (Returnstate st v m) (Returnstate st' v m). -(* Lemma match_stackframes_equiv stf1 stf2 stf3: match_stackframes stf1 stf2 -> equiv_stackframe stf2 stf3 -> match_stackframes stf1 stf3. Proof. @@ -120,7 +122,6 @@ Proof. destruct 1; intros EQ; inv EQ; econstructor; eauto. intros; eapply eqlive_reg_triv_trans; eauto. Qed. -*) Lemma eqlive_match_stackframes stf1 stf2 stf3: eqlive_stackframes stf1 stf2 -> match_stackframes stf2 stf3 -> match_stackframes stf1 stf3. @@ -187,8 +188,8 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. -Hypothesis all_fundef_are_checked: forall b f, - Genv.find_funct_ptr pge b = Some f -> liveness_ok_fundef f. +Hypothesis all_fundef_liveness_ok: forall b fd, + Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd. Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s. Proof. @@ -296,7 +297,8 @@ Qed. Lemma s_find_function_preserved sp svos rs0 m0 fd: s_find_function pge ge sp svos rs0 m0 = Some fd -> exists fd', s_find_function tpge tge sp svos rs0 m0 = Some fd' - /\ transf_fundef fd = OK fd'. + /\ transf_fundef fd = OK fd' + /\ liveness_ok_fundef fd. Proof. Local Hint Resolve symbols_preserved_RTL. unfold s_find_function. destruct svos; simpl. @@ -304,66 +306,76 @@ Proof. destruct (sval_eval _ _ _ _); try congruence. intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. + eapply find_funct_liveness_ok; eauto. + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. - intros; apply function_ptr_preserved; auto. + intros; exploit function_ptr_preserved; eauto. + intros (fd' & X). eexists. intuition eauto. Qed. -Lemma sem_istate_simu dupmap sp st st' rs m is: +Lemma sem_istate_simu f dupmap sp st st' rs m is: sem_istate ge sp st rs m is -> - s_istate_simu dupmap st st' -> + liveness_ok_function f -> + s_istate_simu f dupmap st st' -> exists is', - sem_istate tge sp st' rs m is' /\ - istate_simu dupmap is is'. + sem_istate tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. - intros SEM X; eapply X; eauto. + intros SEM LIVE X; eapply X; eauto. Qed. -(* + Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> + liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> - s_istate_simu dupmap st st' -> + s_istate_simu f dupmap st st' -> sfval_simu dupmap st.(the_pc) st'.(the_pc) sv sv' -> sfval_sem pge ge sp st stk f rs0 m0 sv rs m t s -> exists s', sfval_sem tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct. - intros FUN STK SIS SFV. destruct SFV. + intros FUN LIVE STK SIS SFV. destruct SFV. - (* Snone *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. + { admit. } + { eapply eqlive_reg_refl. } - (* Scall *) intros SEM; inv SEM; simpl. exploit s_find_function_preserved; eauto. - intros (fd' & FIND & TRANSF). + intros (fd' & FIND & TRANSF & LIVE'). erewrite <- function_sig_preserved; eauto. eexists; split; econstructor; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. + simpl. repeat (econstructor; eauto). + admit. - (* Sreturn *) intros SEM; inv SEM; simpl. eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. erewrite <- sval_eval_preserved; eauto. -Qed. +Admitted. (* The main theorem on simulation of symbolic states ! *) Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> + liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> sem_state pge ge sp st stk f rs m t s -> - s_state_simu dupmap st st' -> + s_state_simu f dupmap st st' -> exists s', sem_state tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. - intros FUNC STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. + intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) exploit sem_istate_simu; eauto. - intros (is' & SEM' & (CONT' & PC' & RS' & M')). + unfold istate_simu; rewrite CONT. + intros (is' & SEM' & (path & PATH & [CONT' PC' RS' M'])). exists (State stk' f' sp (RTLpath.the_pc is') (the_rs is') (the_mem is')). split. + eapply sem_early; auto. congruence. + rewrite M'. econstructor; eauto. - (* sem_normal *) exploit sem_istate_simu; eauto. - intros (is' & SEM' & (CONT' & PC' & RS' & M')). + unfold istate_simu; rewrite CONT. + intros (is' & SEM' & [CONT' PC' RS' M']). + rewrite <- eqlive_reg_triv in RS'. exploit sem_sfval_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). exploit sfval_sem_equiv; eauto. @@ -372,7 +384,6 @@ Proof. + eapply sem_normal; eauto. + eapply match_states_equiv; eauto. Qed. -*) Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: (fn_path f)!pc = Some path -> @@ -382,8 +393,7 @@ Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: match_function dupmap f f' -> liveness_ok_function f -> exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(psize) stk' f' sp rs m pc' t s' /\ match_states s s'. -Admitted. -(* +Proof. intros PATH STEP STACKS DUPPC MATCHF LIVE. exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } unfold symb_step_simu; intros (PATH' & SIMU). @@ -401,7 +411,6 @@ Admitted. exists s'; intuition. eapply match_states_equiv; eauto. Qed. -*) Lemma step_simulation s1 t s1' s2: step ge pge s1 t s1' -> -- cgit From 6302dbfd527ff526af77e16e5fcac05b5b132a52 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 30 Dec 2019 18:18:46 +0100 Subject: RTLpath Scheduling + Liveness: removing admit lemmas --- mppa_k1c/lib/RTLpathSE_theory.v | 13 ++++++++----- mppa_k1c/lib/RTLpathScheduler.v | 24 +++++++++++++----------- 2 files changed, 21 insertions(+), 16 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index aefe43ae..51705c4a 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1039,18 +1039,21 @@ Definition s_istate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: s_ exists is2, sem_istate ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) -Inductive sfval_simu (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := +Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: - srce!opc2 = Some opc1 -> sfval_simu srce opc1 opc2 Snone Snone + srce!opc2 = Some opc1 -> + path_entry (fn_path f) opc1 -> (* redundant check ? *) + sfval_simu f srce opc1 opc2 Snone Snone | Scall_simu sig svos lsv res pc1 pc2: srce!pc2 = Some pc1 -> - sfval_simu srce opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) + path_entry (fn_path f) pc1 -> (* redundant check ? *) + sfval_simu f srce opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) | Sreturn_simu os: - sfval_simu srce opc1 opc2 (Sreturn os) (Sreturn os). + sfval_simu f srce opc1 opc2 (Sreturn os) (Sreturn os). Definition s_state_simu f srce (s1 s2: s_state): Prop := s_istate_simu f srce s1.(internal) s2.(internal) - /\ sfval_simu srce s1.(the_pc) s2.(the_pc) s1.(final) s2.(final). + /\ sfval_simu f srce s1.(the_pc) s2.(the_pc) s1.(final) s2.(final). Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, symb_step f1 pc1 = Some st1 -> diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 15eb1398..3e9861d9 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -257,7 +257,7 @@ Qed. Theorem transf_final_states s1 s2 r: final_state s1 r -> match_states s1 s2 -> final_state s2 r. Proof. - unfold final_state. + unfold final_state. intros H; inv H. intros H; inv H; simpl in * |- *; try congruence. inv H1. @@ -327,31 +327,33 @@ Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> - s_istate_simu f dupmap st st' -> - sfval_simu dupmap st.(the_pc) st'.(the_pc) sv sv' -> + (* s_istate_simu f dupmap st st' -> *) + sfval_simu f dupmap st.(the_pc) st'.(the_pc) sv sv' -> sfval_sem pge ge sp st stk f rs0 m0 sv rs m t s -> exists s', sfval_sem tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct. - intros FUN LIVE STK SIS SFV. destruct SFV. - - (* Snone *) intros SEM; inv SEM; simpl. + intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. + - (* Snone *) + exploit initialize_path; eauto. + intros (path & PATH). eexists; split; econstructor; eauto. - { admit. } - { eapply eqlive_reg_refl. } - - (* Scall *) intros SEM; inv SEM; simpl. + eapply eqlive_reg_refl. + - (* Scall *) exploit s_find_function_preserved; eauto. intros (fd' & FIND & TRANSF & LIVE'). erewrite <- function_sig_preserved; eauto. + exploit initialize_path; eauto. + intros (path & PATH). eexists; split; econstructor; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. + simpl. repeat (econstructor; eauto). - admit. - - (* Sreturn *) intros SEM; inv SEM; simpl. + - (* Sreturn *) eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. erewrite <- sval_eval_preserved; eauto. -Admitted. +Qed. (* The main theorem on simulation of symbolic states ! *) Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: -- cgit From 6929164cef5d45aea7759f9add953c689c80c41f Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 3 Jan 2020 10:54:21 +0100 Subject: RTLpath: removing some admits. --- mppa_k1c/lib/RTLpathLiveproofs.v | 50 ++++++++++++++++++++++++++++++++++++---- 1 file changed, 45 insertions(+), 5 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index 899f023b..2d2456ae 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -358,6 +358,33 @@ Proof. rewrite regset_add_spec. intuition subst. Qed. +Lemma eqlive_eval_builtin_args alive rs1 rs2 sp m args vargs: + eqlive_reg alive rs1 rs2 -> + (forall r, List.In r (params_of_builtin_args args) -> alive r) -> + Events.eval_builtin_args ge (fun r => rs1 # r) sp m args vargs -> + Events.eval_builtin_args ge (fun r => rs2 # r) sp m args vargs. +Admitted. + +Lemma exit_checker_eqlive_builtin_res (pm: path_map) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg): + exit_checker pm (reg_list_add (params_of_builtin_res res) alive) pc tt = Some tt -> + eqlive_reg (ext alive) rs1 rs2 -> + exists path, pm!pc = Some path /\ (forall vres, eqlive_reg (ext path.(input_regs)) (regmap_setres res vres rs1) (regmap_setres res vres rs2)). +Admitted. + +Lemma exit_list_checker_eqlive (pm: path_map) (alive: Regset.t) (tbl: list node) rs1 rs2 pc: forall n, + exit_list_checker pm alive tbl = true -> + eqlive_reg (ext alive) rs1 rs2 -> + list_nth_z tbl n = Some pc -> + exists path, pm!pc = Some path /\ eqlive_reg (ext path.(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 inst_checker_eqlive (f: function) sp alive pc i rs1 rs2 m stk1 stk2 t s1: list_forall2 eqlive_stackframes stk1 stk2 -> eqlive_reg (ext alive) rs1 rs2 -> @@ -370,8 +397,8 @@ Proof. intros STACKS EQLIVE LIVENESS PC; destruct 1 as [i' sp pc rs1 m st1| sp pc rs1 m sig ros args res pc' fd| - st1 pc rs1 m sig ros args fd m'|  - sp pc rs1 m ef args res pc' vargs t vres m'| + st1 pc rs1 m sig ros args fd m'| + sp pc rs1 m ef args res pc' vargs t vres m'| sp pc rs1 m arg tbl n pc' | st1 pc rs1 m optr m']; try_simplify_someHyps. @@ -415,9 +442,22 @@ Proof. eapply eqlive_states_call; eauto. eapply find_function_liveness_ok; eauto. + (* Ibuiltin *) - admit. + repeat inversion_ASSERT. intros. + exploit exit_checker_eqlive_builtin_res; eauto. + intros (path & PATH & EQLIVE2). + eexists; split. + - eapply exec_Ibuiltin; eauto. + eapply eqlive_eval_builtin_args; eauto. + intros; eapply list_mem_correct; eauto. + - repeat (econstructor; simpl; eauto). + (* Ijumptable *) - admit. + repeat inversion_ASSERT. intros. + exploit exit_list_checker_eqlive; eauto. + intros (path & PATH & EQLIVE2). + eexists; split. + - eapply exec_Ijumptable; eauto. + erewrite <- EQLIVE; eauto. + - repeat (econstructor; simpl; eauto). + (* Ireturn *) repeat inversion_ASSERT. intros. eexists; split. @@ -426,7 +466,7 @@ Proof. * erewrite (EQLIVE r); eauto. eapply eqlive_states_return; eauto. * eapply eqlive_states_return; eauto. -Admitted. +Qed. Lemma path_step_eqlive path stk1 f sp rs1 m pc t s1 stk2 rs2: path_step ge pge (psize path) stk1 f sp rs1 m pc t s1 -> -- cgit From 9711d7b0d4b112a7bad11cd176e881eb46d327d7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 4 Jan 2020 18:44:44 +0100 Subject: RTLpath: removing some admits. --- configure | 2 +- mppa_k1c/lib/RTLpathLivegen.v | 11 ++++---- mppa_k1c/lib/RTLpathLiveproofs.v | 54 +++++++++++++++++++++++++++++----------- 3 files changed, 47 insertions(+), 20 deletions(-) diff --git a/configure b/configure index 5065782e..23a48300 100755 --- a/configure +++ b/configure @@ -849,7 +849,7 @@ ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure EXECUTE=k1-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __K1C_COS__ SIMU=k1-cluster -- -BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathLivegen.v RTLpathLiveproofs.v RTLpathSE_theory.v\\ +BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v RTLpath.v RTLpathLivegen.v RTLpathLiveproofs.v RTLpathSE_theory.v RTLpathScheduler.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 4ec854d3..1db90271 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -112,10 +112,11 @@ Definition reg_option_mem (or: option reg) (alive: Regset.t) := Definition reg_sum_mem (ros: reg + ident) (alive: Regset.t) := match ros with inl r => Regset.mem r alive | inr s => true end. -Fixpoint reg_list_add (rl: list reg) (alive: Regset.t) {struct rl} : Regset.t := - match rl with - | nil => alive - | r1 :: rs => reg_list_add rs (Regset.add r1 alive) +(* 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. Fixpoint exit_list_checker (pm: path_map) (alive: Regset.t) (l: list node): bool := @@ -160,7 +161,7 @@ Definition inst_checker (pm: path_map) (alive: Regset.t) (i: instruction): optio Some tt | Ibuiltin ef args res pc' => ASSERT list_mem (params_of_builtin_args args) alive IN - exit_checker pm (reg_list_add (params_of_builtin_res res) alive) pc' tt + exit_checker pm (reg_builtin_res res alive) pc' tt | Ijumptable arg tbl => ASSERT Regset.mem arg alive IN ASSERT exit_list_checker pm alive tbl IN diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index 2d2456ae..1cdd6168 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -45,16 +45,6 @@ Proof. unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto. Qed. -(* -Lemma eqlive_reg_notin alive rs r v: ~(alive r) -> eqlive_reg alive rs (rs # r <- v). -Proof. - unfold eqlive_reg; intros NOTALIVE r0 ALIVE. - destruct (Pos.eq_dec r r0) as [H|H]. - - subst; tauto. - - rewrite Regmap.gso; auto. -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. @@ -358,18 +348,54 @@ Proof. rewrite regset_add_spec. intuition subst. 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 -> - (forall r, List.In r (params_of_builtin_args args) -> alive r) -> 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. -Admitted. +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. Lemma exit_checker_eqlive_builtin_res (pm: path_map) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg): - exit_checker pm (reg_list_add (params_of_builtin_res res) alive) pc tt = Some tt -> + exit_checker pm (reg_builtin_res res alive) pc tt = Some tt -> eqlive_reg (ext alive) rs1 rs2 -> exists path, pm!pc = Some path /\ (forall vres, eqlive_reg (ext path.(input_regs)) (regmap_setres res vres rs1) (regmap_setres res vres rs2)). -Admitted. +Proof. + destruct res; simpl. + - intros; exploit exit_checker_eqlive_ext1; eauto. + - intros; exploit exit_checker_eqlive; eauto. + intros (path & PATH & EQLIVE). + eexists; intuition eauto. + - intros; exploit exit_checker_eqlive; eauto. + intros (path & PATH & EQLIVE). + eexists; intuition eauto. +Qed. Lemma exit_list_checker_eqlive (pm: path_map) (alive: Regset.t) (tbl: list node) rs1 rs2 pc: forall n, exit_list_checker pm alive tbl = true -> -- cgit From 38258dfd105ca8ec019f3606fd48d64bb355cb5e Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 6 Jan 2020 08:30:31 +0100 Subject: RTLpath: small simplification of comparison between symbolic executions --- mppa_k1c/lib/RTLpathSE_theory.v | 14 ++++++-------- mppa_k1c/lib/RTLpathScheduler.v | 28 ++++++++++++---------------- 2 files changed, 18 insertions(+), 24 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 51705c4a..e3242e8f 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1017,18 +1017,18 @@ End SymbValPreserved. Require Import RTLpathLivegen RTLpathLiveproofs. -Record istate_simulive alive (srce: PTree.t node) (ist1 ist2: istate): Prop := - { simu_continue: ist1.(continue) = ist2.(continue); - simu_the_pc: srce!(ist2.(RTLpath.the_pc)) = Some ist1.(RTLpath.the_pc); - simu_the_rs: eqlive_reg alive ist1.(the_rs) ist2.(the_rs); - simu_the_mem: ist1.(the_mem) = ist2.(the_mem) }. +Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := + is1.(continue) = is2.(continue) + /\ eqlive_reg alive is1.(the_rs) is2.(the_rs) + /\ is1.(the_mem) = is2.(the_mem). Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := if is1.(continue) then istate_simulive (fun _ => True) srce is1 is2 else exists path, f.(fn_path)!(is1.(RTLpath.the_pc)) = Some path - /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2. + /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 + /\ srce!(is2.(RTLpath.the_pc)) = Some is1.(RTLpath.the_pc). (* NOTE: a pure semantic definition on [s_istate], for a total freedom in refinements *) Definition s_istate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: s_istate): Prop := @@ -1042,11 +1042,9 @@ Definition s_istate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: s_ Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: srce!opc2 = Some opc1 -> - path_entry (fn_path f) opc1 -> (* redundant check ? *) sfval_simu f srce opc1 opc2 Snone Snone | Scall_simu sig svos lsv res pc1 pc2: srce!pc2 = Some pc1 -> - path_entry (fn_path f) pc1 -> (* redundant check ? *) sfval_simu f srce opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) | Sreturn_simu os: sfval_simu f srce opc1 opc2 (Sreturn os) (Sreturn os). diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 3e9861d9..fb2be412 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -31,11 +31,9 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := 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 -> - path_entry (fn_path f1) pc1 -> - path_entry (fn_path f2) pc2 /\ symb_step_simu dupmap f1 f2 pc1 pc2 + dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; + dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> symb_step_simu dupmap f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) @@ -335,7 +333,7 @@ Proof. Local Hint Resolve transf_fundef_correct. intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. - (* Snone *) - exploit initialize_path; eauto. + exploit initialize_path. { eapply dupmap_path_entry1; eauto. } intros (path & PATH). eexists; split; econstructor; eauto. eapply eqlive_reg_refl. @@ -343,7 +341,7 @@ Proof. exploit s_find_function_preserved; eauto. intros (fd' & FIND & TRANSF & LIVE'). erewrite <- function_sig_preserved; eauto. - exploit initialize_path; eauto. + exploit initialize_path. { eapply dupmap_path_entry1; eauto. } intros (path & PATH). eexists; split; econstructor; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. @@ -368,7 +366,7 @@ Proof. - (* sem_early *) exploit sem_istate_simu; eauto. unfold istate_simu; rewrite CONT. - intros (is' & SEM' & (path & PATH & [CONT' PC' RS' M'])). + intros (is' & SEM' & (path & PATH & (CONT' & RS' & M') & PC')). exists (State stk' f' sp (RTLpath.the_pc is') (the_rs is') (the_mem is')). split. + eapply sem_early; auto. congruence. @@ -376,7 +374,7 @@ Proof. - (* sem_normal *) exploit sem_istate_simu; eauto. unfold istate_simu; rewrite CONT. - intros (is' & SEM' & [CONT' PC' RS' M']). + intros (is' & SEM' & (CONT' & RS' & M')). rewrite <- eqlive_reg_triv in RS'. exploit sem_sfval_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). @@ -396,16 +394,14 @@ Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s: liveness_ok_function f -> exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(psize) stk' f' sp rs m pc' t s' /\ match_states s s'. Proof. - intros PATH STEP STACKS DUPPC MATCHF LIVE. - exploit dupmap_correct; eauto. { unfold path_entry; try congruence. } - unfold symb_step_simu; intros (PATH' & SIMU). - exploit initialize_path; eauto. - clear PATH'; intros (path' & PATH'). + intros PATH STEP STACKS DUPPC MATCHF LIVE. + exploit initialize_path. { eapply dupmap_path_entry2; eauto. } + intros (path' & PATH'). exists path'. exploit (symb_step_correct f pc pge ge sp path stk rs m t s); eauto. intros (st & SYMB & SEM); clear STEP. - exploit SIMU; eauto. - intros (st' & SYMB' & SIMU'); clear SIMU SYMB. + exploit dupmap_correct; eauto. + clear SYMB; intros (st' & SYMB & SIMU). exploit sem_symb_state_simu; eauto. intros (s0 & SEM0 & MATCH). exploit symb_step_exact; eauto. -- cgit From 49be035b2ef75041fb4c7259bd288cee55a4e441 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 7 Feb 2020 16:12:45 +0100 Subject: add a TODO --- mppa_k1c/lib/RTLpathSE_theory.v | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index e3242e8f..0fd729ed 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1024,6 +1024,10 @@ Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := if is1.(continue) then + (* TODO: il faudra raffiner le (fun _ => True) si on veut autoriser l'oracle à + ajouter du "code mort" sur des registres non utilisés (loop invariant code motion à la David) + Typiquement, pour connaître la frame des registres vivants, il faudra faire une propagation en arrière + sur la dernière instruction du superblock. *) istate_simulive (fun _ => True) srce is1 is2 else exists path, f.(fn_path)!(is1.(RTLpath.the_pc)) = Some path -- cgit From d760be4554fa65f8fcfa9232f3d9ff7f9183f452 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 10 Apr 2020 11:41:20 +0200 Subject: Compatibility Coq 8.11.0 --- backend/Duplicateaux.ml | 4 ++-- backend/Inliningproof.v | 8 ++++---- backend/ValueAnalysis.v | 8 ++++---- configure | 2 +- lib/Heaps.v | 6 +++--- lib/Ordered.v | 10 ++++++---- 6 files changed, 20 insertions(+), 18 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 9ff2ae55..a64f4862 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -5,9 +5,9 @@ let rec make_identity_ptree_rec = function | [] -> PTree.empty | m::lm -> let (n, _) = m in PTree.set n n (make_identity_ptree_rec lm) -let make_identity_ptree f = make_identity_ptree_rec (PTree.elements (fn_code f)) +let make_identity_ptree f = make_identity_ptree_rec (PTree.elements f.fn_code) (* For now, identity function *) let duplicate_aux f = let pTreeId = make_identity_ptree f - in (((fn_code f), (fn_entrypoint f)), pTreeId) + in ((f.fn_code, f.fn_entrypoint), pTreeId) diff --git a/backend/Inliningproof.v b/backend/Inliningproof.v index 181f40bf..cc84b1cc 100644 --- a/backend/Inliningproof.v +++ b/backend/Inliningproof.v @@ -744,7 +744,7 @@ Lemma match_stacks_free_right: match_stacks F m m1' stk stk' sp. Proof. intros. eapply match_stacks_invariant; eauto. - intros. eapply Mem.perm_free_1; eauto. + intros. eapply Mem.perm_free_1; eauto with ordered_type. intros. eapply Mem.perm_free_3; eauto. Qed. @@ -1043,7 +1043,7 @@ Proof. eapply match_stacks_bound with (bound := sp'). eapply match_stacks_invariant; eauto. intros. eapply Mem.perm_free_3; eauto. - intros. eapply Mem.perm_free_1; eauto. + intros. eapply Mem.perm_free_1; eauto with ordered_type. intros. eapply Mem.perm_free_3; eauto. erewrite Mem.nextblock_free; eauto. red in VB; xomega. eapply agree_val_regs; eauto. @@ -1135,7 +1135,7 @@ Proof. eapply match_stacks_bound with (bound := sp'). eapply match_stacks_invariant; eauto. intros. eapply Mem.perm_free_3; eauto. - intros. eapply Mem.perm_free_1; eauto. + intros. eapply Mem.perm_free_1; eauto with ordered_type. intros. eapply Mem.perm_free_3; eauto. erewrite Mem.nextblock_free; eauto. red in VB; xomega. destruct or; simpl. apply agree_val_reg; auto. auto. @@ -1182,7 +1182,7 @@ Proof. subst b1. rewrite D in H8; inv H8. eelim Plt_strict; eauto. intros. eapply Mem.perm_alloc_1; eauto. intros. exploit Mem.perm_alloc_inv. eexact A. eauto. - rewrite dec_eq_false; auto. + rewrite dec_eq_false; auto with ordered_type. auto. auto. auto. eauto. auto. rewrite H5. apply agree_regs_init_regs. eauto. auto. inv H1; auto. congruence. auto. eapply Mem.valid_new_block; eauto. diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v index 8dbb67a7..2b233900 100644 --- a/backend/ValueAnalysis.v +++ b/backend/ValueAnalysis.v @@ -1148,10 +1148,10 @@ Proof. - constructor. - assert (Plt sp bound') by eauto with va. eapply sound_stack_public_call; eauto. apply IHsound_stack; intros. - apply INV. xomega. rewrite SAME; auto. xomega. auto. auto. + apply INV. xomega. rewrite SAME; auto with ordered_type. xomega. auto. auto. - assert (Plt sp bound') by eauto with va. eapply sound_stack_private_call; eauto. apply IHsound_stack; intros. - apply INV. xomega. rewrite SAME; auto. xomega. auto. auto. + apply INV. xomega. rewrite SAME; auto with ordered_type. xomega. auto. auto. apply bmatch_ext with m; auto. intros. apply INV. xomega. auto. auto. auto. Qed. @@ -1362,7 +1362,7 @@ Proof. apply sound_stack_exten with bc. apply sound_stack_inv with m. auto. intros. apply Q. red. eapply Plt_trans; eauto. - rewrite C; auto. + rewrite C; auto with ordered_type. exact AA. * (* public builtin call *) exploit anonymize_stack; eauto. @@ -1381,7 +1381,7 @@ Proof. apply sound_stack_exten with bc. apply sound_stack_inv with m. auto. intros. apply Q. red. eapply Plt_trans; eauto. - rewrite C; auto. + rewrite C; auto with ordered_type. exact AA. } unfold transfer_builtin in TR. diff --git a/configure b/configure index 23a48300..d7cc2567 100755 --- a/configure +++ b/configure @@ -568,7 +568,7 @@ missingtools=false echo "Testing Coq... " | tr -d '\n' coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p') case "$coq_ver" in - 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10) + 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10|8.11.0) echo "version $coq_ver -- good!";; ?*) echo "version $coq_ver -- UNSUPPORTED" diff --git a/lib/Heaps.v b/lib/Heaps.v index 9fa07a1d..85343998 100644 --- a/lib/Heaps.v +++ b/lib/Heaps.v @@ -256,14 +256,14 @@ Proof. eapply gt_heap_trans with y; eauto. red; auto. - intuition. eapply lt_heap_trans; eauto. red; auto. - eapply gt_heap_trans; eauto. red; auto. + eapply gt_heap_trans; eauto. red; auto with ordered_type. - intuition. eapply gt_heap_trans; eauto. red; auto. - rewrite e3 in *; simpl in *. intuition. eapply lt_heap_trans with y; eauto. red; auto. eapply gt_heap_trans; eauto. red; auto. - intuition. eapply lt_heap_trans with y; eauto. red; auto. - eapply gt_heap_trans; eauto. red; auto. + eapply gt_heap_trans; eauto. red; auto with ordered_type. eapply gt_heap_trans with x; eauto. red; auto. - rewrite e3 in *; simpl in *; intuition. eapply gt_heap_trans; eauto. red; auto. @@ -308,7 +308,7 @@ Proof. intros. unfold insert. case_eq (partition x h). intros a b EQ; simpl. assert (E.eq y x \/ ~E.eq y x). - destruct (E.compare y x); auto. + destruct (E.compare y x); auto with ordered_type. right; red; intros. elim (E.lt_not_eq l). apply E.eq_sym; auto. destruct H0. tauto. diff --git a/lib/Ordered.v b/lib/Ordered.v index bcf24cbd..1adbd330 100644 --- a/lib/Ordered.v +++ b/lib/Ordered.v @@ -21,6 +21,8 @@ Require Import Coqlib. Require Import Maps. Require Import Integers. +Create HintDb ordered_type. + (** The ordered type of positive numbers *) Module OrderedPositive <: OrderedType. @@ -173,17 +175,17 @@ Definition eq (x y: t) := Lemma eq_refl : forall x : t, eq x x. Proof. - intros; split; auto. + intros; split; auto with ordered_type. Qed. Lemma eq_sym : forall x y : t, eq x y -> eq y x. Proof. - unfold eq; intros. intuition auto. + unfold eq; intros. intuition auto with ordered_type. Qed. Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z. Proof. - unfold eq; intros. intuition eauto. + unfold eq; intros. intuition eauto with ordered_type. Qed. Definition lt (x y: t) := @@ -201,7 +203,7 @@ Proof. case (A.compare (fst x) (fst z)); intro. assumption. generalize (A.lt_not_eq H2); intro. elim H5. - apply A.eq_trans with (fst z). auto. auto. + apply A.eq_trans with (fst z). auto. auto with ordered_type. generalize (@A.lt_not_eq (fst z) (fst y)); intro. elim H5. apply A.lt_trans with (fst x); auto. apply A.eq_sym; auto. -- cgit From ef9103f022d76a81f5ceac0c13781cfd6535116e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 10 Apr 2020 12:26:11 +0200 Subject: Fixing the merge. Needs to add notrap load semantics + prove other admitted stuff --- mppa_k1c/lib/RTLpath.v | 71 ++++++++++++++++++++++------------------- mppa_k1c/lib/RTLpathLivegen.v | 15 +++++---- mppa_k1c/lib/RTLpathSE_theory.v | 46 +++++++++++++------------- 3 files changed, 71 insertions(+), 61 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 37ee5e70..e2e6659f 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -27,6 +27,8 @@ Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL Linking. +Declare Scope option_monad_scope. + Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end) (at level 200, X ident, A at level 100, B at level 200) : option_monad_scope. @@ -41,19 +43,21 @@ Local Open Scope option_monad_scope. (** Internal instruction = instruction with a default successor in a path. *) +(* FIXME - change successor based on prediction *) Definition default_succ (i: instruction): option node := match i with | Inop s => Some s | Iop op args res s => Some s - | Iload chunk addr args dst s => Some s + | Iload _ chunk addr args dst s => Some s | Istore chunk addr args src s => Some s - | Icond cond args ifso ifnot => Some ifnot + | Icond cond args ifso ifnot _ => Some ifnot | _ => None (* TODO: we could choose a successor for jumptable ? *) end. +(* FIXME - change early_exit based on prediction *) Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *) match i with - | Icond cond args ifso ifnot => Some ifso + | Icond cond args ifso ifnot _ => Some ifso | _ => None end. @@ -139,13 +143,14 @@ Coercion program_RTL: program >-> RTL.program. Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. +(* FIXME - prediction + no trapping load *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with | Inop pc' => Some (ist true pc' rs m) | Iop op args res pc' => SOME v <- eval_operation ge sp op rs##args m IN Some (ist true pc' (rs#res <- v) m) - | Iload chunk addr args dst pc' => + | Iload _ chunk addr args dst pc' => SOME a <- eval_addressing ge sp addr rs##args IN SOME v <- Mem.loadv chunk m a IN Some (ist true pc' (rs#dst <- v) m) @@ -153,7 +158,7 @@ Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem) SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN Some (ist true pc' rs m') - | Icond cond args ifso ifnot => + | Icond cond args ifso ifnot _ => SOME b <- eval_condition cond rs##args m IN Some (ist (negb b) (if b then ifso else ifnot) rs m) | _ => None (* TODO jumptable ? *) @@ -363,7 +368,7 @@ This way can be viewed as a correctness property: all transitions in RTLpath are *) -Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond. +Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond: core. Lemma istep_correct ge i stack (f:function) sp rs m st : istep ge i sp rs m = Some st -> @@ -373,7 +378,7 @@ Proof. destruct i; simpl; try congruence; simplify_SOME x. Qed. -Local Hint Resolve star_refl. +Local Hint Resolve star_refl: core. Lemma isteps_correct ge path stack f sp: forall rs m pc st, isteps ge path f sp rs m pc = Some st -> @@ -411,7 +416,7 @@ Proof. eapply istep_correct; eauto. Qed. -Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro. +Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro: core. Section CORRECTNESS. @@ -443,7 +448,7 @@ Proof. intros (cuint & tf & H1 & H2 & H3); subst; auto. Qed. -Local Hint Resolve istep_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match. +Local Hint Resolve istep_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match: core. Lemma path_last_step_correct stack f sp pc rs m t s: path_last_step ge pge stack f sp pc rs m t s -> @@ -464,7 +469,7 @@ Proof. auto. Qed. -Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return. +Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return: core. Lemma step_correct s t s': step ge pge s t s' -> plus RTL.step ge s t s'. Proof. @@ -527,8 +532,8 @@ Lemma wellformed_suffix_path c pm path path': exists pc', nth_default_succ c (path-path') pc = Some pc' /\ wellformed_path c pm path' pc'. Proof. induction 1 as [|m]. - + intros; cutrewrite (path'-path'=O)%nat; [simpl;eauto|omega]. - + intros pc WF; cutrewrite (S m-path'=S (m-path'))%nat; [simpl;eauto|omega]. + + intros. enough (path'-path'=0)%nat as ->; [simpl;eauto|omega]. + + intros pc WF; enough (S m-path'=S (m-path'))%nat as ->; [simpl;eauto|omega]. inversion WF; subst; clear WF; intros; simplify_someHyps. intros; simplify_someHyps; eauto. Qed. @@ -547,7 +552,7 @@ Proof. set (ps:=path.(psize)). exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps O); omega || eauto. destruct 1 as (pc' & NTH_SUCC & WF'); auto. - cutrewrite (ps - 0 = ps)%nat in NTH_SUCC; try omega. + assert (ps - 0 = ps)%nat as HH by omega. rewrite HH in NTH_SUCC. clear HH. unfold nth_default_succ_inst. inversion WF'; clear WF'; subst. simplify_someHyps; eauto. Qed. @@ -564,7 +569,7 @@ Proof. set (ps:=path0.(psize)). intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps (ps-path)); eauto. { omega. } destruct 1 as (pc' & NTH_SUCC & WF'). - cutrewrite (ps - (ps - path) = path)%nat in NTH_SUCC; try omega. + assert (ps - (ps - path) = path)%nat as HH by omega. rewrite HH in NTH_SUCC. clear HH. unfold nth_default_succ_inst. inversion WF'; clear WF'; subst. { omega. } simplify_someHyps; eauto. @@ -574,7 +579,7 @@ Lemma initialize_path (*c*) pm n: path_entry (*c*) pm n -> exists path, pm!n = S Proof. unfold path_entry; destruct pm!n; eauto. intuition congruence. Qed. -Local Hint Resolve fn_entry_point_wf. +Local Hint Resolve fn_entry_point_wf: core. Local Opaque path_entry. Lemma istep_successors ge i sp rs m st: @@ -642,7 +647,6 @@ Proof. induction path as [|path]; simpl. - intros; try_simplify_someHyps; try congruence. - intros rs m pc st; inversion_SOME i; inversion_SOME st0. - cutrewrite (path-0=path)%nat; try omega. destruct (continue st0) eqn: CONT. + intros STEP PC STEPS CONT0. exploit IHpath; eauto. clear STEPS. @@ -671,8 +675,8 @@ Proof. Qed. Inductive is_early_exit pc: instruction -> Prop := - | Icond_early_exit cond args ifnot: - is_early_exit pc (Icond cond args pc ifnot) + | Icond_early_exit cond args ifnot predict: + is_early_exit pc (Icond cond args pc ifnot predict) . (* TODO add jumptable here ? *) Lemma istep_early_exit ge i sp rs m st : @@ -680,7 +684,7 @@ Lemma istep_early_exit ge i sp rs m st : st.(continue) = false -> st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. Proof. - Local Hint Resolve Icond_early_exit. + Local Hint Resolve Icond_early_exit: core. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. destruct b; simpl in * |- *; intuition eauto. congruence. @@ -727,7 +731,7 @@ Lemma wf_stackframe_nil: wf_stackframe nil. Proof. unfold wf_stackframe; simpl. tauto. Qed. -Local Hint Resolve wf_stackframe_nil. +Local Hint Resolve wf_stackframe_nil: core. Lemma wf_stackframe_cons st stack: wf_stackframe (st::stack) <-> (wf_stf st) /\ wf_stackframe stack. @@ -764,6 +768,7 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := /\ wf_stackframe (stack_of s2). (** *** Auxiliary lemmas of completeness *) +(* FIXME - needs to update to take into account the load_notrap *) Lemma istep_complete t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> (fn_code f)!pc = Some i -> @@ -772,7 +777,7 @@ Lemma istep_complete t i stack f sp rs m pc s': Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. -Qed. +Admitted. Lemma stuttering path idx stack f sp rs m pc st t s1': isteps ge (path.(psize)-(S idx)) f sp rs m pc = Some st -> @@ -790,7 +795,7 @@ Proof. intros (SILENT & st0 & STEP0 & EQ). intuition; subst; unfold match_inst_states; simpl. intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto. - set (ps:=path.(psize)). cutrewrite (ps - idx = S (ps - (S idx)))%nat; try omega. + set (ps:=path.(psize)). enough (ps - idx = S (ps - (S idx)))%nat as ->; try omega. erewrite <- isteps_step_right; eauto. Qed. @@ -804,7 +809,7 @@ Lemma normal_exit path stack f sp rs m pc st t s1': (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' /\ (exists idx', match_states idx' s1' s2'). Proof. - Local Hint Resolve istep_successors list_nth_z_in. (* Hint for path_entry proofs *) + Local Hint Resolve istep_successors list_nth_z_in: core. (* Hint for path_entry proofs *) intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto. intros (i & Hi & SUCCS). unfold nth_default_succ_inst in Hi. @@ -819,7 +824,7 @@ Proof. destruct (initialize_path (*fn_code f*) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. exists (path0.(psize)); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. + * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto. - (* Icall *) @@ -842,7 +847,7 @@ Proof. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0.(psize); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. + * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - (* Ijumptable *) intros; exploit exec_Ijumptable; eauto. @@ -851,7 +856,7 @@ Proof. destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto. exists path0.(psize); intuition eauto. econstructor; eauto. - * cutrewrite (path0.(psize)-path0.(psize)=0)%nat; simpl; eauto || omega. + * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - (* Ireturn *) intros; exploit exec_Ireturn; eauto. @@ -874,11 +879,11 @@ Lemma path_step_complete stack f sp rs m pc t s1' idx path st: ) /\ match_states idx' s1' s2'. Proof. - Local Hint Resolve exec_early_exit exec_normal_exit. + Local Hint Resolve exec_early_exit exec_normal_exit: core. intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. destruct idx as [ | idx]. + (* path_step on normal_exit *) - cutrewrite (path.(psize)-0=path.(psize))%nat in PSTEP; [|omega]. + assert (path.(psize)-0=path.(psize))%nat as HH by omega. rewrite HH in PSTEP. clear HH. exploit normal_exit; eauto. intros (s2' & LSTEP & (idx' & MATCH)). exists idx'; exists s2'; intuition eauto. @@ -919,7 +924,7 @@ Proof. - simpl; eauto. * (* single step case *) exploit (stuttering path1 ps stack f sp (the_rs st) (the_mem st) (the_pc st)); simpl; auto. - - { rewrite Hpath1size; cutrewrite (S ps-S ps=0)%nat; try omega. simpl; eauto. } + - { rewrite Hpath1size; enough (S ps-S ps=0)%nat as ->; try omega. simpl; eauto. } - omega. - simpl; eauto. - simpl; eauto. @@ -944,7 +949,7 @@ Proof. destruct (initialize_path (*fn_code f*) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto. exists path.(psize). constructor; auto. econstructor; eauto. - - set (ps:=path.(psize)). cutrewrite (ps-ps=O)%nat; simpl; eauto. + - set (ps:=path.(psize)). enough (ps-ps=O)%nat as ->; simpl; eauto. omega. - simpl; auto. + (* exec_function_external *) @@ -963,7 +968,7 @@ Proof. destruct (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto. exists path.(psize). constructor; auto. econstructor; eauto. - - set (ps:=path.(psize)). cutrewrite (ps-ps=O)%nat; simpl; eauto. + - set (ps:=path.(psize)). enough (ps-ps=O)%nat as ->; simpl; eauto. omega. - simpl; auto. Qed. @@ -974,7 +979,7 @@ Lemma step_complete s1 t s1' idx s2: RTL.step ge s1 t s1' -> exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'. Proof. - Local Hint Resolve plus_one plus_two exec_path. + Local Hint Resolve plus_one plus_two exec_path: core. unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1. - clear His1; destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst; simpl in * |- *. intros STEP; exploit path_step_complete; eauto. @@ -1000,7 +1005,7 @@ Proof. - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2. destruct s2; simpl in * |- *; inversion H0; subst. constructor. - - Local Hint Resolve star_refl. + - Local Hint Resolve star_refl: core. intros; exploit step_complete; eauto. destruct 1 as (idx' & s2' & X). exists idx'. exists s2'. intuition (subst; eauto). diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 1db90271..6a1d7176 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -44,27 +44,28 @@ Proof. inversion_ASSERT; try_simplify_someHyps. Qed. +(* FIXME - what about trap? *) Definition iinst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option (Regset.t * node) := match i with | Inop pc' => Some (alive, pc') | Iop op args dst pc' => ASSERT list_mem args alive IN Some (Regset.add dst alive, pc') - | Iload chunk addr args dst pc' => + | Iload _ chunk addr args dst pc' => ASSERT list_mem args alive IN Some (Regset.add dst alive, pc') | Istore chunk addr args src pc' => ASSERT Regset.mem src alive IN ASSERT list_mem args alive IN Some (alive, pc') - | Icond cond args ifso ifnot => + | Icond cond args ifso ifnot _ => ASSERT list_mem args alive IN exit_checker pm alive ifso (alive, ifnot) | _ => None (* TODO jumptable ? *) end. -Local Hint Resolve exit_checker_path_entry. +Local Hint Resolve exit_checker_path_entry: core. Lemma iinst_checker_path_entry (pm: path_map) (alive: Regset.t) (i: instruction) res pc: iinst_checker pm alive i = Some res -> @@ -147,7 +148,7 @@ Proof. firstorder (subst; eauto). Qed. -Local Hint Resolve exit_list_checker_correct. +Local Hint Resolve exit_list_checker_correct: core. Definition inst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option unit := match i with @@ -358,7 +359,7 @@ Proof. Qed. -Local Hint Resolve symbols_preserved senv_preserved. +Local Hint Resolve symbols_preserved senv_preserved: core. Lemma transf_program_RTL_correct: forward_simulation (RTL.semantics prog) (RTL.semantics (program_RTL tprog)). @@ -383,6 +384,8 @@ Proof. erewrite eval_operation_preserved; eauto. + (* Iload *) eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto. + + (* Iload notrap1 *) admit. + + (* Iload notrap2 *) admit. + (* Istore *) eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto. + (* Icall *) @@ -408,7 +411,7 @@ Proof. eapply external_call_symbols_preserved; eauto. + (* exec_return *) eapply RTL.exec_return; eauto. -Qed. +Admitted. Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (RTLpath.semantics tprog). diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 0fd729ed..cf277368 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -167,7 +167,7 @@ Lemma exclude_early_NYE ge sp st rs0 m0 rs m pc: has_not_yet_exit ge sp (exits st) rs0 m0 -> False. Proof. - Local Hint Resolve is_tail_in. + Local Hint Resolve is_tail_in: core. destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. unfold has_not_yet_exit; intros NYE; rewrite NYE in EVAL; eauto. try congruence. @@ -180,7 +180,7 @@ Lemma sem_istate_exclude_incompatible_continue ge sp st rs m is1 is2: sem_istate ge sp st rs m is2 -> False. Proof. - Local Hint Resolve exclude_early_NYE. + Local Hint Resolve exclude_early_NYE: core. unfold sem_istate. intros CONT1 CONT2. rewrite CONT1, CONT2; simpl. @@ -192,7 +192,7 @@ Lemma sem_istate_determ_continue ge sp st rs m is1 is2: sem_istate ge sp st rs m is2 -> is1.(continue) = is2.(continue). Proof. - Local Hint Resolve sem_istate_exclude_incompatible_continue. + Local Hint Resolve sem_istate_exclude_incompatible_continue: core. destruct (Bool.bool_dec is1.(continue) is2.(continue)) as [|H]; auto. intros H1 H2. assert (absurd: False); intuition. destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. @@ -208,9 +208,10 @@ Proof. 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. +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. + Local Hint Resolve is_tail_cons: core. induction 1 as [|i l1 l3 T1 IND]; simpl; auto. intros T2; inversion T2; subst; auto. Qed. @@ -227,7 +228,7 @@ 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. + Local Hint Resolve is_tail_in: core. unfold has_not_yet_exit in NYE. rewrite NYE in EVAL; eauto. try congruence. @@ -238,7 +239,7 @@ Lemma sem_early_exit_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : sem_early_exit ge sp lx rs0 m0 rs2 m2 pc2 -> pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. - Local Hint Resolve exit_cond_determ eq_sym. + Local Hint Resolve exit_cond_determ eq_sym: core. destruct 1 as [ext1 lx1 TAIL1 NYE1 EVAL1 SEM1 PC1]; subst. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. assert (X:ext1=ext2). @@ -300,6 +301,7 @@ Definition smem_set (st:local_istate) (sm:smem) := Definition sist (st: s_istate) (pc: node) (nxt: local_istate): option s_istate := Some {| the_pc := pc; exits := st.(exits); curr:= nxt |}. +(* FIXME - add notrap *) Definition symb_istep (i: instruction) (st: s_istate): option s_istate := match i with | Inop pc' => @@ -309,7 +311,7 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in sist st pc' next - | Iload chunk addr args dst pc' => + | Iload _ chunk addr args dst pc' => let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in @@ -319,7 +321,7 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := let vargs := list_val_inj (List.map prev.(the_sreg) args) in let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in sist st pc' next - | Icond cond args ifso ifnot => + | Icond cond args ifso ifnot _ => let prev := st.(curr) in let vargs := list_val_inj (List.map prev.(the_sreg) args) in let ex := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in @@ -370,7 +372,7 @@ Lemma symb_istep_preserv_has_aborted i ge sp rs0 m0 st st': symb_istep i st = Some st' -> has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. Proof. - Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted. + Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted: core. unfold has_aborted. destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps; intuition eauto. (* COND *) @@ -451,8 +453,8 @@ Proof. erewrite list_sval_eval_inj; simpl; auto. erewrite ADD; simpl; auto. } + (* COND *) - Local Hint Resolve is_tail_refl. - Local Hint Unfold sem_local_istate. + Local Hint Resolve is_tail_refl: core. + Local Hint Unfold sem_local_istate: core. inversion_SOME b; intros COND. { destruct b; simpl; unfold sem_istate, sem_local_istate; simpl. - intros; econstructor; eauto; simpl. @@ -496,7 +498,7 @@ Lemma symb_isteps_correct_false ge sp path f rs0 m0 st' is: symb_isteps path f st = Some st' -> sem_istate ge sp st' rs0 m0 is. Proof. - Local Hint Resolve symb_istep_preserv_early_exits. + Local Hint Resolve symb_istep_preserv_early_exits: core. intros CONT; unfold sem_istate; rewrite CONT. induction path; simpl. + unfold sist; try_simplify_someHyps. @@ -508,7 +510,7 @@ Lemma symb_isteps_preserv_has_aborted ge sp path f rs0 m0 st': forall st, symb_isteps path f st = Some st' -> has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. Proof. - Local Hint Resolve symb_istep_preserv_has_aborted. + Local Hint Resolve symb_istep_preserv_has_aborted: core. induction path; simpl. + unfold sist; try_simplify_someHyps. + intros st; inversion_SOME i. @@ -544,7 +546,7 @@ Lemma symb_isteps_correct_true ge sp path (f:function) rs0 m0: forall st is, (isteps ge path f sp is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) . Proof. - Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF. + Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF: core. induction path; simpl. + intros st is CONT INV WF; unfold sem_istate, sist in * |- *; @@ -725,7 +727,7 @@ Proof. + (* Ibuiltin *) admit. + (* Ijumptable *) admit. + (* Ireturn *) - intros; subst. cutrewrite (v=regmap_optget or Vundef rs). + intros; subst. enough (v=regmap_optget or Vundef rs) as ->. * eapply exec_Ireturn; eauto. * intros; destruct or; simpl; congruence. Admitted. @@ -777,7 +779,7 @@ Theorem symb_step_correct f pc pge ge sp path stack rs m t s: path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, symb_step f pc = Some st /\ sem_state pge ge sp st stack f rs m t s. Proof. - Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits. + Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits: core. intros PATH STEP; unfold symb_step; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. @@ -849,13 +851,13 @@ Qed. Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk. Proof. - Local Hint Resolve equiv_stackframe_refl. + 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. + Local Hint Resolve equiv_stack_refl: core. induction s; simpl; constructor; auto. Qed. @@ -889,7 +891,7 @@ Lemma sfval_sem_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: (forall r, rs1#r = rs2#r) -> exists s', equiv_state s s' /\ sfval_sem pge ge sp st stack f rs0 m0 sv rs2 m t s'. Proof. - Local Hint Resolve equiv_stack_refl. + Local Hint Resolve equiv_stack_refl: core. destruct 1. - (* Snone *) intros; eexists; econstructor. + eapply State_equiv; eauto. @@ -913,7 +915,7 @@ Theorem symb_step_exact f pc pge ge sp path stack st 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_sem_istate. + Local Hint Resolve init_sem_istate: core. unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. @@ -967,7 +969,7 @@ Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_s Lemma sval_eval_preserved sp sv rs0 m0: sval_eval ge sp sv rs0 m0 = sval_eval ge' sp sv rs0 m0. Proof. - Local Hint Resolve symbols_preserved_RTL. + Local Hint Resolve symbols_preserved_RTL: core. induction sv using sval_mut with (P0 := fun lsv => list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0) (P1 := fun sm => smem_eval ge sp sm rs0 m0 = smem_eval ge' sp sm rs0 m0); simpl; auto. + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. -- cgit From 27119766fc56f3ae09f599ab7b5c6b1cbcd359ea Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 13 Apr 2020 09:55:23 +0200 Subject: allow Coq 8.11.1 --- configure | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure b/configure index 93d3c242..c8efcb25 100755 --- a/configure +++ b/configure @@ -568,7 +568,7 @@ missingtools=false echo "Testing Coq... " | tr -d '\n' coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p') case "$coq_ver" in - 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0) + 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0|8.11.1) echo "version $coq_ver -- good!";; ?*) echo "version $coq_ver -- UNSUPPORTED" -- cgit From 98afe356cbad770511d243be8382a661e3ffe64e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 15 Apr 2020 19:37:40 +0200 Subject: Some progress on notrap load --- mppa_k1c/lib/RTLpath.v | 34 ++++++++++++++++++++++++++-------- mppa_k1c/lib/RTLpathLiveproofs.v | 40 ++++++++++++++++++++++++++++++---------- mppa_k1c/lib/RTLpathSE_theory.v | 10 +++++----- 3 files changed, 61 insertions(+), 23 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index e2e6659f..6aa0258e 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -143,17 +143,26 @@ Coercion program_RTL: program >-> RTL.program. Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. -(* FIXME - prediction + no trapping load *) +(* FIXME - prediction *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with | Inop pc' => Some (ist true pc' rs m) | Iop op args res pc' => SOME v <- eval_operation ge sp op rs##args m IN Some (ist true pc' (rs#res <- v) m) - | Iload _ chunk addr args dst pc' => + | Iload TRAP chunk addr args dst pc' => SOME a <- eval_addressing ge sp addr rs##args IN SOME v <- Mem.loadv chunk m a IN Some (ist true pc' (rs#dst <- v) m) + | Iload NOTRAP chunk addr args dst pc' => + let default_state := ist true pc' rs#dst <- (default_notrap_load_value chunk) m in + match (eval_addressing ge sp addr rs##args) with + | None => Some default_state + | Some a => match (Mem.loadv chunk m a) with + | None => Some default_state + | Some v => Some (ist true pc' (rs#dst <- v) m) + end + end | Istore chunk addr args src pc' => SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN @@ -352,6 +361,14 @@ Ltac simplify_someHyp := | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H end. +Ltac explore_destruct := + repeat (match goal with + | [H: ?expr = ?val |- context[match ?expr with | _ => _ end]] => rewrite H + | [H: match ?var with | _ => _ end |- _] => destruct var + | [ |- context[match ?m with | _ => _ end] ] => destruct m + | _ => discriminate + end). + Ltac simplify_someHyps := repeat (simplify_someHyp; simpl in * |- *). @@ -368,7 +385,7 @@ This way can be viewed as a correctness property: all transitions in RTLpath are *) -Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond: core. +Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond RTL.exec_Iload_notrap1 RTL.exec_Iload_notrap2: core. Lemma istep_correct ge i stack (f:function) sp rs m st : istep ge i sp rs m = Some st -> @@ -376,6 +393,7 @@ Lemma istep_correct ge i stack (f:function) sp rs m st : RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). Proof. destruct i; simpl; try congruence; simplify_SOME x. + 1-3: explore_destruct; simplify_SOME x. Qed. Local Hint Resolve star_refl: core. @@ -587,7 +605,7 @@ Lemma istep_successors ge i sp rs m st: In (the_pc st) (successors_instr i). Proof. destruct i; simpl; try congruence; simplify_SOME x. - destruct x; simpl in * |- *; intuition congruence. + all: explore_destruct; simplify_SOME x. Qed. Lemma istep_normal_exit ge i sp rs m st: @@ -596,7 +614,7 @@ Lemma istep_normal_exit ge i sp rs m st: default_succ i = Some st.(the_pc). Proof. destruct i; simpl; try congruence; simplify_SOME x. - destruct x; simpl in * |- *; try congruence. + all: explore_destruct; simplify_SOME x. Qed. Lemma isteps_normal_exit ge path f sp: forall rs m pc st, @@ -686,8 +704,7 @@ Lemma istep_early_exit ge i sp rs m st : Proof. Local Hint Resolve Icond_early_exit: core. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. - destruct b; simpl in * |- *; intuition eauto. - congruence. + all: explore_destruct; simplify_SOME b; try discriminate. Qed. Section COMPLETENESS. @@ -777,7 +794,8 @@ Lemma istep_complete t i stack f sp rs m pc s': Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. -Admitted. + all: explore_destruct; simplify_SOME a. +Qed. Lemma stuttering path idx stack f sp rs m pc st t s1': isteps ge (path.(psize)-(S idx)) f sp rs m pc = Some st -> diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index 1cdd6168..a807279b 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -125,15 +125,34 @@ Proof. intuition. - (* Iload *) inversion_ASSERT; try_simplify_someHyps. - inversion_SOME a0. intros EVAL. - erewrite <- eqlive_reg_listmem; eauto. - try_simplify_someHyps. - inversion_SOME v; try_simplify_someHyps. - repeat (econstructor; simpl; eauto). - eapply eqlive_reg_update. - eapply eqlive_reg_monotonic; eauto. - intros r0; rewrite regset_add_spec. - intuition. + destruct t. (* TODO - simplify that proof ? *) + + inversion_SOME a0. intros EVAL. + erewrite <- eqlive_reg_listmem; eauto. + try_simplify_someHyps. + inversion_SOME v; try_simplify_someHyps. + repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + + erewrite <- (eqlive_reg_listmem _ _ rs1 rs2); eauto. + destruct (eval_addressing _ _ _ _). + * destruct (Memory.Mem.loadv _ _ _). + ** intros. inv H1. repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + ** intros. inv H1. repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. + * intros. inv H1. repeat (econstructor; simpl; eauto). + eapply eqlive_reg_update. + eapply eqlive_reg_monotonic; eauto. + intros r0; rewrite regset_add_spec. + intuition. - (* Istore *) (repeat inversion_ASSERT); try_simplify_someHyps. inversion_SOME a0. intros EVAL. @@ -186,6 +205,7 @@ Proof. intros EQLIVE. set (tmp := istep ge i sp rs2). destruct i; simpl; try_simplify_someHyps; repeat (inversion_ASSERT || inversion_SOME b); try_simplify_someHyps; try congruence. + 1-3: explore_destruct; simpl; try_simplify_someHyps; repeat (inversion_ASSERT || inversion_SOME b); try_simplify_someHyps; try congruence. (* Icond *) unfold tmp; clear tmp; simpl. intros EVAL; erewrite <- eqlive_reg_listmem; eauto. @@ -558,4 +578,4 @@ Proof. * eapply eqlive_states_intro; eauto. Qed. -End LivenessProperties. \ No newline at end of file +End LivenessProperties. diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index cf277368..ebea3b33 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -416,10 +416,10 @@ Proof. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. - + (* LOAD *) - inversion_SOME a0; intros ADD. + + (* LOAD *) admit. (* FIXME *) +(* inversion_SOME a0; intros ADD. { inversion_SOME v; intros LOAD; simpl. - - unfold sem_istate, sem_local_istate; simpl; intuition. + - explore_destruct; unfold sem_istate, sem_local_istate; simpl; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. @@ -434,7 +434,7 @@ Proof. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite list_sval_eval_inj; simpl; auto. - rewrite ADD; simpl; auto. } + rewrite ADD; simpl; auto. } *) + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. @@ -471,7 +471,7 @@ Proof. unfold seval_condition. erewrite list_sval_eval_inj; simpl; auto. try_simplify_someHyps. } -Qed. +Admitted. Lemma symb_istep_correct_None ge sp i st rs0 m0 rs m: sem_local_istate ge sp (st.(curr)) rs0 m0 rs m -> -- cgit From a1c535b67ea8d66caa771a7a360c11990ff9c461 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Apr 2020 23:12:10 +0200 Subject: notrap cases --- mppa_k1c/lib/RTLpathLivegen.v | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 6a1d7176..80df6e7c 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -384,8 +384,10 @@ Proof. erewrite eval_operation_preserved; eauto. + (* Iload *) eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto. - + (* Iload notrap1 *) admit. - + (* Iload notrap2 *) admit. + + (* Iload notrap1 *) eapply exec_Iload_notrap1; eauto. + erewrite eval_addressing_preserved; eauto. + + (* Iload notrap2 *) eapply exec_Iload_notrap2; eauto. + erewrite eval_addressing_preserved; eauto. + (* Istore *) eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto. + (* Icall *) @@ -411,7 +413,7 @@ Proof. eapply external_call_symbols_preserved; eauto. + (* exec_return *) eapply RTL.exec_return; eauto. -Admitted. +Qed. Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (RTLpath.semantics tprog). -- cgit From d575d205b80d2a26ecbd14cfeb1aad477b196f43 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 21 Apr 2020 14:17:00 +0200 Subject: Added and cleaned some comments on RTLpath.v --- mppa_k1c/lib/RTLpath.v | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 6aa0258e..35f6e715 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -144,6 +144,7 @@ Coercion program_RTL: program >-> RTL.program. Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. (* FIXME - prediction *) +(* Internal step through the path *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with | Inop pc' => Some (ist true pc' rs m) @@ -175,6 +176,7 @@ Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem) (** Execution of a path in a single step *) +(* Executes until a state [st] is reached where st.(continue) is false *) Fixpoint isteps ge (path:nat) (f: function) sp rs m pc: option istate := match path with | O => Some (ist true pc rs m) @@ -244,6 +246,7 @@ Definition state_RTL (s: state): RTL.state := end. Coercion state_RTL: state >-> RTL.state. +(* Used to execute the last instruction of a path (isteps is only in charge of executing the instructions before the last) *) Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop := | exec_istate i sp pc rs m st: (fn_code f)!pc = Some i -> @@ -281,6 +284,7 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m E0 (Returnstate stack (regmap_optget or Vundef rs) m'). +(* Executes an entire path *) Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := | exec_early_exit st: isteps ge path f sp rs m pc = Some st -> @@ -292,6 +296,7 @@ Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> path_step ge pge path stack f sp rs m pc t s. +(* Either internal path execution, or the usual exec_function / exec_return borrowed from RTL *) Inductive step ge pge: state -> trace -> state -> Prop := | exec_path path stack f sp rs m pc t s: (fn_path f)!pc = Some path -> @@ -387,6 +392,7 @@ This way can be viewed as a correctness property: all transitions in RTLpath are Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond RTL.exec_Iload_notrap1 RTL.exec_Iload_notrap2: core. +(* istep reflects RTL.step *) Lemma istep_correct ge i stack (f:function) sp rs m st : istep ge i sp rs m = Some st -> forall pc, (fn_code f)!pc = Some i -> @@ -398,6 +404,7 @@ Qed. Local Hint Resolve star_refl: core. +(* isteps reflects a star relation on RTL.step *) Lemma isteps_correct ge path stack f sp: forall rs m pc st, isteps ge path f sp rs m pc = Some st -> star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). @@ -692,6 +699,7 @@ Proof. destruct (continue st0) eqn: CONT0; eauto. Qed. +(* FIXME - add prediction *) Inductive is_early_exit pc: instruction -> Prop := | Icond_early_exit cond args ifnot predict: is_early_exit pc (Icond cond args pc ifnot predict) @@ -785,7 +793,6 @@ Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop := /\ wf_stackframe (stack_of s2). (** *** Auxiliary lemmas of completeness *) -(* FIXME - needs to update to take into account the load_notrap *) Lemma istep_complete t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> (fn_code f)!pc = Some i -> -- cgit From ab8be4ed9c30c97098fad4f88809d4c9f831a3a1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Apr 2020 15:12:00 +0200 Subject: Trying to cut sem_istate_exclude_abort into lemmas --- mppa_k1c/lib/RTLpathSE_theory.v | 76 ++++++++++++++++++++++++++++++++++++----- 1 file changed, 68 insertions(+), 8 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index ebea3b33..eaecbdee 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -20,6 +20,7 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ (** * Syntax and semantics of symbolic values *) +(* FIXME - might have to add non trapping loads *) (* symbolic value *) Inductive sval := | Sinput (r: reg) @@ -45,6 +46,7 @@ Fixpoint list_val_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. +(* FIXME - add non trapping load eval *) Fixpoint sval_eval (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := match sv with | Sinput r => Some (rs0#r) @@ -99,10 +101,14 @@ Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) r eval_condition cond args m. Definition has_not_yet_exit ge sp (lx: list exit_istate) rs0 m0: Prop := - forall ext, List.In ext lx -> seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. + forall ext, List.In ext lx -> + seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. + Definition has_aborted_on_test ge sp (lx: list exit_istate) rs0 m0: Prop := - exists ext, List.In ext lx /\ seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. + exists ext, List.In ext lx + /\ seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. + Inductive sem_early_exit (ge: RTL.genv) (sp:val) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := SEE_intro ext lx': @@ -263,16 +269,70 @@ Proof. intuition. Qed. -Lemma sem_istate_exclude_abort ge sp st rs m is: - has_aborted ge sp st rs m -> +Lemma sem_istate_exclude_abort_basic_continue ge sp st rs m is: + continue is = true -> + has_aborted_on_basic ge sp (curr st) rs m -> + sem_istate ge sp st rs m is -> False. +Proof. + intros CONT ABORT SEM. unfold sem_istate in SEM. rewrite CONT in SEM. + destruct SEM as (SEML & _ & _). inversion SEML as (SEML1 & SEML2 & SEML3); clear SEML. + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. + - contradiction. + - rewrite SEML2 in ABORT2. discriminate. + - inv ABORT3. rewrite SEML3 in H. discriminate. +Qed. + +Lemma sem_istate_exclude_abort_basic_nocontinue ge sp st rs0 m0 is: + continue is = false -> + has_aborted_on_basic ge sp (curr st) rs0 m0 -> + sem_istate ge sp st rs0 m0 is -> False. +Proof. +(* intros CONT ABORT SEM. unfold sem_istate in SEM. rewrite CONT in SEM. + destruct st as [pc exits iis]. simpl in *. + destruct is as [cont pc' rs' m']. simpl in *. + inv SEM. inversion H2 as (SEML1 & SEML2 & SEML3); clear H2. + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. + - Search iis. + - rewrite SEML2 in ABORT2. discriminate. + - inv ABORT3. rewrite SEML3 in H. discriminate. *) +Admitted. + + +Lemma sem_istate_exclude_abort_basic ge sp st rs m is: + has_aborted_on_basic ge sp (curr st) rs m -> + sem_istate ge sp st rs m is -> False. +Proof. + intros. destruct (continue is) eqn:CONT. + - eapply sem_istate_exclude_abort_basic_continue; eauto. + - eapply sem_istate_exclude_abort_basic_nocontinue; eauto. +Qed. + +Lemma sem_istate_exclude_abort_test ge sp st rs m is: + has_aborted_on_test ge sp (exits st) rs m -> sem_istate ge sp st rs m is -> False. Proof. - intros [[AB|[AB|AB]]|AB]; - ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; - [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). - (* TODO: découper cette grosse preuve à faire en lemmes *) Admitted. + Lemma sem_istate_exclude_abort ge sp st rs m is: + has_aborted ge sp st rs m -> + sem_istate ge sp st rs m is -> False. + Proof. + intros ABORT SEM. + inv ABORT. + - eapply sem_istate_exclude_abort_basic; eauto. + - eapply sem_istate_exclude_abort_test; eauto. +Qed. + + +(* + intros [[AB|[AB|AB]]|AB]; + ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; + [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). + (* TODO: découper cette grosse preuve à faire en lemmes *) + Admitted. + *) + + Definition istate_eq_option ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -- cgit From a6612cfead341de00ef37e4dec03dfc0f63733d8 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Apr 2020 15:16:11 +0200 Subject: list_val_inj --> list_sval_inj --- mppa_k1c/lib/RTLpathSE_theory.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index eaecbdee..1caea687 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -38,10 +38,10 @@ 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_val_inj (l: list sval): list_sval := +Fixpoint list_sval_inj (l: list sval): list_sval := match l with | nil => Snil - | v::l => Scons v (list_val_inj l) + | v::l => Scons v (list_sval_inj l) end. Local Open Scope option_monad_scope. @@ -368,22 +368,22 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := sist st pc' st.(curr) | Iop op args dst pc' => let prev := st.(curr) in - let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let vargs := list_sval_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in sist st pc' next | Iload _ chunk addr args dst pc' => let prev := st.(curr) in - let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let vargs := list_sval_inj (List.map prev.(the_sreg) args) in let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in sist st pc' next | Istore chunk addr args src pc' => let prev := st.(curr) in - let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let vargs := list_sval_inj (List.map prev.(the_sreg) args) in let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in sist st pc' next | Icond cond args ifso ifnot _ => let prev := st.(curr) in - let vargs := list_val_inj (List.map prev.(the_sreg) args) in + let vargs := list_sval_inj (List.map prev.(the_sreg) args) in let ex := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in Some {| the_pc := ifnot; exits := ex::st.(exits); curr := prev |} | _ => None (* TODO jumptable ? *) @@ -391,7 +391,7 @@ Definition symb_istep (i: instruction) (st: s_istate): option s_istate := Lemma list_sval_eval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: (forall r : reg, sval_eval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> - list_sval_eval ge sp (list_val_inj (map sreg l)) rs0 m0 = Some (rs ## l). + list_sval_eval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; auto. inversion_SOME v. @@ -738,7 +738,7 @@ Definition symb_fstep (i: instruction) (prev: local_istate): sfval := match i with | Icall sig ros args res pc => let svos := sum_left_map prev.(the_sreg) ros in - let sargs := list_val_inj (List.map prev.(the_sreg) args) in + let sargs := list_sval_inj (List.map prev.(the_sreg) args) in Scall sig svos sargs res pc | Ireturn or => let sor := SOME r <- or IN Some (prev.(the_sreg) r) in -- cgit From d765f2d91cadfaa325ef623edec6caeac1729906 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Apr 2020 18:13:59 +0200 Subject: Started general renaming for clearer and more understandable code --- mppa_k1c/lib/RTLpath.v | 112 ++++++------- mppa_k1c/lib/RTLpathLiveproofs.v | 34 ++-- mppa_k1c/lib/RTLpathSE_theory.v | 347 +++++++++++++++++++-------------------- 3 files changed, 243 insertions(+), 250 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 35f6e715..7134de10 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -141,36 +141,36 @@ Coercion program_RTL: program >-> RTL.program. (* Semantics of internal instructions (mimicking RTL semantics) *) -Record istate := ist { continue: bool; the_pc: node; the_rs: regset; the_mem: mem }. +Record istate := mk_istate { icontinue: bool; ipc: node; irs: regset; imem: mem }. (* FIXME - prediction *) (* Internal step through the path *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with - | Inop pc' => Some (ist true pc' rs m) + | Inop pc' => Some (mk_istate true pc' rs m) | Iop op args res pc' => SOME v <- eval_operation ge sp op rs##args m IN - Some (ist true pc' (rs#res <- v) m) + Some (mk_istate true pc' (rs#res <- v) m) | Iload TRAP chunk addr args dst pc' => SOME a <- eval_addressing ge sp addr rs##args IN SOME v <- Mem.loadv chunk m a IN - Some (ist true pc' (rs#dst <- v) m) + Some (mk_istate true pc' (rs#dst <- v) m) | Iload NOTRAP chunk addr args dst pc' => - let default_state := ist true pc' rs#dst <- (default_notrap_load_value chunk) m in + let default_state := mk_istate true pc' rs#dst <- (default_notrap_load_value chunk) m in match (eval_addressing ge sp addr rs##args) with | None => Some default_state | Some a => match (Mem.loadv chunk m a) with | None => Some default_state - | Some v => Some (ist true pc' (rs#dst <- v) m) + | Some v => Some (mk_istate true pc' (rs#dst <- v) m) end end | Istore chunk addr args src pc' => SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN - Some (ist true pc' rs m') + Some (mk_istate true pc' rs m') | Icond cond args ifso ifnot _ => SOME b <- eval_condition cond rs##args m IN - Some (ist (negb b) (if b then ifso else ifnot) rs m) + Some (mk_istate (negb b) (if b then ifso else ifnot) rs m) | _ => None (* TODO jumptable ? *) end. @@ -179,12 +179,12 @@ Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem) (* Executes until a state [st] is reached where st.(continue) is false *) Fixpoint isteps ge (path:nat) (f: function) sp rs m pc: option istate := match path with - | O => Some (ist true pc rs m) + | O => Some (mk_istate true pc rs m) | S p => SOME i <- (fn_code f)!pc IN SOME st <- istep ge i sp rs m IN - if st.(continue) then - isteps ge p f sp st.(the_rs) st.(the_mem) st.(the_pc) + if (icontinue st) then + isteps ge p f sp (irs st) (imem st) (ipc st) else Some st end. @@ -252,7 +252,7 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me (fn_code f)!pc = Some i -> istep ge i sp rs m = Some st -> path_last_step ge pge stack f sp pc rs m - E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + E0 (State stack f sp (ipc st) (irs st) (imem st)) | exec_Icall sp pc rs m sig ros args res pc' fd: (fn_code f)!pc = Some(Icall sig ros args res pc') -> find_function pge ros rs = Some fd -> @@ -288,12 +288,12 @@ Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> me Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop := | exec_early_exit st: isteps ge path f sp rs m pc = Some st -> - st.(continue) = false -> - path_step ge pge path stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) + (icontinue st) = false -> + path_step ge pge path stack f sp rs m pc E0 (State stack f sp (ipc st) (irs st) (imem st)) | exec_normal_exit st t s: isteps ge path f sp rs m pc = Some st -> - st.(continue) = true -> - path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem) t s -> + (icontinue st) = true -> + path_last_step ge pge stack f sp (ipc st) (irs st) (imem st) t s -> path_step ge pge path stack f sp rs m pc t s. (* Either internal path execution, or the usual exec_function / exec_return borrowed from RTL *) @@ -396,7 +396,7 @@ Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL Lemma istep_correct ge i stack (f:function) sp rs m st : istep ge i sp rs m = Some st -> forall pc, (fn_code f)!pc = Some i -> - RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)). Proof. destruct i; simpl; try congruence; simplify_SOME x. 1-3: explore_destruct; simplify_SOME x. @@ -407,12 +407,12 @@ Local Hint Resolve star_refl: core. (* isteps reflects a star relation on RTL.step *) Lemma isteps_correct ge path stack f sp: forall rs m pc st, isteps ge path f sp rs m pc = Some st -> - star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)). Proof. induction path; simpl; try_simplify_someHyps. inversion_SOME i; intros Hi. inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:cont. + destruct (icontinue st0) eqn:cont. + intros; eapply star_step. - eapply istep_correct; eauto. - simpl; eauto. @@ -425,13 +425,13 @@ Qed. Lemma isteps_correct_early_exit ge path stack f sp: forall rs m pc st, isteps ge path f sp rs m pc = Some st -> - st.(continue) = false -> - plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + st.(icontinue) = false -> + plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)). Proof. destruct path; simpl; try_simplify_someHyps; try congruence. inversion_SOME i; intros Hi. inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:cont. + destruct (icontinue st0) eqn:cont. + intros; eapply plus_left. - eapply istep_correct; eauto. - eapply isteps_correct; eauto. @@ -609,7 +609,7 @@ Local Opaque path_entry. Lemma istep_successors ge i sp rs m st: istep ge i sp rs m = Some st -> - In (the_pc st) (successors_instr i). + In (ipc st) (successors_instr i). Proof. destruct i; simpl; try congruence; simplify_SOME x. all: explore_destruct; simplify_SOME x. @@ -617,23 +617,23 @@ Qed. Lemma istep_normal_exit ge i sp rs m st: istep ge i sp rs m = Some st -> - st.(continue) = true -> - default_succ i = Some st.(the_pc). + st.(icontinue) = true -> + default_succ i = Some st.(ipc). Proof. destruct i; simpl; try congruence; simplify_SOME x. all: explore_destruct; simplify_SOME x. Qed. Lemma isteps_normal_exit ge path f sp: forall rs m pc st, - st.(continue) = true -> + st.(icontinue) = true -> isteps ge path f sp rs m pc = Some st -> - nth_default_succ (fn_code f) path pc = Some st.(the_pc). + nth_default_succ (fn_code f) path pc = Some st.(ipc). Proof. induction path; simpl. { try_simplify_someHyps. } intros rs m pc st CONT; try_simplify_someHyps. inversion_SOME i; intros Hi. inversion_SOME st0; intros Hst0. - destruct (continue st0) eqn:X; try congruence. + destruct (icontinue st0) eqn:X; try congruence. try_simplify_someHyps. intros; erewrite istep_normal_exit; eauto. Qed. @@ -644,9 +644,9 @@ Qed. *) Lemma isteps_step_right ge path f sp: forall rs m pc st i, isteps ge path f sp rs m pc = Some st -> - st.(continue) = true -> - (fn_code f)!(st.(the_pc)) = Some i -> - istep ge i sp st.(the_rs) st.(the_mem) = isteps ge (S path) f sp rs m pc. + st.(icontinue) = true -> + (fn_code f)!(st.(ipc)) = Some i -> + istep ge i sp st.(irs) st.(imem) = isteps ge (S path) f sp rs m pc. Proof. induction path. + simpl; intros; try_simplify_someHyps. simplify_SOME st. @@ -654,25 +654,25 @@ Proof. + intros rs m pc st i H. simpl in H. generalize H; clear H; simplify_SOME xx. - destruct (continue xx0) eqn: CONTxx0. + destruct (icontinue xx0) eqn: CONTxx0. * intros; erewrite IHpath; eauto. * intros; congruence. Qed. Lemma isteps_inversion_early ge path f sp: forall rs m pc st, isteps ge path f sp rs m pc = Some st -> - (continue st)=false -> + (icontinue st)=false -> exists st0 i path0, (path > path0)%nat /\ isteps ge path0 f sp rs m pc = Some st0 /\ - st0.(continue) = true /\ - (fn_code f)!(st0.(the_pc)) = Some i /\ - istep ge i sp st0.(the_rs) st0.(the_mem) = Some st. + st0.(icontinue) = true /\ + (fn_code f)!(st0.(ipc)) = Some i /\ + istep ge i sp st0.(irs) st0.(imem) = Some st. Proof. induction path as [|path]; simpl. - intros; try_simplify_someHyps; try congruence. - intros rs m pc st; inversion_SOME i; inversion_SOME st0. - destruct (continue st0) eqn: CONT. + destruct (icontinue st0) eqn: CONT. + intros STEP PC STEPS CONT0. exploit IHpath; eauto. clear STEPS. intros (st1 & i0 & path0 & BOUND & STEP1 & CONT1 & X1 & X2); auto. @@ -687,7 +687,7 @@ Qed. Lemma isteps_resize ge path0 path1 f sp rs m pc st: (path0 <= path1)%nat -> isteps ge path0 f sp rs m pc = Some st -> - (continue st)=false -> + (icontinue st)=false -> isteps ge path1 f sp rs m pc = Some st. Proof. induction 1 as [|path1]; simpl; auto. @@ -696,7 +696,7 @@ Proof. induction path1 as [|path]; simpl; auto. - intros; try_simplify_someHyps; try congruence. - intros rs m pc st; inversion_SOME i; inversion_SOME st0; intros; try_simplify_someHyps. - destruct (continue st0) eqn: CONT0; eauto. + destruct (icontinue st0) eqn: CONT0; eauto. Qed. (* FIXME - add prediction *) @@ -707,8 +707,8 @@ Inductive is_early_exit pc: instruction -> Prop := Lemma istep_early_exit ge i sp rs m st : istep ge i sp rs m = Some st -> - st.(continue) = false -> - st.(the_rs) = rs /\ st.(the_mem) = m /\ is_early_exit st.(the_pc) i. + st.(icontinue) = false -> + st.(irs) = rs /\ st.(imem) = m /\ is_early_exit st.(ipc) i. Proof. Local Hint Resolve Icond_early_exit: core. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. @@ -782,7 +782,7 @@ Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop := (fn_path f)!pc = Some path -> (idx <= path.(psize))%nat -> isteps ge (path.(psize)-idx) f sp rs m pc = Some s2 -> - s1 = State stack f sp s2.(the_pc) s2.(the_rs) s2.(the_mem) -> + s1 = State stack f sp s2.(ipc) s2.(irs) s2.(imem) -> match_inst_states_goal idx s1 (State stack f sp pc rs m). Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop := @@ -797,7 +797,7 @@ Lemma istep_complete t i stack f sp rs m pc s': RTL.step ge (State stack f sp pc rs m) t s' -> (fn_code f)!pc = Some i -> default_succ i <> None -> - t = E0 /\ exists st, istep ge i sp rs m = Some st /\ s'=(State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)). + t = E0 /\ exists st, istep ge i sp rs m = Some st /\ s'=(State stack f sp st.(ipc) st.(irs) st.(imem)). Proof. intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence; (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto. @@ -808,8 +808,8 @@ Lemma stuttering path idx stack f sp rs m pc st t s1': isteps ge (path.(psize)-(S idx)) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (S idx <= path.(psize))%nat -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + st.(icontinue) = true -> + RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' -> t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m). Proof. intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path.(psize)-(S idx))); omega || eauto. @@ -827,11 +827,11 @@ Qed. Lemma normal_exit path stack f sp rs m pc st t s1': isteps ge path.(psize) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> - st.(continue) = true -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + st.(icontinue) = true -> + RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' -> wf_stackframe stack -> exists s2', - (path_last_step ge pge stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s2' + (path_last_step ge pge stack f sp st.(ipc) st.(irs) st.(imem)) t s2' /\ (exists idx', match_states idx' s1' s2'). Proof. Local Hint Resolve istep_successors list_nth_z_in: core. (* Hint for path_entry proofs *) @@ -846,7 +846,7 @@ Proof. exploit (exec_istate ge pge); eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. - destruct (initialize_path (*fn_code f*) (fn_path f) (the_pc st0)) as (path0 & Hpath0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) (ipc st0)) as (path0 & Hpath0); eauto. exists (path0.(psize)); intuition eauto. econstructor; eauto. * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. @@ -893,19 +893,19 @@ Lemma path_step_complete stack f sp rs m pc t s1' idx path st: isteps ge (path.(psize)-idx) f sp rs m pc = Some st -> (fn_path f)!pc = Some path -> (idx <= path.(psize))%nat -> - RTL.step ge (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) t s1' -> + RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' -> wf_stackframe stack -> exists idx' s2', (path_step ge pge path.(psize) stack f sp rs m pc t s2' \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat) - \/ (exists path', path_step ge pge path.(psize) stack f sp rs m pc E0 (State stack f sp st.(the_pc) st.(the_rs) st.(the_mem)) - /\ (fn_path f)!(the_pc st) = Some path' /\ path'.(psize) = O - /\ path_step ge pge path'.(psize) stack f sp st.(the_rs) st.(the_mem) st.(the_pc) t s2') + \/ (exists path', path_step ge pge path.(psize) stack f sp rs m pc E0 (State stack f sp st.(ipc) st.(irs) st.(imem)) + /\ (fn_path f)!(ipc st) = Some path' /\ path'.(psize) = O + /\ path_step ge pge path'.(psize) stack f sp st.(irs) st.(imem) st.(ipc) t s2') ) /\ match_states idx' s1' s2'. Proof. Local Hint Resolve exec_early_exit exec_normal_exit: core. - intros PSTEP PATH BOUND RSTEP WF; destruct (st.(continue)) eqn: CONT. + intros PSTEP PATH BOUND RSTEP WF; destruct (st.(icontinue)) eqn: CONT. destruct idx as [ | idx]. + (* path_step on normal_exit *) assert (path.(psize)-0=path.(psize))%nat as HH by omega. rewrite HH in PSTEP. clear HH. @@ -920,7 +920,7 @@ Proof. exploit (isteps_resize ge (path.(psize) - idx)%nat path.(psize)); eauto; try omega. clear PSTEP; intros PSTEP. (* TODO for clarification: move the assert below into a separate lemma *) - assert (HPATH0: exists path0, (fn_path f)!(the_pc st) = Some path0). + assert (HPATH0: exists path0, (fn_path f)!(ipc st) = Some path0). { clear RSTEP. exploit isteps_inversion_early; eauto. intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). @@ -948,7 +948,7 @@ Proof. - simpl; eauto. - simpl; eauto. * (* single step case *) - exploit (stuttering path1 ps stack f sp (the_rs st) (the_mem st) (the_pc st)); simpl; auto. + exploit (stuttering path1 ps stack f sp (irs st) (imem st) (ipc st)); simpl; auto. - { rewrite Hpath1size; enough (S ps-S ps=0)%nat as ->; try omega. simpl; eauto. } - omega. - simpl; eauto. diff --git a/mppa_k1c/lib/RTLpathLiveproofs.v b/mppa_k1c/lib/RTLpathLiveproofs.v index a807279b..8e2c22cd 100644 --- a/mppa_k1c/lib/RTLpathLiveproofs.v +++ b/mppa_k1c/lib/RTLpathLiveproofs.v @@ -98,10 +98,10 @@ Proof. Qed. Record eqlive_istate alive (st1 st2: istate): Prop := - { eqlive_continue: continue st1 = continue st2; - eqlive_the_pc: the_pc st1 = the_pc st2; - eqlive_the_rs: eqlive_reg alive (the_rs st1) (the_rs st2); - eqlive_the_mem: (the_mem st1) = (the_mem st2) }. + { eqlive_continue: icontinue st1 = icontinue st2; + eqlive_ipc: ipc st1 = ipc st2; + eqlive_irs: eqlive_reg alive (irs st1) (irs st2); + eqlive_imem: (imem st1) = (imem st2) }. Lemma iinst_checker_eqlive ge sp pm alive i res rs1 rs2 m st1: eqlive_reg (ext alive) rs1 rs2 -> @@ -174,8 +174,8 @@ Qed. Lemma iinst_checker_istep_continue ge sp pm alive i res rs m st: iinst_checker pm alive i = Some res -> istep ge i sp rs m = Some st -> - continue st = true -> - (snd res)=(the_pc st). + icontinue st = true -> + (snd res)=(ipc st). Proof. intros; exploit iinst_checker_default_succ; eauto. erewrite istep_normal_exit; eauto. @@ -199,8 +199,8 @@ Lemma iinst_checker_eqlive_stopped ge sp pm alive i res rs1 rs2 m st1: eqlive_reg (ext alive) rs1 rs2 -> istep ge i sp rs1 m = Some st1 -> iinst_checker pm alive i = Some res -> - continue st1 = false -> - exists path st2, pm!(the_pc st1) = Some path /\ istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. + icontinue st1 = false -> + exists path st2, pm!(ipc st1) = Some path /\ istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. Proof. intros EQLIVE. set (tmp := istep ge i sp rs2). @@ -220,7 +220,7 @@ Lemma ipath_checker_eqlive_normal ge ps (f:function) sp pm: forall alive pc res eqlive_reg (ext alive) rs1 rs2 -> ipath_checker ps f pm alive pc = Some res -> isteps ge ps f sp rs1 m pc = Some st1 -> - continue st1 = true -> + icontinue st1 = true -> exists st2, isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext (fst res)) st1 st2. Proof. induction ps as [|ps]; simpl; try_simplify_someHyps. @@ -233,7 +233,7 @@ Proof. destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]). try_simplify_someHyps. rewrite <- CONT, <- MEM, <- PC. - destruct (continue st0) eqn:CONT'. + destruct (icontinue st0) eqn:CONT'. * intros; exploit iinst_checker_istep_continue; eauto. rewrite <- PC; intros X; rewrite X in * |-. eauto. * try_simplify_someHyps. @@ -243,14 +243,14 @@ Qed. Lemma ipath_checker_isteps_continue ge ps (f:function) sp pm: forall alive pc res rs m st, ipath_checker ps f pm alive pc = Some res -> isteps ge ps f sp rs m pc = Some st -> - continue st = true -> - (snd res)=(the_pc st). + icontinue st = true -> + (snd res)=(ipc st). Proof. induction ps as [|ps]; simpl; try_simplify_someHyps. inversion_SOME i; try_simplify_someHyps. inversion_SOME res0. inversion_SOME st0. - destruct (continue st0) eqn:CONT'. + destruct (icontinue st0) eqn:CONT'. - intros; exploit iinst_checker_istep_continue; eauto. intros EQ; rewrite EQ in * |-; clear EQ; eauto. - try_simplify_someHyps; congruence. @@ -260,15 +260,15 @@ Lemma ipath_checker_eqlive_stopped ge ps (f:function) sp pm: forall alive pc res eqlive_reg (ext alive) rs1 rs2 -> ipath_checker ps f pm alive pc = Some res -> isteps ge ps f sp rs1 m pc = Some st1 -> - continue st1 = false -> - exists path st2, pm!(the_pc st1) = Some path /\ isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. + icontinue st1 = false -> + exists path st2, pm!(ipc st1) = Some path /\ isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2. Proof. induction ps as [|ps]; simpl; try_simplify_someHyps; try congruence. inversion_SOME i; try_simplify_someHyps. inversion_SOME res0. inversion_SOME st0. intros. - destruct (continue st0) eqn:CONT'; try_simplify_someHyps; intros. + destruct (icontinue st0) eqn:CONT'; try_simplify_someHyps; intros. * intros; exploit iinst_checker_eqlive; eauto. destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]). exploit iinst_checker_istep_continue; eauto. @@ -452,7 +452,7 @@ Proof. intros PC ISTEP. erewrite inst_checker_from_iinst_checker; eauto. inversion_SOME res. intros. - destruct (continue st1) eqn: CONT. + destruct (icontinue st1) eqn: CONT. - (* CONT => true *) exploit iinst_checker_eqlive; eauto. destruct 1 as (st2 & ISTEP2 & [CONT' PC2 RS MEM]). diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 1caea687..9c0c0c0a 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -47,103 +47,104 @@ Fixpoint list_sval_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. (* FIXME - add non trapping load eval *) -Fixpoint sval_eval (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := +Fixpoint seval_sval (ge: RTL.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 <- list_sval_eval ge sp l rs0 m0 IN - SOME m <- smem_eval ge sp sm rs0 m0 IN + SOME args <- seval_list_sval ge sp l rs0 m0 IN + SOME m <- seval_smem ge sp sm rs0 m0 IN eval_operation ge sp op args m | Sload sm chunk addr lsv => - SOME args <- list_sval_eval ge sp lsv rs0 m0 IN + SOME args <- seval_list_sval ge sp lsv rs0 m0 IN SOME a <- eval_addressing ge sp addr args IN - SOME m <- smem_eval ge sp sm rs0 m0 IN + SOME m <- seval_smem ge sp sm rs0 m0 IN Mem.loadv chunk m a end -with list_sval_eval (ge: RTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := +with seval_list_sval (ge: RTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := match lsv with | Snil => Some nil | Scons sv lsv' => - SOME v <- sval_eval ge sp sv rs0 m0 IN - SOME lv <- list_sval_eval ge sp lsv' rs0 m0 IN + SOME v <- seval_sval ge sp sv rs0 m0 IN + SOME lv <- seval_list_sval ge sp lsv' rs0 m0 IN Some (v::lv) end -with smem_eval (ge: RTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem := +with seval_smem (ge: RTL.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 <- list_sval_eval ge sp lsv rs0 m0 IN + SOME args <- seval_list_sval ge sp lsv rs0 m0 IN SOME a <- eval_addressing ge sp addr args IN - SOME m <- smem_eval ge sp sm rs0 m0 IN - SOME sv <- sval_eval ge sp srce rs0 m0 IN + SOME m <- seval_smem ge sp sm rs0 m0 IN + SOME sv <- seval_sval ge sp srce rs0 m0 IN Mem.storev chunk m a sv end. (* Syntax and Semantics of local symbolic internal states *) -Record local_istate := { pre: RTL.genv -> val -> regset -> mem -> Prop; the_sreg: reg -> sval; the_smem: smem }. +(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) +Record sistate_local := { si_pre: RTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }. -Definition sem_local_istate (ge: RTL.genv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := - st.(pre) ge sp rs0 m0 - /\ smem_eval ge sp st.(the_smem) rs0 m0 = Some m - /\ forall (r:reg), sval_eval ge sp (st.(the_sreg) r) rs0 m0 = Some (rs#r). +Definition smatch_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := + st.(si_pre) ge sp rs0 m0 + /\ seval_smem ge sp st.(si_smem) rs0 m0 = Some m + /\ forall (r:reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r). -Definition has_aborted_on_basic (ge: RTL.genv) (sp:val) (st: local_istate) (rs0: regset) (m0: mem): Prop := - ~(st.(pre) ge sp rs0 m0) - \/ smem_eval ge sp st.(the_smem) rs0 m0 = None - \/ exists (r: reg), sval_eval ge sp (st.(the_sreg) r) rs0 m0 = None. +Definition sabort_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem): Prop := + ~(st.(si_pre) ge sp rs0 m0) + \/ seval_smem ge sp st.(si_smem) rs0 m0 = None + \/ exists (r: reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = None. (* Syntax and semantics of symbolic exit states *) -Record exit_istate := { the_cond: condition; cond_args: list_sval; exit_ist: local_istate; cond_ifso: node }. +Record sistate_exit := { si_cond: condition; si_scondargs: list_sval; si_elocal: sistate_local; si_ifso: node }. Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := - SOME args <- list_sval_eval ge sp lsv rs0 m0 IN - SOME m <- smem_eval ge sp sm rs0 m0 IN + SOME args <- seval_list_sval ge sp lsv rs0 m0 IN + SOME m <- seval_smem ge sp sm rs0 m0 IN eval_condition cond args m. -Definition has_not_yet_exit ge sp (lx: list exit_istate) rs0 m0: Prop := +Definition all_fallthrough ge sp (lx: list sistate_exit) rs0 m0: Prop := forall ext, List.In ext lx -> - seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false. + seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some false. -Definition has_aborted_on_test ge sp (lx: list exit_istate) rs0 m0: Prop := +Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := exists ext, List.In ext lx - /\ seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = None. + /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. -Inductive sem_early_exit (ge: RTL.genv) (sp:val) (lx:list exit_istate) (rs0: regset) (m0: mem) rs m pc: Prop := - SEE_intro ext lx': +Inductive smatch_exits (ge: RTL.genv) (sp:val) (lx:list sistate_exit) (rs0: regset) (m0: mem) rs m pc: Prop := + smatch_exits_intro ext lx': is_tail (ext::lx') lx -> - has_not_yet_exit ge sp lx' rs0 m0 -> - seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some true -> - sem_local_istate ge sp ext.(exit_ist) rs0 m0 rs m -> - ext.(cond_ifso) = pc -> - sem_early_exit ge sp lx rs0 m0 rs m pc. + all_fallthrough ge sp lx' rs0 m0 -> + seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some true -> + smatch_local ge sp ext.(si_elocal) rs0 m0 rs m -> + ext.(si_ifso) = pc -> + smatch_exits ge sp lx rs0 m0 rs m pc. (** * Syntax and Semantics of symbolic internal state *) -Record s_istate := { the_pc: node; exits: list exit_istate; curr: local_istate }. +Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. -Definition sem_istate (ge: RTL.genv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem) (is: istate): Prop := - if (is.(continue)) +Definition smatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := + if (is.(icontinue)) then - sem_local_istate ge sp st.(curr) rs0 m0 is.(the_rs) is.(the_mem) - /\ st.(the_pc) = is.(RTLpath.the_pc) - /\ has_not_yet_exit ge sp st.(exits) rs0 m0 - else sem_early_exit ge sp st.(exits) rs0 m0 is.(the_rs) is.(the_mem) is.(RTLpath.the_pc). + smatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) + /\ st.(si_pc) = is.(ipc) + /\ all_fallthrough ge sp st.(si_exits) rs0 m0 + else smatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) is.(ipc). -Definition has_aborted (ge: RTL.genv) (sp:val) (st: s_istate) (rs0: regset) (m0: mem): Prop := - has_aborted_on_basic ge sp st.(curr) rs0 m0 - \/ has_aborted_on_test ge sp st.(exits) rs0 m0. +Definition sabort (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := + sabort_local ge sp st.(si_local) rs0 m0 + \/ sabort_exits ge sp st.(si_exits) rs0 m0. -Definition sem_option_istate ge sp (st: s_istate) rs0 m0 (ois: option istate): Prop := +Definition smatch_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => sem_istate ge sp st rs0 m0 is - | None => has_aborted ge sp st rs0 m0 + | Some is => smatch ge sp st rs0 m0 is + | None => sabort ge sp st rs0 m0 end. -Definition sem_option2_istate ge sp (ost: option s_istate) rs0 m0 (ois: option istate) : Prop := +Definition smatch_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => sem_option_istate ge sp st rs0 m0 ois + | Some st => smatch_opt ge sp st rs0 m0 ois | None => ois=None end. @@ -154,10 +155,10 @@ Definition sem_option2_istate ge sp (ost: option s_istate) rs0 m0 (ois: option i *) Definition istate_eq ist1 ist2 := - ist1.(continue) = ist2.(continue) /\ - ist1.(RTLpath.the_pc) = ist2.(RTLpath.the_pc) /\ - (forall r, (ist1.(the_rs)#r) = ist2.(the_rs)#r) /\ - ist1.(the_mem) = ist2.(the_mem). + ist1.(icontinue) = ist2.(icontinue) /\ + ist1.(ipc) = ist2.(ipc) /\ + (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\ + ist1.(imem) = ist2.(imem). (* TODO: is it useful Lemma has_not_yet_exit_cons_split ge sp ext lx rs0 m0: @@ -168,48 +169,48 @@ Proof. Qed. *) -Lemma exclude_early_NYE ge sp st rs0 m0 rs m pc: - sem_early_exit ge sp (exits st) rs0 m0 rs m pc -> - has_not_yet_exit ge sp (exits st) rs0 m0 -> +Lemma all_fallthrough_noexit ge sp st rs0 m0 rs m pc: + smatch_exits ge sp (si_exits st) rs0 m0 rs m pc -> + all_fallthrough ge sp (si_exits st) rs0 m0 -> False. Proof. Local Hint Resolve is_tail_in: core. destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. - unfold has_not_yet_exit; intros NYE; rewrite NYE in EVAL; eauto. + unfold all_fallthrough; intros NYE; rewrite NYE in EVAL; eauto. try congruence. Qed. -Lemma sem_istate_exclude_incompatible_continue ge sp st rs m is1 is2: - is1.(continue) = true -> - is2.(continue) = false -> - sem_istate ge sp st rs m is1 -> - sem_istate ge sp st rs m is2 -> +Lemma smatch_exclude_incompatible_continue ge sp st rs m is1 is2: + is1.(icontinue) = true -> + is2.(icontinue) = false -> + smatch ge sp st rs m is1 -> + smatch ge sp st rs m is2 -> False. Proof. - Local Hint Resolve exclude_early_NYE: core. - unfold sem_istate. + Local Hint Resolve all_fallthrough_noexit: core. + unfold smatch. intros CONT1 CONT2. rewrite CONT1, CONT2; simpl. intuition eauto. Qed. -Lemma sem_istate_determ_continue ge sp st rs m is1 is2: - sem_istate ge sp st rs m is1 -> - sem_istate ge sp st rs m is2 -> - is1.(continue) = is2.(continue). +Lemma smatch_determ_continue ge sp st rs m is1 is2: + smatch ge sp st rs m is1 -> + smatch ge sp st rs m is2 -> + is1.(icontinue) = is2.(icontinue). Proof. - Local Hint Resolve sem_istate_exclude_incompatible_continue: core. - destruct (Bool.bool_dec is1.(continue) is2.(continue)) as [|H]; auto. + Local Hint Resolve smatch_exclude_incompatible_continue: core. + destruct (Bool.bool_dec is1.(icontinue) is2.(icontinue)) as [|H]; auto. intros H1 H2. assert (absurd: False); intuition. - destruct (continue is1) eqn: His1, (continue is2) eqn: His2; eauto. + destruct (icontinue is1) eqn: His1, (icontinue is2) eqn: His2; eauto. Qed. -Lemma sem_local_istate_determ ge sp st rs0 m0 rs1 m1 rs2 m2: - sem_local_istate ge sp st rs0 m0 rs1 m1 -> - sem_local_istate ge sp st rs0 m0 rs2 m2 -> +Lemma smatch_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: + smatch_local ge sp st rs0 m0 rs1 m1 -> + smatch_local ge sp st rs0 m0 rs2 m2 -> (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. - unfold sem_local_istate. intuition try congruence. + unfold smatch_local. intuition try congruence. generalize (H5 r); rewrite H4; congruence. Qed. @@ -226,23 +227,23 @@ 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) -> - has_not_yet_exit ge sp lx1 rs0 m0 -> - seval_condition ge sp (the_cond ext1) (cond_args ext1) (the_smem (exit_ist ext1)) rs0 m0 = Some true -> - has_not_yet_exit ge sp lx2 rs0 m0 -> + 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 has_not_yet_exit in NYE. + unfold all_fallthrough in NYE. rewrite NYE in EVAL; eauto. try congruence. Qed. -Lemma sem_early_exit_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : - sem_early_exit ge sp lx rs0 m0 rs1 m1 pc1 -> - sem_early_exit ge sp lx rs0 m0 rs2 m2 pc2 -> +Lemma smatch_exits_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : + smatch_exits ge sp lx rs0 m0 rs1 m1 pc1 -> + smatch_exits ge sp lx 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. @@ -250,31 +251,31 @@ Proof. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. assert (X:ext1=ext2). - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. - - subst. destruct (sem_local_istate_determ ge sp (exit_ist ext2) rs0 m0 rs1 m1 rs2 m2); auto. + - subst. destruct (smatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. -Lemma sem_istate_determ ge sp st rs m is1 is2: - sem_istate ge sp st rs m is1 -> - sem_istate ge sp st rs m is2 -> +Lemma smatch_determ ge sp st rs m is1 is2: + smatch ge sp st rs m is1 -> + smatch ge sp st rs m is2 -> istate_eq is1 is2. Proof. unfold istate_eq. intros SEM1 SEM2. - exploit (sem_istate_determ_continue ge sp st rs m is1 is2); eauto. - intros CONTEQ. unfold sem_istate in * |-. rewrite CONTEQ in * |- *. - destruct (continue is2). - - destruct (sem_local_istate_determ ge sp (curr st) rs m (the_rs is1) (the_mem is1) (the_rs is2) (the_mem is2)); + exploit (smatch_determ_continue ge sp st rs m is1 is2); eauto. + intros CONTEQ. unfold smatch in * |-. rewrite CONTEQ in * |- *. + destruct (icontinue is2). + - destruct (smatch_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); intuition (try congruence). - - destruct (sem_early_exit_determ ge sp (exits st) rs m (the_rs is1) (the_mem is1) (RTLpath.the_pc is1) (the_rs is2) (the_mem is2) (RTLpath.the_pc is2)); + - destruct (smatch_exits_determ ge sp (si_exits st) rs m (irs is1) (imem is1) (ipc is1) (irs is2) (imem is2) (ipc is2)); intuition. Qed. -Lemma sem_istate_exclude_abort_basic_continue ge sp st rs m is: - continue is = true -> - has_aborted_on_basic ge sp (curr st) rs m -> - sem_istate ge sp st rs m is -> False. +Lemma smatch_exclude_sabort_local_continue ge sp st rs m is: + icontinue is = true -> + sabort_local ge sp (si_local st) rs m -> + smatch ge sp st rs m is -> False. Proof. - intros CONT ABORT SEM. unfold sem_istate in SEM. rewrite CONT in SEM. + intros CONT ABORT SEM. unfold smatch in SEM. rewrite CONT in SEM. destruct SEM as (SEML & _ & _). inversion SEML as (SEML1 & SEML2 & SEML3); clear SEML. inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. - contradiction. @@ -282,10 +283,10 @@ Proof. - inv ABORT3. rewrite SEML3 in H. discriminate. Qed. -Lemma sem_istate_exclude_abort_basic_nocontinue ge sp st rs0 m0 is: - continue is = false -> - has_aborted_on_basic ge sp (curr st) rs0 m0 -> - sem_istate ge sp st rs0 m0 is -> False. +Lemma smatch_exclude_sabort_local_nocontinue ge sp st rs0 m0 is: + icontinue is = false -> + sabort_local ge sp (si_local st) rs0 m0 -> + smatch ge sp st rs0 m0 is -> False. Proof. (* intros CONT ABORT SEM. unfold sem_istate in SEM. rewrite CONT in SEM. destruct st as [pc exits iis]. simpl in *. @@ -298,99 +299,91 @@ Proof. Admitted. -Lemma sem_istate_exclude_abort_basic ge sp st rs m is: - has_aborted_on_basic ge sp (curr st) rs m -> - sem_istate ge sp st rs m is -> False. +Lemma smatch_exclude_sabort_local ge sp st rs m is: + sabort_local ge sp (si_local st) rs m -> + smatch ge sp st rs m is -> False. Proof. - intros. destruct (continue is) eqn:CONT. - - eapply sem_istate_exclude_abort_basic_continue; eauto. - - eapply sem_istate_exclude_abort_basic_nocontinue; eauto. + intros. destruct (icontinue is) eqn:CONT. + - eapply smatch_exclude_sabort_local_continue; eauto. + - eapply smatch_exclude_sabort_local_nocontinue; eauto. Qed. -Lemma sem_istate_exclude_abort_test ge sp st rs m is: - has_aborted_on_test ge sp (exits st) rs m -> - sem_istate ge sp st rs m is -> False. +Lemma smatch_exclude_sabort_exits ge sp st rs m is: + sabort_exits ge sp (si_exits st) rs m -> + smatch ge sp st rs m is -> False. Proof. Admitted. - Lemma sem_istate_exclude_abort ge sp st rs m is: - has_aborted ge sp st rs m -> - sem_istate ge sp st rs m is -> False. - Proof. +Lemma smatch_exclude_sabort ge sp st rs m is: + sabort ge sp st rs m -> + smatch ge sp st rs m is -> False. +Proof. intros ABORT SEM. inv ABORT. - - eapply sem_istate_exclude_abort_basic; eauto. - - eapply sem_istate_exclude_abort_test; eauto. + - eapply smatch_exclude_sabort_local; eauto. + - eapply smatch_exclude_sabort_exits; eauto. Qed. - -(* - intros [[AB|[AB|AB]]|AB]; - ( unfold sem_istate, sem_local_istate; destruct (continue is) eqn: CONT; - [ intros (PRE & REG & MEM) | intros [ext lx' TAIL NYE COND (PRE®&MEM) PC] ]; intuition try congruence). - (* TODO: découper cette grosse preuve à faire en lemmes *) - Admitted. - *) - - -Definition istate_eq_option ist1 oist := +Definition istate_eq_opt ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -Lemma sem_option_istate_determ ge sp st rs m ois is: - sem_option_istate ge sp st rs m ois -> - sem_istate ge sp st rs m is -> - istate_eq_option is ois. +Lemma smatch_opt_determ ge sp st rs m ois is: + smatch_opt ge sp st rs m ois -> + smatch ge sp st rs m is -> + istate_eq_opt is ois. Proof. destruct ois as [is1|]; simpl; eauto. - - intros; eexists; intuition; eapply sem_istate_determ; eauto. - - intros; exploit sem_istate_exclude_abort; eauto. destruct 1. + - intros; eexists; intuition; eapply smatch_determ; eauto. + - intros; exploit smatch_exclude_sabort; eauto. destruct 1. Qed. (** * Symbolic execution of one internal step *) -Definition sreg_set (st:local_istate) (r:reg) (sv:sval) := - {| pre:=(fun ge sp rs m => sval_eval ge sp (st.(the_sreg) r) rs m <> None /\ (st.(pre) ge sp rs m)); - the_sreg:=fun y => if Pos.eq_dec r y then sv else st.(the_sreg) y; - the_smem:= st.(the_smem)|}. +Definition slocal_set_sreg (st:sistate_local) (r:reg) (sv:sval) := + {| si_pre:=(fun ge sp rs m => seval_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 smem_set (st:local_istate) (sm:smem) := - {| pre:=(fun ge sp rs m => smem_eval ge sp st.(the_smem) rs m <> None /\ (st.(pre) ge sp rs m)); - the_sreg:= st.(the_sreg); - the_smem:= sm |}. +Definition slocal_set_smem (st:sistate_local) (sm:smem) := + {| si_pre:=(fun ge sp rs m => seval_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 (st: s_istate) (pc: node) (nxt: local_istate): option s_istate := - Some {| the_pc := pc; exits := st.(exits); curr:= nxt |}. +Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option sistate := + Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. (* FIXME - add notrap *) -Definition symb_istep (i: instruction) (st: s_istate): option s_istate := +Definition symb_istep (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => - sist st pc' st.(curr) + sist_set_local st pc' st.(si_local) | Iop op args dst pc' => - let prev := st.(curr) in - let vargs := list_sval_inj (List.map prev.(the_sreg) args) in - let next := sreg_set prev dst (Sop op vargs prev.(the_smem)) in - sist st pc' next + 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 + sist_set_local st pc' next | Iload _ chunk addr args dst pc' => - let prev := st.(curr) in - let vargs := list_sval_inj (List.map prev.(the_sreg) args) in - let next := sreg_set prev dst (Sload prev.(the_smem) chunk addr vargs) in - sist st pc' next + 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) chunk addr vargs) in + sist_set_local st pc' next | Istore chunk addr args src pc' => - let prev := st.(curr) in - let vargs := list_sval_inj (List.map prev.(the_sreg) args) in - let next := smem_set prev (Sstore prev.(the_smem) chunk addr vargs (prev.(the_sreg) src)) in - sist st pc' next + let prev := st.(si_local) in + let vargs := list_sval_inj (List.map prev.(si_sreg) args) in + let next := slocal_set_smem prev (Sstore prev.(si_smem) chunk addr vargs (prev.(si_sreg) src)) in + sist_set_local st pc' next | Icond cond args ifso ifnot _ => - let prev := st.(curr) in - let vargs := list_sval_inj (List.map prev.(the_sreg) args) in - let ex := {| the_cond:=cond; cond_args:=vargs; exit_ist := prev; cond_ifso := ifso |} in - Some {| the_pc := ifnot; exits := ex::st.(exits); curr := prev |} + 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 (* TODO jumptable ? *) end. +(* TODO TODO - continue renaming *) + Lemma list_sval_eval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: - (forall r : reg, sval_eval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> + (forall r : reg, seval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> list_sval_eval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; auto. @@ -553,7 +546,7 @@ Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := end. Lemma symb_isteps_correct_false ge sp path f rs0 m0 st' is: - is.(continue)=false -> + is.(icontinue)=false -> forall st, sem_istate ge sp st rs0 m0 is -> symb_isteps path f st = Some st' -> sem_istate ge sp st' rs0 m0 is. @@ -599,7 +592,7 @@ Proof. Qed. Lemma symb_isteps_correct_true ge sp path (f:function) rs0 m0: forall st is, - is.(continue)=true -> + is.(icontinue)=true -> sem_istate ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(the_pc) <> None -> sem_option2_istate ge sp (symb_isteps path f st) rs0 m0 @@ -619,7 +612,7 @@ Proof. exploit symb_istep_correct; eauto. inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - inversion_SOME is1; intros His1;rewrite His1; simpl. - * destruct (continue is1) eqn:CONT1. + * destruct (icontinue is1) eqn:CONT1. (* continue is0 = true *) intros; eapply IHpath; eauto. destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. @@ -677,7 +670,7 @@ Inductive sfval := Definition s_find_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := match svos with - | inl sv => SOME v <- sval_eval ge sp sv rs0 m0 IN Genv.find_funct pge v + | inl sv => SOME v <- seval ge sp sv rs0 m0 IN Genv.find_funct pge v | inr symb => SOME b <- Genv.find_symbol pge symb IN Genv.find_funct_ptr pge b end. @@ -714,7 +707,7 @@ Inductive sfval_sem (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: s_istate) s | 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 => sval_eval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> + match osv with Some sv => seval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> sfval_sem pge ge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . @@ -723,11 +716,11 @@ Record s_state := { internal:> s_istate; final: sfval }. Inductive sem_state pge (ge: RTL.genv) (sp:val) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | sem_early is: - is.(continue) = false -> + is.(icontinue) = false -> sem_istate ge sp st rs0 m0 is -> sem_state pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) | sem_normal is t s: - is.(continue) = true -> + is.(icontinue) = true -> sem_istate ge sp st rs0 m0 is -> sfval_sem pge ge sp st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> sem_state pge ge sp st stack f rs0 m0 t s @@ -871,7 +864,7 @@ Proof. exploit symb_istep_correct; eauto. simpl. erewrite Hst', ISTEP; simpl. clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i. - intro SEM; destruct (ist.(continue)) eqn: CONT. + intro SEM; destruct (ist.(icontinue)) eqn: CONT. { (* continue ist = true *) eapply sem_normal; simpl; eauto. unfold sem_istate in SEM. @@ -1026,8 +1019,8 @@ Variable ge ge': RTL.genv. Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. -Lemma sval_eval_preserved sp sv rs0 m0: - sval_eval ge sp sv rs0 m0 = sval_eval ge' sp sv rs0 m0. +Lemma seval_preserved sp sv rs0 m0: + seval ge sp sv rs0 m0 = seval ge' sp sv rs0 m0. Proof. Local Hint Resolve symbols_preserved_RTL: core. induction sv using sval_mut with (P0 := fun lsv => list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0) @@ -1039,7 +1032,7 @@ Proof. erewrite <- eval_addressing_preserved; eauto. destruct (eval_addressing _ sp _ _); auto. rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (sval_eval _ _ _ _); auto. + + rewrite IHsv; clear IHsv. destruct (seval _ _ _ _); auto. rewrite IHsv0; auto. + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. erewrite <- eval_addressing_preserved; eauto. @@ -1052,7 +1045,7 @@ Lemma list_sval_eval_preserved sp lsv rs0 m0: list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0. Proof. induction lsv; simpl; auto. - rewrite sval_eval_preserved. destruct (sval_eval _ _ _ _); auto. + rewrite seval_preserved. destruct (seval _ _ _ _); auto. rewrite IHlsv; auto. Qed. @@ -1064,7 +1057,7 @@ Proof. erewrite <- eval_addressing_preserved; eauto. destruct (eval_addressing _ sp _ _); auto. rewrite IHsm; clear IHsm. destruct (smem_eval _ _ _ _); auto. - rewrite sval_eval_preserved; auto. + rewrite seval_preserved; auto. Qed. Lemma seval_condition_preserved sp cond lsv sm rs0 m0: @@ -1080,12 +1073,12 @@ End SymbValPreserved. Require Import RTLpathLivegen RTLpathLiveproofs. Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := - is1.(continue) = is2.(continue) + is1.(icontinue) = is2.(icontinue) /\ eqlive_reg alive is1.(the_rs) is2.(the_rs) /\ is1.(the_mem) = is2.(the_mem). Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := - if is1.(continue) then + if is1.(icontinue) then (* TODO: il faudra raffiner le (fun _ => True) si on veut autoriser l'oracle à ajouter du "code mort" sur des registres non utilisés (loop invariant code motion à la David) Typiquement, pour connaître la frame des registres vivants, il faudra faire une propagation en arrière -- cgit From fc113f5e7fc1a7d4baac9ea7f0d63ec2c2bdcb90 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 23 Apr 2020 12:15:51 +0200 Subject: Finished renaming --- mppa_k1c/lib/RTLpathSE_theory.v | 606 ++++++++++++++++++++-------------------- mppa_k1c/lib/RTLpathScheduler.v | 70 ++--- 2 files changed, 339 insertions(+), 337 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 9c0c0c0a..e38e66dd 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -83,7 +83,7 @@ with seval_smem (ge: RTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): opti (* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) Record sistate_local := { si_pre: RTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }. -Definition smatch_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := +Definition simatch_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := st.(si_pre) ge sp rs0 m0 /\ seval_smem ge sp st.(si_smem) rs0 m0 = Some m /\ forall (r:reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r). @@ -111,46 +111,46 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. -Inductive smatch_exits (ge: RTL.genv) (sp:val) (lx:list sistate_exit) (rs0: regset) (m0: mem) rs m pc: Prop := - smatch_exits_intro ext lx': +Inductive simatch_exits (ge: RTL.genv) (sp:val) (lx:list sistate_exit) (rs0: regset) (m0: mem) rs m pc: Prop := + simatch_exits_intro ext lx': is_tail (ext::lx') lx -> all_fallthrough ge sp lx' rs0 m0 -> seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some true -> - smatch_local ge sp ext.(si_elocal) rs0 m0 rs m -> + simatch_local ge sp ext.(si_elocal) rs0 m0 rs m -> ext.(si_ifso) = pc -> - smatch_exits ge sp lx rs0 m0 rs m pc. + simatch_exits ge sp lx rs0 m0 rs m pc. (** * Syntax and Semantics of symbolic internal state *) Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. -Definition smatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := +Definition simatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(icontinue)) then - smatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) + simatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) /\ st.(si_pc) = is.(ipc) /\ all_fallthrough ge sp st.(si_exits) rs0 m0 - else smatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) is.(ipc). + else simatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) is.(ipc). Definition sabort (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := sabort_local ge sp st.(si_local) rs0 m0 \/ sabort_exits ge sp st.(si_exits) rs0 m0. -Definition smatch_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := +Definition simatch_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => smatch ge sp st rs0 m0 is + | Some is => simatch ge sp st rs0 m0 is | None => sabort ge sp st rs0 m0 end. -Definition smatch_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := +Definition simatch_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => smatch_opt ge sp st rs0 m0 ois + | Some st => simatch_opt ge sp st rs0 m0 ois | None => ois=None end. (** * An internal state represents a parallel program ! - We prove below that the semantics [sem_option_istate] is deterministic. + We prove below that the semantics [simatch_opt] is deterministic. *) @@ -162,15 +162,15 @@ Definition istate_eq ist1 ist2 := (* TODO: is it useful Lemma has_not_yet_exit_cons_split ge sp ext lx rs0 m0: - has_not_yet_exit ge sp (ext::lx) rs0 m0 <-> - (seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(the_smem) rs0 m0 = Some false /\ has_not_yet_exit ge sp lx rs0 m0). + all_fallthrough ge sp (ext::lx) rs0 m0 <-> + (seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(si_smem) rs0 m0 = Some false /\ all_fallthrough ge sp lx rs0 m0). Proof. - unfold has_not_yet_exit; simpl; intuition (subst; auto). + unfold all_fallthrough; simpl; intuition (subst; auto). Qed. *) Lemma all_fallthrough_noexit ge sp st rs0 m0 rs m pc: - smatch_exits ge sp (si_exits st) rs0 m0 rs m pc -> + simatch_exits ge sp (si_exits st) rs0 m0 rs m pc -> all_fallthrough ge sp (si_exits st) rs0 m0 -> False. Proof. @@ -180,37 +180,37 @@ Proof. try congruence. Qed. -Lemma smatch_exclude_incompatible_continue ge sp st rs m is1 is2: +Lemma simatch_exclude_incompatible_continue ge sp st rs m is1 is2: is1.(icontinue) = true -> is2.(icontinue) = false -> - smatch ge sp st rs m is1 -> - smatch ge sp st rs m is2 -> + simatch ge sp st rs m is1 -> + simatch ge sp st rs m is2 -> False. Proof. Local Hint Resolve all_fallthrough_noexit: core. - unfold smatch. + unfold simatch. intros CONT1 CONT2. rewrite CONT1, CONT2; simpl. intuition eauto. Qed. -Lemma smatch_determ_continue ge sp st rs m is1 is2: - smatch ge sp st rs m is1 -> - smatch ge sp st rs m is2 -> +Lemma simatch_determ_continue ge sp st rs m is1 is2: + simatch ge sp st rs m is1 -> + simatch ge sp st rs m is2 -> is1.(icontinue) = is2.(icontinue). Proof. - Local Hint Resolve smatch_exclude_incompatible_continue: core. + Local Hint Resolve simatch_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 smatch_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: - smatch_local ge sp st rs0 m0 rs1 m1 -> - smatch_local ge sp st rs0 m0 rs2 m2 -> +Lemma simatch_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: + simatch_local ge sp st rs0 m0 rs1 m1 -> + simatch_local ge sp st rs0 m0 rs2 m2 -> (forall r, rs1#r = rs2#r) /\ m1 = m2. Proof. - unfold smatch_local. intuition try congruence. + unfold simatch_local. intuition try congruence. generalize (H5 r); rewrite H4; congruence. Qed. @@ -241,9 +241,9 @@ Proof. try congruence. Qed. -Lemma smatch_exits_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : - smatch_exits ge sp lx rs0 m0 rs1 m1 pc1 -> - smatch_exits ge sp lx rs0 m0 rs2 m2 pc2 -> +Lemma simatch_exits_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : + simatch_exits ge sp lx rs0 m0 rs1 m1 pc1 -> + simatch_exits ge sp lx 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. @@ -251,31 +251,31 @@ Proof. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. assert (X:ext1=ext2). - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. - - subst. destruct (smatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. + - subst. destruct (simatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. -Lemma smatch_determ ge sp st rs m is1 is2: - smatch ge sp st rs m is1 -> - smatch ge sp st rs m is2 -> +Lemma simatch_determ ge sp st rs m is1 is2: + simatch ge sp st rs m is1 -> + simatch ge sp st rs m is2 -> istate_eq is1 is2. Proof. unfold istate_eq. intros SEM1 SEM2. - exploit (smatch_determ_continue ge sp st rs m is1 is2); eauto. - intros CONTEQ. unfold smatch in * |-. rewrite CONTEQ in * |- *. + exploit (simatch_determ_continue ge sp st rs m is1 is2); eauto. + intros CONTEQ. unfold simatch in * |-. rewrite CONTEQ in * |- *. destruct (icontinue is2). - - destruct (smatch_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); + - destruct (simatch_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); intuition (try congruence). - - destruct (smatch_exits_determ ge sp (si_exits st) rs m (irs is1) (imem is1) (ipc is1) (irs is2) (imem is2) (ipc is2)); + - destruct (simatch_exits_determ ge sp (si_exits st) rs m (irs is1) (imem is1) (ipc is1) (irs is2) (imem is2) (ipc is2)); intuition. Qed. -Lemma smatch_exclude_sabort_local_continue ge sp st rs m is: +Lemma simatch_exclude_sabort_local_continue ge sp st rs m is: icontinue is = true -> sabort_local ge sp (si_local st) rs m -> - smatch ge sp st rs m is -> False. + simatch ge sp st rs m is -> False. Proof. - intros CONT ABORT SEM. unfold smatch in SEM. rewrite CONT in SEM. + intros CONT ABORT SEM. unfold simatch in SEM. rewrite CONT in SEM. destruct SEM as (SEML & _ & _). inversion SEML as (SEML1 & SEML2 & SEML3); clear SEML. inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. - contradiction. @@ -283,13 +283,13 @@ Proof. - inv ABORT3. rewrite SEML3 in H. discriminate. Qed. -Lemma smatch_exclude_sabort_local_nocontinue ge sp st rs0 m0 is: +Lemma simatch_exclude_sabort_local_nocontinue ge sp st rs0 m0 is: icontinue is = false -> sabort_local ge sp (si_local st) rs0 m0 -> - smatch ge sp st rs0 m0 is -> False. + simatch ge sp st rs0 m0 is -> False. Proof. -(* intros CONT ABORT SEM. unfold sem_istate in SEM. rewrite CONT in SEM. - destruct st as [pc exits iis]. simpl in *. +(* intros CONT ABORT SEM. unfold simatch in SEM. rewrite CONT in SEM. + destruct st as [pc si_exits iis]. simpl in *. destruct is as [cont pc' rs' m']. simpl in *. inv SEM. inversion H2 as (SEML1 & SEML2 & SEML3); clear H2. inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. @@ -299,42 +299,42 @@ Proof. Admitted. -Lemma smatch_exclude_sabort_local ge sp st rs m is: +Lemma simatch_exclude_sabort_local ge sp st rs m is: sabort_local ge sp (si_local st) rs m -> - smatch ge sp st rs m is -> False. + simatch ge sp st rs m is -> False. Proof. intros. destruct (icontinue is) eqn:CONT. - - eapply smatch_exclude_sabort_local_continue; eauto. - - eapply smatch_exclude_sabort_local_nocontinue; eauto. + - eapply simatch_exclude_sabort_local_continue; eauto. + - eapply simatch_exclude_sabort_local_nocontinue; eauto. Qed. -Lemma smatch_exclude_sabort_exits ge sp st rs m is: +Lemma simatch_exclude_sabort_exits ge sp st rs m is: sabort_exits ge sp (si_exits st) rs m -> - smatch ge sp st rs m is -> False. + simatch ge sp st rs m is -> False. Proof. Admitted. -Lemma smatch_exclude_sabort ge sp st rs m is: +Lemma simatch_exclude_sabort ge sp st rs m is: sabort ge sp st rs m -> - smatch ge sp st rs m is -> False. + simatch ge sp st rs m is -> False. Proof. intros ABORT SEM. inv ABORT. - - eapply smatch_exclude_sabort_local; eauto. - - eapply smatch_exclude_sabort_exits; eauto. + - eapply simatch_exclude_sabort_local; eauto. + - eapply simatch_exclude_sabort_exits; eauto. Qed. Definition istate_eq_opt ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -Lemma smatch_opt_determ ge sp st rs m ois is: - smatch_opt ge sp st rs m ois -> - smatch ge sp st rs m is -> +Lemma simatch_opt_determ ge sp st rs m ois is: + simatch_opt ge sp st rs m ois -> + simatch ge sp st rs m is -> istate_eq_opt is ois. Proof. destruct ois as [is1|]; simpl; eauto. - - intros; eexists; intuition; eapply smatch_determ; eauto. - - intros; exploit smatch_exclude_sabort; eauto. destruct 1. + - intros; eexists; intuition; eapply simatch_determ; eauto. + - intros; exploit simatch_exclude_sabort; eauto. destruct 1. Qed. (** * Symbolic execution of one internal step *) @@ -353,7 +353,7 @@ Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. (* FIXME - add notrap *) -Definition symb_istep (i: instruction) (st: sistate): option sistate := +Definition sistep (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => sist_set_local st pc' st.(si_local) @@ -380,11 +380,10 @@ Definition symb_istep (i: instruction) (st: sistate): option sistate := | _ => None (* TODO jumptable ? *) end. -(* TODO TODO - continue renaming *) -Lemma list_sval_eval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: - (forall r : reg, seval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> - list_sval_eval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). +Lemma seval_list_sval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs: + (forall r : reg, seval_sval ge sp (sreg r) rs0 m0 = Some (rs # r)) -> + seval_list_sval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l). Proof. intros H; induction l as [|r l]; simpl; auto. inversion_SOME v. @@ -393,242 +392,242 @@ Proof. try_simplify_someHyps. Qed. -Lemma symb_istep_preserv_early_exits i ge sp st rs0 m0 rs m pc st': - sem_early_exit ge sp st.(exits) rs0 m0 rs m pc -> - symb_istep i st = Some st' -> - sem_early_exit ge sp st'.(exits) rs0 m0 rs m pc. +Lemma sistep_preserves_simatch_exits i ge sp st rs0 m0 rs m pc st': + simatch_exits ge sp st.(si_exits) rs0 m0 rs m pc -> + sistep i st = Some st' -> + simatch_exits ge sp st'.(si_exits) rs0 m0 rs m pc. Proof. intros H. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps. destruct H as [ext lx TAIL NYE COND INV PC]. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. -Lemma sreg_set_preserv_has_aborted ge sp st rs0 m0 r sv: - has_aborted_on_basic ge sp st rs0 m0 -> - has_aborted_on_basic ge sp (sreg_set st r sv) rs0 m0. +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 has_aborted_on_basic. simpl; intuition. + 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 smem_set_preserv_has_aborted ge sp st rs0 m0 m: - has_aborted_on_basic ge sp st rs0 m0 -> - has_aborted_on_basic ge sp (smem_set st m) rs0 m0. +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 has_aborted_on_basic. simpl; intuition. + unfold sabort_local. simpl; intuition. Qed. -Lemma symb_istep_preserv_has_aborted i ge sp rs0 m0 st st': - symb_istep i st = Some st' -> - has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. +Lemma sistep_preserves_sabort i ge sp rs0 m0 st st': + sistep i st = Some st' -> + sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. Proof. - Local Hint Resolve sreg_set_preserv_has_aborted smem_set_preserv_has_aborted: core. - unfold has_aborted. - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps; intuition eauto. + Local Hint Resolve slocal_set_sreg_preserves_sabort_local slocal_set_smem_preserves_sabort_local: core. + unfold sabort. + destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps; intuition eauto. (* COND *) - right. unfold has_aborted_on_test in * |- *. + right. unfold sabort_exits in * |- *. destruct H as (ext & H1 & H2). simpl; exists ext; intuition. Qed. -Lemma symb_istep_WF i st: - symb_istep i st = None -> default_succ i = None. +Lemma sistep_WF i st: + sistep i st = None -> default_succ i = None. Proof. - destruct i; simpl; unfold sist; simpl; congruence. + destruct i; simpl; unfold sist_set_local; simpl; congruence. Qed. -Lemma symb_istep_default_succ i st st': - symb_istep i st = Some st' -> default_succ i = Some (st'.(the_pc)). +Lemma sistep_default_succ i st st': + sistep i st = Some st' -> default_succ i = Some (st'.(si_pc)). Proof. - destruct i; simpl; unfold sist; simpl; try congruence; + destruct i; simpl; unfold sist_set_local; simpl; try congruence; intro H; inversion_clear H; simpl; auto. Qed. -Lemma symb_istep_correct ge sp i st rs0 m0 rs m: - sem_local_istate ge sp st.(curr) rs0 m0 rs m -> - has_not_yet_exit ge sp st.(exits) rs0 m0 -> - sem_option2_istate ge sp (symb_istep i st) rs0 m0 (istep ge i sp rs m). +Lemma sistep_correct ge sp i st rs0 m0 rs m: + simatch_local ge sp st.(si_local) rs0 m0 rs m -> + all_fallthrough ge sp st.(si_exits) rs0 m0 -> + simatch_opt2 ge sp (sistep i st) rs0 m0 (istep ge i sp rs m). Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. + (* Nop *) - unfold sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + unfold simatch, simatch_local; simpl; try_simplify_someHyps. + (* Op *) inversion_SOME v; intros OP; simpl. - - unfold sem_istate, sem_local_istate; simpl; intuition. + - unfold simatch, simatch_local; simpl; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl; auto. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - - unfold has_aborted, has_aborted_on_basic; simpl. + - unfold sabort, sabort_local; simpl. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite list_sval_eval_inj; simpl; auto. + simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. + (* LOAD *) admit. (* FIXME *) (* inversion_SOME a0; intros ADD. { inversion_SOME v; intros LOAD; simpl. - - explore_destruct; unfold sem_istate, sem_local_istate; simpl; intuition. + - explore_destruct; unfold simatch, simatch_local; simpl; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - - unfold has_aborted, has_aborted_on_basic; simpl. + - unfold sabort, sabort_local; simpl. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite list_sval_eval_inj; simpl; auto. + simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } - { unfold has_aborted, has_aborted_on_basic; simpl. + { unfold sabort, sabort_local; simpl. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite list_sval_eval_inj; simpl; auto. + simpl. erewrite seval_list_sval_inj; simpl; auto. rewrite ADD; simpl; auto. } *) + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. - - unfold sem_istate, sem_local_istate; simpl; intuition. + - unfold simatch, simatch_local; simpl; intuition. * congruence. - * erewrite list_sval_eval_inj; simpl; auto. + * erewrite seval_list_sval_inj; simpl; auto. erewrite REG. try_simplify_someHyps. - - unfold has_aborted, has_aborted_on_basic; simpl. + - unfold sabort, sabort_local; simpl. left. right. left. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. erewrite REG. try_simplify_someHyps. } - { unfold has_aborted, has_aborted_on_basic; simpl. + { unfold sabort, sabort_local; simpl. left. right. left. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. erewrite ADD; simpl; auto. } + (* COND *) Local Hint Resolve is_tail_refl: core. - Local Hint Unfold sem_local_istate: core. + Local Hint Unfold simatch_local: core. inversion_SOME b; intros COND. - { destruct b; simpl; unfold sem_istate, sem_local_istate; simpl. + { destruct b; simpl; unfold simatch, simatch_local; simpl. - intros; econstructor; eauto; simpl. unfold seval_condition. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - - intuition. unfold has_not_yet_exit in * |- *. simpl. + - intuition. unfold all_fallthrough in * |- *. simpl. intuition. subst. simpl. unfold seval_condition. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } - { unfold has_aborted, has_aborted_on_test; simpl. + { unfold sabort, sabort_exits; simpl. right. eexists; intuition eauto. simpl. unfold seval_condition. - erewrite list_sval_eval_inj; simpl; auto. + erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } Admitted. -Lemma symb_istep_correct_None ge sp i st rs0 m0 rs m: - sem_local_istate ge sp (st.(curr)) rs0 m0 rs m -> - symb_istep i st = None -> +Lemma sistep_correct_None ge sp i st rs0 m0 rs m: + simatch_local ge sp (st.(si_local)) rs0 m0 rs m -> + sistep i st = None -> istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). - destruct i; simpl; unfold sist, sem_istate, sem_local_istate; simpl; try_simplify_someHyps. + destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps. Qed. (** * Symbolic execution of the internal steps of a path *) -Fixpoint symb_isteps (path:nat) (f: function) (st: s_istate): option s_istate := +Fixpoint sisteps (path:nat) (f: function) (st: sistate): option sistate := match path with | O => Some st | S p => - SOME i <- (fn_code f)!(st.(the_pc)) IN - SOME st1 <- symb_istep i st IN - symb_isteps p f st1 + SOME i <- (fn_code f)!(st.(si_pc)) IN + SOME st1 <- sistep i st IN + sisteps p f st1 end. -Lemma symb_isteps_correct_false ge sp path f rs0 m0 st' is: +Lemma sisteps_correct_false ge sp path f rs0 m0 st' is: is.(icontinue)=false -> - forall st, sem_istate ge sp st rs0 m0 is -> - symb_isteps path f st = Some st' -> - sem_istate ge sp st' rs0 m0 is. + forall st, simatch ge sp st rs0 m0 is -> + sisteps path f st = Some st' -> + simatch ge sp st' rs0 m0 is. Proof. - Local Hint Resolve symb_istep_preserv_early_exits: core. - intros CONT; unfold sem_istate; rewrite CONT. + Local Hint Resolve sistep_preserves_simatch_exits: core. + intros CONT; unfold simatch; rewrite CONT. induction path; simpl. - + unfold sist; try_simplify_someHyps. + + unfold sist_set_local; try_simplify_someHyps. + intros st; inversion_SOME i. inversion_SOME st1; eauto. Qed. -Lemma symb_isteps_preserv_has_aborted ge sp path f rs0 m0 st': forall st, - symb_isteps path f st = Some st' -> - has_aborted ge sp st rs0 m0 -> has_aborted ge sp st' rs0 m0. +Lemma sisteps_preserves_sabort ge sp path f rs0 m0 st': forall st, + sisteps path f st = Some st' -> + sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. Proof. - Local Hint Resolve symb_istep_preserv_has_aborted: core. + Local Hint Resolve sistep_preserves_sabort: core. induction path; simpl. - + unfold sist; try_simplify_someHyps. + + unfold sist_set_local; try_simplify_someHyps. + intros st; inversion_SOME i. inversion_SOME st1; eauto. Qed. -Lemma symb_isteps_WF path f: forall st, - symb_isteps path f st = None -> nth_default_succ (fn_code f) path st.(the_pc) = None. +Lemma sisteps_WF path f: forall st, + sisteps path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None. Proof. induction path; simpl. - + unfold sist. intuition congruence. - + intros st; destruct ((fn_code f) ! (the_pc st)); simpl; try tauto. - destruct (symb_istep i st) as [st1|] eqn: Hst1; simpl. - - intros; erewrite symb_istep_default_succ; eauto. - - intros; erewrite symb_istep_WF; eauto. + + unfold sist_set_local. intuition congruence. + + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try tauto. + destruct (sistep i st) as [st1|] eqn: Hst1; simpl. + - intros; erewrite sistep_default_succ; eauto. + - intros; erewrite sistep_WF; eauto. Qed. -Lemma symb_isteps_default_succ path f st': forall st, - symb_isteps path f st = Some st' -> nth_default_succ (fn_code f) path st.(the_pc) = Some st'.(the_pc). +Lemma sisteps_default_succ path f st': forall st, + sisteps 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. intros st H. inversion_clear H; simpl; try congruence. - + intros st; destruct ((fn_code f) ! (the_pc st)); simpl; try congruence. - destruct (symb_istep i st) as [st1|] eqn: Hst1; simpl; try congruence. - intros; erewrite symb_istep_default_succ; eauto. + + 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 (sistep i st) as [st1|] eqn: Hst1; simpl; try congruence. + intros; erewrite sistep_default_succ; eauto. Qed. -Lemma symb_isteps_correct_true ge sp path (f:function) rs0 m0: forall st is, +Lemma sisteps_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(icontinue)=true -> - sem_istate ge sp st rs0 m0 is -> - nth_default_succ (fn_code f) path st.(the_pc) <> None -> - sem_option2_istate ge sp (symb_isteps path f st) rs0 m0 - (isteps ge path f sp is.(the_rs) is.(the_mem) is.(RTLpath.the_pc)) + simatch ge sp st rs0 m0 is -> + nth_default_succ (fn_code f) path st.(si_pc) <> None -> + simatch_opt2 ge sp (sisteps path f st) rs0 m0 + (isteps ge path f sp is.(irs) is.(imem) is.(ipc)) . Proof. - Local Hint Resolve symb_isteps_correct_false symb_isteps_preserv_has_aborted symb_isteps_WF: core. + Local Hint Resolve sisteps_correct_false sisteps_preserves_sabort sisteps_WF: core. induction path; simpl. + intros st is CONT INV WF; - unfold sem_istate, sist in * |- *; + unfold simatch, sist_set_local in * |- *; try_simplify_someHyps. simpl. destruct is; simpl in * |- *; subst; intuition auto. - + intros st is CONT; unfold sem_istate at 1; rewrite CONT. + + intros st is CONT; unfold simatch at 1; rewrite CONT. intros (LOCAL & PC & NYE) WF. rewrite <- PC. inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. - exploit symb_istep_correct; eauto. + exploit sistep_correct; eauto. inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - inversion_SOME is1; intros His1;rewrite His1; simpl. * destruct (icontinue is1) eqn:CONT1. - (* continue is0 = true *) + (* icontinue is0 = true *) intros; eapply IHpath; eauto. - destruct i; simpl in * |- *; unfold sist in * |- *; try_simplify_someHyps. - (* continue is0 = false -> EARLY EXIT *) - destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. - destruct WF. erewrite symb_istep_default_succ; eauto. + destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps. + (* icontinue is0 = false -> EARLY EXIT *) + destruct (sisteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. + destruct WF. erewrite sistep_default_succ; eauto. (* try_simplify_someHyps; eauto. *) - * destruct (symb_isteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. + * destruct (sisteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. - intros His1;rewrite His1; simpl; auto. Qed. (** REM: in the following two unused lemmas *) -Lemma symb_isteps_right_assoc_decompose f path: forall st st', - symb_isteps (S path) f st = Some st' -> - exists st0, symb_isteps path f st = Some st0 /\ symb_isteps 1%nat f st0 = Some st'. +Lemma sisteps_right_assoc_decompose f path: forall st st', + sisteps (S path) f st = Some st' -> + exists st0, sisteps path f st = Some st0 /\ sisteps 1%nat f st0 = Some st'. Proof. induction path; simpl; eauto. intros st st'. @@ -637,10 +636,10 @@ Proof. try_simplify_someHyps; eauto. Qed. -Lemma symb_isteps_right_assoc_compose f path: forall st st0 st', - symb_isteps path f st = Some st0 -> - symb_isteps 1%nat f st0 = Some st' -> - symb_isteps (S path) f st = Some st'. +Lemma sisteps_right_assoc_compose f path: forall st st0 st', + sisteps path f st = Some st0 -> + sisteps 1%nat f st0 = Some st' -> + sisteps (S path) f st = Some st'. Proof. induction path. + intros st st0 st' H. simpl in H. @@ -668,20 +667,20 @@ Inductive sfval := | Sreturn: option sval -> sfval . -Definition s_find_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := +Definition sfind_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef := match svos with - | inl sv => SOME v <- seval ge sp sv rs0 m0 IN Genv.find_funct pge v + | inl sv => SOME v <- seval_sval ge sp sv rs0 m0 IN Genv.find_funct pge v | inr symb => SOME b <- Genv.find_symbol pge symb IN Genv.find_funct_ptr pge b end. -Inductive sfval_sem (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: s_istate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: - sfval_sem pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(the_pc) rs m) + sfmatch pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(si_pc) rs m) | exec_Scall rs m sig svos lsv args res pc fd: - s_find_function pge ge sp svos rs0 m0 = Some fd -> + sfind_function pge ge sp svos rs0 m0 = Some fd -> funsig fd = sig -> - list_sval_eval ge sp lsv rs0 m0 = Some args -> - sfval_sem pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + seval_list_sval ge sp lsv rs0 m0 = Some args -> + sfmatch pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) (* TODO: | exec_Itailcall stk pc rs m sig ros args fd m': @@ -707,52 +706,52 @@ Inductive sfval_sem (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: s_istate) s | 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 => seval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> - sfval_sem pge ge sp st stack f rs0 m0 (Sreturn osv) rs m + match osv with Some sv => seval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> + sfmatch pge ge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . -Record s_state := { internal:> s_istate; final: sfval }. +Record sstate := { internal:> sistate; final: sfval }. -Inductive sem_state pge (ge: RTL.genv) (sp:val) (st: s_state) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := - | sem_early is: +Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := + | smatch_early is: is.(icontinue) = false -> - sem_istate ge sp st rs0 m0 is -> - sem_state pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(RTLpath.the_pc) is.(the_rs) is.(the_mem)) - | sem_normal is t s: + simatch ge sp st rs0 m0 is -> + smatch pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) + | smatch_normal is t s: is.(icontinue) = true -> - sem_istate ge sp st rs0 m0 is -> - sfval_sem pge ge sp st stack f rs0 m0 st.(final) is.(the_rs) is.(the_mem) t s -> - sem_state pge ge sp st stack f rs0 m0 t s + simatch ge sp st rs0 m0 is -> + sfmatch pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> + smatch pge ge sp st stack f rs0 m0 t s . (** * Symbolic execution of final step *) -Definition symb_fstep (i: instruction) (prev: local_istate): sfval := +Definition sfstep (i: instruction) (prev: sistate_local): sfval := match i with | Icall sig ros args res pc => - let svos := sum_left_map prev.(the_sreg) ros in - let sargs := list_sval_inj (List.map prev.(the_sreg) args) in + let svos := sum_left_map prev.(si_sreg) ros in + let sargs := list_sval_inj (List.map prev.(si_sreg) args) in Scall sig svos sargs res pc | Ireturn or => - let sor := SOME r <- or IN Some (prev.(the_sreg) r) in + let sor := SOME r <- or IN Some (prev.(si_sreg) r) in Sreturn sor | _ => Snone end. -Lemma symb_fstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: +Lemma sfstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> - pc = st.(the_pc) -> - sem_local_istate ge sp (curr st) rs0 m0 rs m -> + pc = st.(si_pc) -> + simatch_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> - symb_istep i st = None -> - sfval_sem pge ge sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s. + sistep i st = None -> + sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. - + (* Snone *) destruct i; simpl in Hi |- *; unfold sist in Hi; try congruence. + + (* Snone *) destruct i; simpl in Hi |- *; unfold sist_set_local in Hi; try congruence. + (* Icall *) intros; eapply exec_Scall; auto. - destruct ros; simpl in * |- *; auto. rewrite REG; auto. - - erewrite list_sval_eval_inj; simpl; auto. + - erewrite seval_list_sval_inj; simpl; auto. + (* Itaillcall *) admit. + (* Ibuiltin *) admit. + (* Ijumptable *) admit. @@ -760,19 +759,19 @@ Proof. destruct or; simpl; auto. Admitted. -Lemma symb_fstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: +Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> - pc = st.(the_pc) -> - sem_local_istate ge sp (curr st) rs0 m0 rs m -> - sfval_sem pge ge sp st stack f rs0 m0 (symb_fstep i (curr st)) rs m t s -> - symb_istep i st = None -> + pc = st.(si_pc) -> + simatch_local ge sp (si_local st) rs0 m0 rs m -> + sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s -> + sistep i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct i as [ | | | |(*Icall*)sig ros args res pc'| | | | |(*Ireturn*)or]; - subst; try_simplify_someHyps; try (unfold sist in Hi; try congruence); inversion LAST; subst; clear LAST; simpl in * |- *. + subst; try_simplify_someHyps; try (unfold sist_set_local in Hi; try congruence); inversion LAST; subst; clear LAST; simpl in * |- *. + (* Icall *) - erewrite list_sval_eval_inj in * |- ; simpl; try_simplify_someHyps; auto. + erewrite seval_list_sval_inj in * |- ; simpl; try_simplify_someHyps; auto. intros; eapply exec_Icall; eauto. destruct ros; simpl in * |- *; auto. rewrite REG in * |- ; auto. @@ -788,22 +787,22 @@ Admitted. (** * Main function of the symbolic execution *) -Definition init_local_istate := {| pre:= fun _ _ _ _ => True; the_sreg:= fun r => Sinput r; the_smem:= Sinit |}. +Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. -Definition init_istate pc := {| the_pc:= pc; exits:=nil; curr:= init_local_istate |}. +Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}. -Lemma init_sem_istate ge sp pc rs m: sem_istate ge sp (init_istate pc) rs m (ist true pc rs m). +Lemma init_simatch ge sp pc rs m: simatch ge sp (init_sistate pc) rs m (mk_istate true pc rs m). Proof. - unfold sem_istate, sem_local_istate, has_not_yet_exit; simpl. intuition. + unfold simatch, simatch_local, all_fallthrough; simpl. intuition. Qed. -Definition symb_step (f: function) (pc:node): option s_state := +Definition sstep (f: function) (pc:node): option sstate := SOME path <- (fn_path f)!pc IN - SOME st <- symb_isteps path.(psize) f (init_istate pc) IN - SOME i <- (fn_code f)!(st.(the_pc)) IN - Some (match symb_istep i st with + SOME st <- sisteps path.(psize) f (init_sistate pc) IN + SOME i <- (fn_code f)!(st.(si_pc)) IN + Some (match sistep i st with | Some st' => {| internal := st'; final := Snone |} - | None => {| internal := st; final := symb_fstep i st.(curr) |} + | None => {| internal := st; final := sfstep i st.(si_local) |} end). Lemma final_node_path_simpl f path pc: @@ -816,67 +815,70 @@ 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.(the_pc) -> - symb_istep i st = Some st' -> + pc = st.(si_pc) -> + sistep i st = Some st' -> path_last_step ge pge stack f sp pc rs m t s -> - exists ist, istep ge i sp rs m = Some ist /\ t = E0 /\ s = (State stack f sp ist.(RTLpath.the_pc) ist.(RTLpath.the_rs) ist.(the_mem)). + 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 [symb_step]) -(symb_step is a correct over-approximation) +(* NB: each concrete execution can be executed on the symbolic state (produced from [sstep]) +(sstep is a correct over-approximation) *) -Theorem symb_step_correct f pc pge ge sp path stack rs m t s: +Theorem sstep_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, symb_step f pc = Some st /\ sem_state pge ge sp st stack f rs m t s. + exists st, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. - Local Hint Resolve init_sem_istate symb_istep_preserv_early_exits: core. - intros PATH STEP; unfold symb_step; rewrite PATH; simpl. + Local Hint Resolve init_simatch sistep_preserves_simatch_exits: core. + intros PATH STEP; unfold sstep; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (sisteps_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 sem_option2_istate; destruct (symb_isteps _ _ _) as [st|] eqn: Hst; try congruence); + (rewrite STEPS; unfold simatch_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); (* intro SEM *) - (simpl; unfold sem_istate; simpl; rewrite CONT; intro SEM); + (simpl; unfold simatch; simpl; rewrite CONT; intro SEM); (* intro Hi' *) - ( assert (Hi': (fn_code f) ! (the_pc st) = Some i); + ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); [ unfold nth_default_succ_inst in Hi; - exploit symb_isteps_default_succ; eauto; simpl; + exploit sisteps_default_succ; eauto; simpl; intros DEF; rewrite DEF in Hi; auto | clear Hi; rewrite Hi' ]); (* eexists *) (eexists; constructor; eauto). - (* early *) - eapply sem_early; eauto. - unfold sem_istate; simpl; rewrite CONT. - destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl; eauto. + eapply smatch_early; eauto. + unfold simatch; simpl; rewrite CONT. + destruct (sistep i st) as [st'|] eqn: Hst'; simpl; eauto. - destruct SEM as (SEM & PC & HNYE). - destruct (symb_istep i st) as [st'|] eqn: Hst'; simpl. + destruct (sistep i st) as [st'|] eqn: Hst'; simpl. + (* normal on Snone *) rewrite <- PC in LAST. exploit symb_path_last_step; eauto; simpl. - intros (ist & ISTEP & Ht & Hs); subst. - exploit symb_istep_correct; eauto. simpl. + intros (mk_istate & ISTEP & Ht & Hs); subst. + exploit sistep_correct; eauto. simpl. erewrite Hst', ISTEP; simpl. clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i. - intro SEM; destruct (ist.(icontinue)) eqn: CONT. - { (* continue ist = true *) - eapply sem_normal; simpl; eauto. - unfold sem_istate in SEM. + intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT. + { (* icontinue mk_istate = true *) + eapply smatch_normal; simpl; eauto. + unfold simatch in SEM. rewrite CONT in SEM. destruct SEM as (SEM & PC & HNYE). rewrite <- PC. eapply exec_Snone. } - { eapply sem_early; eauto. } + { eapply smatch_early; eauto. } + (* normal non-Snone instruction *) - eapply sem_normal; eauto. - * unfold sem_istate; simpl; rewrite CONT; intuition. - * simpl. eapply symb_fstep_correct; eauto. + eapply smatch_normal; eauto. + * unfold simatch; simpl; rewrite CONT; intuition. + * simpl. eapply sfstep_correct; eauto. rewrite PC; auto. Qed. @@ -939,10 +941,10 @@ Proof. Qed. *) -Lemma sfval_sem_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: - sfval_sem pge ge sp st stack f rs0 m0 sv rs1 m t s -> +Lemma sfmatch_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + sfmatch 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' /\ sfval_sem pge ge sp st stack f rs0 m0 sv rs2 m t s'. + exists s', equiv_state s s' /\ sfmatch pge ge sp st stack f rs0 m0 sv rs2 m t s'. Proof. Local Hint Resolve equiv_stack_refl: core. destruct 1. @@ -958,57 +960,57 @@ Proof. + eapply exec_Sreturn; eauto. Qed. -(* NB: each execution of a symbolic state (produced from [symb_step]) represents a concrete execution - (symb_step is exact). +(* NB: each execution of a symbolic state (produced from [sstep]) represents a concrete execution + (sstep is exact). *) -Theorem symb_step_exact f pc pge ge sp path stack st rs m t s1: +Theorem sstep_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> - symb_step f pc = Some st -> - sem_state pge ge sp st stack f rs m t s1 -> + sstep f pc = Some st -> + smatch 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_sem_istate: core. - unfold symb_step; intros PATH SSTEP SEM; rewrite PATH in SSTEP. + Local Hint Resolve init_simatch: core. + unfold sstep; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (symb_isteps_correct_true ge sp path.(psize) f rs m (init_istate pc) (ist true pc rs m)); simpl; eauto. + exploit (sisteps_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 (symb_isteps path.(psize) f (init_istate pc)) as [st0|] eqn: Hst0; simpl. + destruct (sisteps path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl. 2:{ (* absurd case *) - exploit symb_isteps_WF; eauto. + exploit sisteps_WF; eauto. simpl; intros NDS; rewrite NDS in Hi; congruence. } - exploit symb_isteps_default_succ; eauto; simpl. + exploit sisteps_default_succ; eauto; simpl. intros NDS; rewrite NDS in Hi. rewrite Hi in SSTEP. intros ISTEPS. try_simplify_someHyps. - destruct (symb_istep i st0) as [st'|] eqn: Hst'; simpl. + destruct (sistep i st0) as [st'|] eqn: Hst'; simpl. + (* normal exit on Snone instruction *) admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. - exploit sem_option_istate_determ; eauto. + exploit simatch_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 sem_option_istate_determ; eauto. + exploit simatch_opt_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - unfold sem_istate in SEM1. + unfold simatch in SEM1. rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). - exploit sfval_sem_equiv; eauto. + exploit sfmatch_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 symb_fstep_complete; eauto. + eapply sfstep_complete; eauto. * congruence. - * unfold sem_local_istate in * |- *. intuition try congruence. + * unfold simatch_local in * |- *. intuition try congruence. Admitted. (** * Simulation of RTLpath code w.r.t symbolic execution *) @@ -1020,43 +1022,43 @@ Variable ge ge': RTL.genv. Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. Lemma seval_preserved sp sv rs0 m0: - seval ge sp sv rs0 m0 = seval ge' sp sv rs0 m0. + seval_sval ge sp sv rs0 m0 = seval_sval ge' sp sv rs0 m0. Proof. Local Hint Resolve symbols_preserved_RTL: core. - induction sv using sval_mut with (P0 := fun lsv => list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0) - (P1 := fun sm => smem_eval ge sp sm rs0 m0 = smem_eval ge' sp sm rs0 m0); simpl; auto. - + rewrite IHsv; clear IHsv. destruct (list_sval_eval _ _ _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (smem_eval _ _ _ _); auto. + induction sv using sval_mut with (P0 := fun lsv => seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv rs0 m0) + (P1 := fun sm => seval_smem ge sp sm rs0 m0 = seval_smem ge' sp sm rs0 m0); simpl; auto. + + rewrite IHsv; clear IHsv. destruct (seval_list_sval _ _ _ _); auto. + rewrite IHsv0; clear IHsv0. destruct (seval_smem _ _ _ _); auto. erewrite eval_operation_preserved; eauto. - + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + + rewrite IHsv0; clear IHsv0. destruct (seval_list_sval _ _ _ _); auto. erewrite <- eval_addressing_preserved; eauto. destruct (eval_addressing _ sp _ _); auto. rewrite IHsv; auto. - + rewrite IHsv; clear IHsv. destruct (seval _ _ _ _); auto. + + rewrite IHsv; clear IHsv. destruct (seval_sval _ _ _ _); auto. rewrite IHsv0; auto. - + rewrite IHsv0; clear IHsv0. destruct (list_sval_eval _ _ _ _); auto. + + rewrite IHsv0; clear IHsv0. destruct (seval_list_sval _ _ _ _); auto. erewrite <- eval_addressing_preserved; eauto. destruct (eval_addressing _ sp _ _); auto. - rewrite IHsv; clear IHsv. destruct (smem_eval _ _ _ _); auto. + rewrite IHsv; clear IHsv. destruct (seval_smem _ _ _ _); auto. rewrite IHsv1; auto. Qed. Lemma list_sval_eval_preserved sp lsv rs0 m0: - list_sval_eval ge sp lsv rs0 m0 = list_sval_eval ge' sp lsv rs0 m0. + seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv rs0 m0. Proof. induction lsv; simpl; auto. - rewrite seval_preserved. destruct (seval _ _ _ _); auto. + rewrite seval_preserved. destruct (seval_sval _ _ _ _); auto. rewrite IHlsv; auto. Qed. Lemma smem_eval_preserved sp sm rs0 m0: - smem_eval ge sp sm rs0 m0 = smem_eval ge' sp sm rs0 m0. + seval_smem ge sp sm rs0 m0 = seval_smem ge' sp sm rs0 m0. Proof. induction sm; simpl; auto. - rewrite list_sval_eval_preserved. destruct (list_sval_eval _ _ _ _); auto. + rewrite list_sval_eval_preserved. destruct (seval_list_sval _ _ _ _); auto. erewrite <- eval_addressing_preserved; eauto. destruct (eval_addressing _ sp _ _); auto. - rewrite IHsm; clear IHsm. destruct (smem_eval _ _ _ _); auto. + rewrite IHsm; clear IHsm. destruct (seval_smem _ _ _ _); auto. rewrite seval_preserved; auto. Qed. @@ -1064,7 +1066,7 @@ 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 (list_sval_eval _ _ _ _); auto. + rewrite list_sval_eval_preserved. destruct (seval_list_sval _ _ _ _); auto. rewrite smem_eval_preserved; auto. Qed. @@ -1074,8 +1076,8 @@ Require Import RTLpathLivegen RTLpathLiveproofs. Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop := is1.(icontinue) = is2.(icontinue) - /\ eqlive_reg alive is1.(the_rs) is2.(the_rs) - /\ is1.(the_mem) = is2.(the_mem). + /\ eqlive_reg alive is1.(irs) is2.(irs) + /\ is1.(imem) = is2.(imem). Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := if is1.(icontinue) then @@ -1085,17 +1087,17 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := sur la dernière instruction du superblock. *) istate_simulive (fun _ => True) srce is1 is2 else - exists path, f.(fn_path)!(is1.(RTLpath.the_pc)) = Some path + exists path, f.(fn_path)!(is1.(ipc)) = Some path /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 - /\ srce!(is2.(RTLpath.the_pc)) = Some is1.(RTLpath.the_pc). + /\ srce!(is2.(ipc)) = Some is1.(ipc). -(* NOTE: a pure semantic definition on [s_istate], for a total freedom in refinements *) -Definition s_istate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: s_istate): Prop := +(* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) +Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sistate): Prop := forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> - forall rs m is1, sem_istate ge1 sp st1 rs m is1 -> - exists is2, sem_istate ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. + forall rs m is1, simatch ge1 sp st1 rs m is1 -> + exists is2, simatch ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := @@ -1108,14 +1110,14 @@ Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node | Sreturn_simu os: sfval_simu f srce opc1 opc2 (Sreturn os) (Sreturn os). -Definition s_state_simu f srce (s1 s2: s_state): Prop := - s_istate_simu f srce s1.(internal) s2.(internal) - /\ sfval_simu f srce s1.(the_pc) s2.(the_pc) s1.(final) s2.(final). +Definition sstate_simu f srce (s1 s2: sstate): Prop := + sistate_simu f srce s1.(internal) s2.(internal) + /\ sfval_simu f srce s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). -Definition symb_step_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, symb_step f1 pc1 = Some st1 -> - exists st2, symb_step f2 pc2 = Some st2 /\ s_state_simu f1 srce st1 st2. +Definition sstep_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := + forall st1, sstep f1 pc1 = Some st1 -> + exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu f1 srce st1 st2. -(* IDEA: we will provide an efficient test for checking "symb_step_simu" property, by hash-consing. - See usage of [symb_step_simu] in [RTLpathScheduler]. -*) \ No newline at end of file +(* IDEA: we will provide an efficient test for checking "sstep_simu" property, by hash-consing. + See usage of [sstep_simu] in [RTLpathScheduler]. +*) diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index fb2be412..4478ee72 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -33,7 +33,7 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> symb_step_simu dupmap f1 f2 pc1 pc2 + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 }. (* TODO: remove these two stub parameters *) @@ -110,13 +110,13 @@ 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. + Local Hint Resolve match_stackframes_equiv: core. induction 1; intros stk3 EQ; inv EQ; econstructor; eauto. 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. + Local Hint Resolve match_stack_equiv: core. destruct 1; intros EQ; inv EQ; econstructor; eauto. intros; eapply eqlive_reg_triv_trans; eauto. Qed. @@ -139,7 +139,7 @@ Qed. Lemma eqlive_match_states s1 s2 s3: eqlive_states s1 s2 -> match_states s2 s3 -> match_states s1 s3. Proof. - Local Hint Resolve eqlive_match_stack. + Local Hint Resolve eqlive_match_stack: core. destruct 1; intros MS; inv MS; try_simplify_someHyps; econstructor; eauto. eapply eqlive_reg_trans; eauto. Qed. @@ -293,15 +293,15 @@ Proof. Qed. Lemma s_find_function_preserved sp svos rs0 m0 fd: - s_find_function pge ge sp svos rs0 m0 = Some fd -> - exists fd', s_find_function tpge tge sp svos rs0 m0 = Some fd' + sfind_function pge ge sp svos rs0 m0 = Some fd -> + exists fd', sfind_function tpge tge sp svos rs0 m0 = Some fd' /\ transf_fundef fd = OK fd' /\ liveness_ok_fundef fd. Proof. - Local Hint Resolve symbols_preserved_RTL. - unfold s_find_function. destruct svos; simpl. - + rewrite (sval_eval_preserved ge tge); auto. - destruct (sval_eval _ _ _ _); try congruence. + Local Hint Resolve symbols_preserved_RTL: core. + unfold sfind_function. destruct svos; simpl. + + rewrite (seval_preserved ge tge); auto. + destruct (seval_sval _ _ _ _); try congruence. intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. eapply find_funct_liveness_ok; eauto. @@ -310,27 +310,27 @@ Proof. intros (fd' & X). eexists. intuition eauto. Qed. -Lemma sem_istate_simu f dupmap sp st st' rs m is: - sem_istate ge sp st rs m is -> +Lemma sistate_simu f dupmap sp st st' rs m is: + simatch ge sp st rs m is -> liveness_ok_function f -> - s_istate_simu f dupmap st st' -> + sistate_simu f dupmap st st' -> exists is', - sem_istate tge sp st' rs m is' /\ istate_simu f dupmap is is'. + simatch tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. intros SEM LIVE X; eapply X; eauto. Qed. -Lemma sem_sfval_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: +Lemma sfmatch_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dupmap f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) - sfval_simu f dupmap st.(the_pc) st'.(the_pc) sv sv' -> - sfval_sem pge ge sp st stk f rs0 m0 sv rs m t s -> - exists s', sfval_sem tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + sfval_simu f dupmap st.(si_pc) st'.(si_pc) sv sv' -> + sfmatch pge ge sp st stk f rs0 m0 sv rs m t s -> + exists s', sfmatch tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. - Local Hint Resolve transf_fundef_correct. + Local Hint Resolve transf_fundef_correct: core. intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. - (* Snone *) exploit initialize_path. { eapply dupmap_path_entry1; eauto. } @@ -350,38 +350,38 @@ Proof. eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. - erewrite <- sval_eval_preserved; eauto. + erewrite <- seval_preserved; eauto. Qed. (* The main theorem on simulation of symbolic states ! *) -Theorem sem_symb_state_simu dupmap f f' stk stk' sp st st' rs m t s: +Theorem smatch_sstate_simu dupmap f f' stk stk' sp st st' rs m t s: match_function dupmap f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> - sem_state pge ge sp st stk f rs m t s -> - s_state_simu f dupmap st st' -> - exists s', sem_state tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. + smatch pge ge sp st stk f rs m t s -> + sstate_simu f dupmap st st' -> + exists s', smatch tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) - exploit sem_istate_simu; eauto. + exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. intros (is' & SEM' & (path & PATH & (CONT' & RS' & M') & PC')). - exists (State stk' f' sp (RTLpath.the_pc is') (the_rs is') (the_mem is')). + exists (State stk' f' sp (ipc is') (irs is') (imem is')). split. - + eapply sem_early; auto. congruence. + + eapply smatch_early; auto. congruence. + rewrite M'. econstructor; eauto. - (* sem_normal *) - exploit sem_istate_simu; eauto. + exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. intros (is' & SEM' & (CONT' & RS' & M')). rewrite <- eqlive_reg_triv in RS'. - exploit sem_sfval_simu; eauto. + exploit sfmatch_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). - exploit sfval_sem_equiv; eauto. + exploit sfmatch_equiv; eauto. clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s1 & EQ & SEM2). exists s1; split. - + eapply sem_normal; eauto. + + eapply smatch_normal; eauto. + eapply match_states_equiv; eauto. Qed. @@ -398,13 +398,13 @@ Proof. exploit initialize_path. { eapply dupmap_path_entry2; eauto. } intros (path' & PATH'). exists path'. - exploit (symb_step_correct f pc pge ge sp path stk rs m t s); eauto. + exploit (sstep_correct f pc pge ge sp path stk rs m t s); eauto. intros (st & SYMB & SEM); clear STEP. exploit dupmap_correct; eauto. clear SYMB; intros (st' & SYMB & SIMU). - exploit sem_symb_state_simu; eauto. + exploit smatch_sstate_simu; eauto. intros (s0 & SEM0 & MATCH). - exploit symb_step_exact; eauto. + exploit sstep_exact; eauto. intros (s' & STEP' & EQ). exists s'; intuition. eapply match_states_equiv; eauto. @@ -417,7 +417,7 @@ Lemma step_simulation s1 t s1' s2: step tge tpge s2 t s2' /\ match_states s1' s2'. Proof. - Local Hint Resolve eqlive_stacks_refl transf_fundef_correct. + Local Hint Resolve eqlive_stacks_refl transf_fundef_correct: core. destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS. (* exec_path *) - try_simplify_someHyps. intros. -- cgit From f6bf6fcb22cc2b4e9e71e5229ed2adbb8cc277c0 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 23 Apr 2020 17:57:58 +0200 Subject: [BROKEN] Added wf predicate sistate - simatch_exclude_abort proven, need to modify some functions --- mppa_k1c/lib/RTLpathSE_theory.v | 198 ++++++++++++++++++++++++---------------- 1 file changed, 121 insertions(+), 77 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index e38e66dd..c503936b 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -110,24 +110,32 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := exists ext, List.In ext lx /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. - -Inductive simatch_exits (ge: RTL.genv) (sp:val) (lx:list sistate_exit) (rs0: regset) (m0: mem) rs m pc: Prop := - simatch_exits_intro ext lx': - is_tail (ext::lx') lx -> - all_fallthrough ge sp lx' rs0 m0 -> - seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some true -> - simatch_local ge sp ext.(si_elocal) rs0 m0 rs m -> - ext.(si_ifso) = pc -> - simatch_exits ge sp lx rs0 m0 rs m pc. - +Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := + simatch_exits_intro: forall ext lx', + lx = ext::lx' -> + all_fallthrough ge sp lx' rs m -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> + simatch_local ge sp (si_elocal ext) rs m rs' m' -> + (si_ifso ext) = pc' -> + simatch_exits ge sp lx rs m rs' m' pc'. + +(* Ensuring there isn't more sistate_local after taking a branch *) +Inductive simatch_exits_bind (ext: sistate_exit) (loc: sistate_local) : Prop := + simatch_exits_bind_intro: + (forall ge sp lx rs m rs' m' pc', simatch_exits ge sp (ext::lx) rs m rs' m' pc' -> si_elocal ext = loc) -> + simatch_exits_bind ext loc. + +Definition sistate_wf exits loc := + exists ext lx, exits = ext::lx /\ simatch_exits_bind ext loc. (** * Syntax and Semantics of symbolic internal state *) -Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. +Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local; + si_wf: sistate_wf si_exits si_local }. Definition simatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(icontinue)) then - simatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) + simatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) /\ st.(si_pc) = is.(ipc) /\ all_fallthrough ge sp st.(si_exits) rs0 m0 else simatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) is.(ipc). @@ -162,7 +170,7 @@ Definition istate_eq ist1 ist2 := (* TODO: is it useful Lemma has_not_yet_exit_cons_split ge sp ext lx rs0 m0: - all_fallthrough ge sp (ext::lx) rs0 m0 <-> + all_fallthrough ge sp (ext::lx) rs0 m0 <-> (seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(si_smem) rs0 m0 = Some false /\ all_fallthrough ge sp lx rs0 m0). Proof. unfold all_fallthrough; simpl; intuition (subst; auto). @@ -177,7 +185,7 @@ Proof. Local Hint Resolve is_tail_in: core. destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. unfold all_fallthrough; intros NYE; rewrite NYE in EVAL; eauto. - try congruence. + try congruence. rewrite lc. econstructor; eauto. Qed. Lemma simatch_exclude_incompatible_continue ge sp st rs m is1 is2: @@ -249,9 +257,9 @@ Proof. Local Hint Resolve exit_cond_determ eq_sym: core. destruct 1 as [ext1 lx1 TAIL1 NYE1 EVAL1 SEM1 PC1]; subst. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. - assert (X:ext1=ext2). - - destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. - - subst. destruct (simatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. + assert (X:ext1=ext2) by congruence. +(*{ destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. } *) + subst. destruct (simatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. Lemma simatch_determ ge sp st rs m is1 is2: @@ -270,58 +278,94 @@ Proof. intuition. Qed. -Lemma simatch_exclude_sabort_local_continue ge sp st rs m is: - icontinue is = true -> +Lemma simatch_local_exclude_sabort_local ge sp st rs m rs' m': + simatch_local ge sp (si_local st) rs m rs' m' -> + all_fallthrough ge sp (si_exits st) rs m -> sabort_local ge sp (si_local st) rs m -> - simatch ge sp st rs m is -> False. + False. Proof. - intros CONT ABORT SEM. unfold simatch in SEM. rewrite CONT in SEM. - destruct SEM as (SEML & _ & _). inversion SEML as (SEML1 & SEML2 & SEML3); clear SEML. - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. - - contradiction. - - rewrite SEML2 in ABORT2. discriminate. - - inv ABORT3. rewrite SEML3 in H. discriminate. + intros SIML ALLF ABORT. inv SIML. destruct H0 as (H0 & H0'). + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. Qed. -Lemma simatch_exclude_sabort_local_nocontinue ge sp st rs0 m0 is: - icontinue is = false -> - sabort_local ge sp (si_local st) rs0 m0 -> - simatch ge sp st rs0 m0 is -> False. +Lemma simatch_local_exclude_sabort_exits ge sp st rs m rs' m': + simatch_local ge sp (si_local st) rs m rs' m' -> + all_fallthrough ge sp (si_exits st) rs m -> + sabort_exits ge sp (si_exits st) rs m -> + False. Proof. -(* intros CONT ABORT SEM. unfold simatch in SEM. rewrite CONT in SEM. - destruct st as [pc si_exits iis]. simpl in *. - destruct is as [cont pc' rs' m']. simpl in *. - inv SEM. inversion H2 as (SEML1 & SEML2 & SEML3); clear H2. - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; clear ABORT. - - Search iis. - - rewrite SEML2 in ABORT2. discriminate. - - inv ABORT3. rewrite SEML3 in H. discriminate. *) -Admitted. + intros SIML ALLF ABORT. inv ABORT. inv H. + unfold all_fallthrough in ALLF. apply ALLF in H0. congruence. +Qed. +Lemma simatch_local_exclude_sabort ge sp st rs m rs' m': + simatch_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. + - eapply simatch_local_exclude_sabort_local; eauto. + - eapply simatch_local_exclude_sabort_exits; eauto. +Qed. -Lemma simatch_exclude_sabort_local ge sp st rs m is: +Lemma simatch_exits_exclude_sabort_local ge sp st rs m rs' m' pc': + simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> sabort_local ge sp (si_local st) rs m -> - simatch ge sp st rs m is -> False. + False. Proof. - intros. destruct (icontinue is) eqn:CONT. - - eapply simatch_exclude_sabort_local_continue; eauto. - - eapply simatch_exclude_sabort_local_nocontinue; eauto. + destruct st as [spc exits loc SISTWF]. simpl. + intros SIMEX ABORT. inv SISTWF. inv H. inv H0. + inversion SIMEX. subst. inv H. + inv H1. eapply H in SIMEX. subst. clear H. + inv H3. destruct H1 as (H1 & H1'). + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. Qed. -Lemma simatch_exclude_sabort_exits ge sp st rs m is: - sabort_exits ge sp (si_exits st) rs m -> - simatch ge sp st rs m is -> False. +(* Weaker version if we don't have the simatch_exits_bind - not used anymore, couldn't get any use from it *) +Remark simatch_exits_exclude_sabort_local' ge sp rs m rs' m' pc' exits ext: + simatch_exits ge sp (ext::exits) rs m rs' m' pc' -> + sabort_local ge sp (si_elocal ext) rs m -> + False. Proof. -Admitted. + intros SIMEX ABORT. + inv SIMEX. inv H. inv H2. destruct H3 as (H3 & H3'). + inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. +Qed. + +Lemma simatch_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': + simatch_exits ge sp exits rs m rs' m' pc' -> + sabort_exits ge sp exits rs m -> + False. +Proof. + intros SIMEX ABORT. + inv ABORT. inv H. inv SIMEX. unfold all_fallthrough in H2. + inv H0. + - congruence. + - apply H2 in H. congruence. +Qed. + +Lemma simatch_exits_exclude_sabort ge sp st rs m rs' m' pc': + simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> + sabort ge sp st rs m -> + False. +Proof. + intros SIME ABORT. + inv ABORT. + - eapply simatch_exits_exclude_sabort_local; eauto. + - eapply simatch_exits_exclude_sabort_exits; eauto. +Qed. Lemma simatch_exclude_sabort ge sp st rs m is: sabort ge sp st rs m -> simatch ge sp st rs m is -> False. Proof. intros ABORT SEM. - inv ABORT. - - eapply simatch_exclude_sabort_local; eauto. - - eapply simatch_exclude_sabort_exits; eauto. + unfold simatch in SEM. destruct icontinue. + - destruct SEM as (SEM1 & SEM2 & SEM3). + eapply simatch_local_exclude_sabort; eauto. + - eapply simatch_exits_exclude_sabort; eauto. Qed. Definition istate_eq_opt ist1 oist := @@ -353,7 +397,7 @@ Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. (* FIXME - add notrap *) -Definition sistep (i: instruction) (st: sistate): option sistate := +Definition sistep (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => sist_set_local st pc' st.(si_local) @@ -420,7 +464,7 @@ Proof. unfold sabort_local. simpl; intuition. Qed. -Lemma sistep_preserves_sabort i ge sp rs0 m0 st st': +Lemma sistep_preserves_sabort i ge sp rs0 m0 st st': sistep i st = Some st' -> sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. Proof. @@ -463,7 +507,7 @@ Proof. subst; rewrite Regmap.gss; simpl; auto. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - - unfold sabort, sabort_local; simpl. + - unfold sabort, sabort_local; simpl. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. @@ -482,7 +526,7 @@ Proof. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } - { unfold sabort, sabort_local; simpl. + { unfold sabort, sabort_local; simpl. left. right. right. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. @@ -527,7 +571,7 @@ Admitted. Lemma sistep_correct_None ge sp i st rs0 m0 rs m: simatch_local ge sp (st.(si_local)) rs0 m0 rs m -> - sistep i st = None -> + sistep i st = None -> istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). @@ -544,9 +588,9 @@ Fixpoint sisteps (path:nat) (f: function) (st: sistate): option sistate := sisteps p f st1 end. -Lemma sisteps_correct_false ge sp path f rs0 m0 st' is: +Lemma sisteps_correct_false ge sp path f rs0 m0 st' is: is.(icontinue)=false -> - forall st, simatch ge sp st rs0 m0 is -> + forall st, simatch ge sp st rs0 m0 is -> sisteps path f st = Some st' -> simatch ge sp st' rs0 m0 is. Proof. @@ -558,7 +602,7 @@ Proof. inversion_SOME st1; eauto. Qed. -Lemma sisteps_preserves_sabort ge sp path f rs0 m0 st': forall st, +Lemma sisteps_preserves_sabort ge sp path f rs0 m0 st': forall st, sisteps path f st = Some st' -> sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. Proof. @@ -592,7 +636,7 @@ Qed. Lemma sisteps_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(icontinue)=true -> - simatch ge sp st rs0 m0 is -> + simatch ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(si_pc) <> None -> simatch_opt2 ge sp (sisteps path f st) rs0 m0 (isteps ge path f sp is.(irs) is.(imem) is.(ipc)) @@ -608,7 +652,7 @@ Proof. intros (LOCAL & PC & NYE) WF. rewrite <- PC. inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. - exploit sistep_correct; eauto. + exploit sistep_correct; eauto. inversion_SOME st1; intros Hst1; erewrite Hst1; simpl. - inversion_SOME is1; intros His1;rewrite His1; simpl. * destruct (icontinue is1) eqn:CONT1. @@ -707,7 +751,7 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac sp = (Vptr stk Ptrofs.zero) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> match osv with Some sv => seval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> - sfmatch pge ge sp st stack f rs0 m0 (Sreturn osv) rs m + sfmatch pge ge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . @@ -716,17 +760,17 @@ Record sstate := { internal:> sistate; final: sfval }. Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | smatch_early is: is.(icontinue) = false -> - simatch ge sp st rs0 m0 is -> + simatch ge sp st rs0 m0 is -> smatch pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) | smatch_normal is t s: is.(icontinue) = true -> - simatch ge sp st rs0 m0 is -> + simatch ge sp st rs0 m0 is -> sfmatch pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> - smatch pge ge sp st stack f rs0 m0 t s + smatch pge ge sp st stack f rs0 m0 t s . (** * Symbolic execution of final step *) -Definition sfstep (i: instruction) (prev: sistate_local): sfval := +Definition sfstep (i: instruction) (prev: sistate_local): sfval := match i with | Icall sig ros args res pc => let svos := sum_left_map prev.(si_sreg) ros in @@ -743,7 +787,7 @@ Lemma sfstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: pc = st.(si_pc) -> simatch_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> - sistep i st = None -> + sistep i st = None -> sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. @@ -818,9 +862,9 @@ Lemma symb_path_last_step i st st' ge pge stack (f:function) sp pc rs m t s: pc = st.(si_pc) -> sistep 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 + 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. @@ -829,7 +873,7 @@ Qed. (* NB: each concrete execution can be executed on the symbolic state (produced from [sstep]) (sstep is a correct over-approximation) *) -Theorem sstep_correct f pc pge ge sp path stack rs m t s: +Theorem sstep_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, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. @@ -846,7 +890,7 @@ Proof. (* intro SEM *) (simpl; unfold simatch; simpl; rewrite CONT; intro SEM); (* intro Hi' *) - ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); + ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); [ unfold nth_default_succ_inst in Hi; exploit sisteps_default_succ; eauto; simpl; intros DEF; rewrite DEF in Hi; auto @@ -963,9 +1007,9 @@ Qed. (* NB: each execution of a symbolic state (produced from [sstep]) represents a concrete execution (sstep is exact). *) -Theorem sstep_exact f pc pge ge sp path stack st rs m t s1: +Theorem sstep_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> - sstep f pc = Some st -> + sstep f pc = Some st -> smatch 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. @@ -1087,7 +1131,7 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := sur la dernière instruction du superblock. *) istate_simulive (fun _ => True) srce is1 is2 else - exists path, f.(fn_path)!(is1.(ipc)) = Some path + 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). @@ -1096,7 +1140,7 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> - forall rs m is1, simatch ge1 sp st1 rs m is1 -> + forall rs m is1, simatch ge1 sp st1 rs m is1 -> exists is2, simatch ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) @@ -1115,7 +1159,7 @@ Definition sstate_simu f srce (s1 s2: sstate): Prop := /\ sfval_simu f srce s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). Definition sstep_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, sstep f1 pc1 = Some st1 -> + forall st1, sstep f1 pc1 = Some st1 -> exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu f1 srce st1 st2. (* IDEA: we will provide an efficient test for checking "sstep_simu" property, by hash-consing. -- cgit From baedc5d727f648327b18c8cb95a27491850fbe2e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 24 Apr 2020 12:34:53 +0200 Subject: Reverting change to sistate --- mppa_k1c/lib/RTLpathSE_theory.v | 46 +++++++---------------------------------- 1 file changed, 7 insertions(+), 39 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index c503936b..66e23cd8 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -112,25 +112,15 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := simatch_exits_intro: forall ext lx', - lx = ext::lx' -> + is_tail (ext::lx') lx -> all_fallthrough ge sp lx' rs m -> seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> simatch_local ge sp (si_elocal ext) rs m rs' m' -> (si_ifso ext) = pc' -> simatch_exits ge sp lx rs m rs' m' pc'. -(* Ensuring there isn't more sistate_local after taking a branch *) -Inductive simatch_exits_bind (ext: sistate_exit) (loc: sistate_local) : Prop := - simatch_exits_bind_intro: - (forall ge sp lx rs m rs' m' pc', simatch_exits ge sp (ext::lx) rs m rs' m' pc' -> si_elocal ext = loc) -> - simatch_exits_bind ext loc. - -Definition sistate_wf exits loc := - exists ext lx, exits = ext::lx /\ simatch_exits_bind ext loc. - (** * Syntax and Semantics of symbolic internal state *) -Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local; - si_wf: sistate_wf si_exits si_local }. +Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }. Definition simatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := if (is.(icontinue)) @@ -185,7 +175,7 @@ Proof. Local Hint Resolve is_tail_in: core. destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. unfold all_fallthrough; intros NYE; rewrite NYE in EVAL; eauto. - try congruence. rewrite lc. econstructor; eauto. + try congruence. (* rewrite lc. econstructor; eauto. *) Qed. Lemma simatch_exclude_incompatible_continue ge sp st rs m is1 is2: @@ -257,8 +247,8 @@ Proof. Local Hint Resolve exit_cond_determ eq_sym: core. destruct 1 as [ext1 lx1 TAIL1 NYE1 EVAL1 SEM1 PC1]; subst. destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. - assert (X:ext1=ext2) by congruence. -(*{ destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. } *) + assert (X:ext1=ext2). + { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. } subst. destruct (simatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. Qed. @@ -315,36 +305,14 @@ Lemma simatch_exits_exclude_sabort_local ge sp st rs m rs' m' pc': sabort_local ge sp (si_local st) rs m -> False. Proof. - destruct st as [spc exits loc SISTWF]. simpl. - intros SIMEX ABORT. inv SISTWF. inv H. inv H0. - inversion SIMEX. subst. inv H. - inv H1. eapply H in SIMEX. subst. clear H. - inv H3. destruct H1 as (H1 & H1'). - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. -Qed. - -(* Weaker version if we don't have the simatch_exits_bind - not used anymore, couldn't get any use from it *) -Remark simatch_exits_exclude_sabort_local' ge sp rs m rs' m' pc' exits ext: - simatch_exits ge sp (ext::exits) rs m rs' m' pc' -> - sabort_local ge sp (si_elocal ext) rs m -> - False. -Proof. - intros SIMEX ABORT. - inv SIMEX. inv H. inv H2. destruct H3 as (H3 & H3'). - inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. -Qed. +Admitted. Lemma simatch_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': simatch_exits ge sp exits rs m rs' m' pc' -> sabort_exits ge sp exits rs m -> False. Proof. - intros SIMEX ABORT. - inv ABORT. inv H. inv SIMEX. unfold all_fallthrough in H2. - inv H0. - - congruence. - - apply H2 in H. congruence. -Qed. +Admitted. Lemma simatch_exits_exclude_sabort ge sp st rs m rs' m' pc': simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> -- cgit From 811b511dd505065b5fed51cae6fe341a89788acb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Apr 2020 19:16:19 +0200 Subject: fixing sabort_exits ? --- mppa_k1c/lib/RTLpathSE_theory.v | 17 ++++++++++++++++- 1 file changed, 16 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 66e23cd8..210377c8 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -105,10 +105,25 @@ 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. - +(* TOO WEAK DEFINITION ! Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := exists ext, List.In ext lx /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. +*) + +Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := + sabort_exits_on_cond: forall ext lx', + is_tail (ext::lx') lx -> + all_fallthrough ge sp lx' rs m -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = None -> + sabort_exits ge sp lx rs m + | sabort_exits_after_cond: forall ext lx', + is_tail (ext::lx') lx -> + all_fallthrough ge sp lx' rs m -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> + sabort_local ge sp (si_elocal ext) rs m -> + sabort_exits ge sp lx rs m. + Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := simatch_exits_intro: forall ext lx', -- cgit From cfcda01c880743abe6eb63e3c1279654b798f2de Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 25 Apr 2020 03:36:43 +0200 Subject: a simpler fix ? --- mppa_k1c/lib/RTLpathSE_theory.v | 20 ++++++++++++++++++-- 1 file changed, 18 insertions(+), 2 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 210377c8..aa2c6aeb 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -111,6 +111,18 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. *) +Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := + sabort_exits_on_cond: forall ext, List.In ext lx -> + seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs m = None -> + sabort_exits ge sp lx rs m + | sabort_exits_after_cond: forall ext, List.In ext lx -> + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> + sabort_local ge sp (si_elocal ext) rs m -> + sabort_exits ge sp lx rs m. + +(* VERSION UNNECESSARILY COMPLEX ? + to try only if the above one does not work... + Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := sabort_exits_on_cond: forall ext lx', is_tail (ext::lx') lx -> @@ -123,6 +135,9 @@ Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: reg seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> sabort_local ge sp (si_elocal ext) rs m -> sabort_exits ge sp lx rs m. +*) + + Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := @@ -299,8 +314,9 @@ Lemma simatch_local_exclude_sabort_exits ge sp st rs m rs' m': sabort_exits ge sp (si_exits st) rs m -> False. Proof. - intros SIML ALLF ABORT. inv ABORT. inv H. - unfold all_fallthrough in ALLF. apply ALLF in H0. congruence. + intros SIML ALLF ABORT. inv ABORT. + - exploit ALLF; eauto. congruence. + - exploit ALLF; eauto. congruence. Qed. Lemma simatch_local_exclude_sabort ge sp st rs m rs' m': -- cgit From 13954b8ed5646e740de695165cf06b0876d69f61 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 25 Apr 2020 04:25:20 +0200 Subject: no: the simpler fix is too simple... --- mppa_k1c/lib/RTLpathSE_theory.v | 36 ++++++++++++++---------------------- 1 file changed, 14 insertions(+), 22 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index aa2c6aeb..67c86219 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -111,18 +111,6 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. *) -Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := - sabort_exits_on_cond: forall ext, List.In ext lx -> - seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs m = None -> - sabort_exits ge sp lx rs m - | sabort_exits_after_cond: forall ext, List.In ext lx -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> - sabort_local ge sp (si_elocal ext) rs m -> - sabort_exits ge sp lx rs m. - -(* VERSION UNNECESSARILY COMPLEX ? - to try only if the above one does not work... - Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := sabort_exits_on_cond: forall ext lx', is_tail (ext::lx') lx -> @@ -135,10 +123,6 @@ Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: reg seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> sabort_local ge sp (si_elocal ext) rs m -> sabort_exits ge sp lx rs m. -*) - - - Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := simatch_exits_intro: forall ext lx', @@ -161,8 +145,8 @@ Definition simatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) else simatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) is.(ipc). Definition sabort (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := - sabort_local ge sp st.(si_local) rs0 m0 - \/ sabort_exits ge sp st.(si_exits) rs0 m0. + (all_fallthrough ge sp st.(si_exits) rs0 m0 /\ sabort_local ge sp st.(si_local) rs0 m0) + \/ sabort_exits ge sp st.(si_exits) rs0 m0. Definition simatch_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with @@ -327,22 +311,30 @@ Lemma simatch_local_exclude_sabort ge sp st rs m rs' m': Proof. intros SIML ALLF ABORT. inv ABORT. - - eapply simatch_local_exclude_sabort_local; eauto. + - intuition; eapply simatch_local_exclude_sabort_local; eauto. - eapply simatch_local_exclude_sabort_exits; eauto. Qed. Lemma simatch_exits_exclude_sabort_local ge sp st rs m rs' m' pc': simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> + all_fallthrough ge sp (si_exits st) rs m -> sabort_local ge sp (si_local st) rs m -> False. Proof. -Admitted. + intros SIML ALLF ABORT. + inv SIML. + exploit ALLF; eauto. + congruence. +Qed. Lemma simatch_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': simatch_exits ge sp exits rs m rs' m' pc' -> sabort_exits ge sp exits rs m -> False. Proof. + intros SIML ABORT. + inv SIML. + inv ABORT. Admitted. Lemma simatch_exits_exclude_sabort ge sp st rs m rs' m' pc': @@ -351,8 +343,8 @@ Lemma simatch_exits_exclude_sabort ge sp st rs m rs' m' pc': False. Proof. intros SIME ABORT. - inv ABORT. - - eapply simatch_exits_exclude_sabort_local; eauto. + inv ABORT. + - intuition. eapply simatch_exits_exclude_sabort_local; eauto. - eapply simatch_exits_exclude_sabort_exits; eauto. Qed. -- cgit From 472b9db723ee17a098f84880f3915cd0d2940e1d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 27 Apr 2020 18:46:56 +0200 Subject: Cutting ssem_exits into two pieces + advancing some proofs --- mppa_k1c/lib/RTLpathSE_theory.v | 378 ++++++++++++++++++++++++++-------------- 1 file changed, 248 insertions(+), 130 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 67c86219..9978b3f3 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -83,7 +83,8 @@ with seval_smem (ge: RTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): opti (* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) Record sistate_local := { si_pre: RTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }. -Definition simatch_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := +(* Predicate on which (rs, m) is a possible final state after evaluating [st] on (rs0, m0) *) +Definition ssem_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop := st.(si_pre) ge sp rs0 m0 /\ seval_smem ge sp st.(si_smem) rs0 m0 = Some m /\ forall (r:reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r). @@ -111,7 +112,7 @@ Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. *) -Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := +(* Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := sabort_exits_on_cond: forall ext lx', is_tail (ext::lx') lx -> all_fallthrough ge sp lx' rs m -> @@ -122,47 +123,79 @@ Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: reg all_fallthrough ge sp lx' rs m -> seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> sabort_local ge sp (si_elocal ext) rs m -> - sabort_exits ge sp lx rs m. + sabort_exits ge sp lx rs m. *) -Inductive simatch_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := - simatch_exits_intro: forall ext lx', - is_tail (ext::lx') lx -> - all_fallthrough ge sp lx' rs m -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> - simatch_local ge sp (si_elocal ext) rs m rs' m' -> - (si_ifso ext) = pc' -> - simatch_exits ge sp lx rs m rs' m' pc'. +(* 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'. + +Definition sabort_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) : Prop := + seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = None + \/ 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 simatch (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem) (is: istate): Prop := +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 (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop := if (is.(icontinue)) then - simatch_local ge sp st.(si_local) rs0 m0 is.(irs) is.(imem) + 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) rs0 m0 - else simatch_exits ge sp st.(si_exits) rs0 m0 is.(irs) is.(imem) 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) (rs0: regset) (m0: mem): Prop := +(* Inductive ssem_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := + ssem_exits_intro: forall ext lx', + is_tail (ext::lx') lx -> + all_fallthrough ge sp lx' rs m -> + 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' -> + ssem_exits ge sp lx rs m rs' m' pc'. *) + +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 sabort (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := (all_fallthrough ge sp st.(si_exits) rs0 m0 /\ sabort_local ge sp st.(si_local) rs0 m0) - \/ sabort_exits ge sp st.(si_exits) rs0 m0. + \/ sabort_exits ge sp st.(si_exits) rs0 m0. *) -Definition simatch_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := +Definition ssem_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => simatch ge sp st rs0 m0 is + | Some is => ssem ge sp st rs0 m0 is | None => sabort ge sp st rs0 m0 end. -Definition simatch_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := +Definition ssem_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => simatch_opt ge sp st rs0 m0 ois + | Some st => ssem_opt ge sp st rs0 m0 ois | None => ois=None end. (** * An internal state represents a parallel program ! - We prove below that the semantics [simatch_opt] is deterministic. + We prove below that the semantics [ssem_opt] is deterministic. *) @@ -181,48 +214,53 @@ Proof. Qed. *) -Lemma all_fallthrough_noexit ge sp st rs0 m0 rs m pc: - simatch_exits ge sp (si_exits st) rs0 m0 rs m pc -> - all_fallthrough ge sp (si_exits st) rs0 m0 -> +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. - destruct 1 as [ext lx' lc NYE' EVAL SEM PC]. subst. - unfold all_fallthrough; intros NYE; rewrite NYE in EVAL; eauto. - try congruence. (* rewrite lc. econstructor; eauto. *) + intros SSEM INE ALLF. + destruct SSEM as (SSEM & SSEM'). + unfold all_fallthrough in ALLF. rewrite ALLF in SSEM; eauto. + discriminate. Qed. -Lemma simatch_exclude_incompatible_continue ge sp st rs m is1 is2: +Lemma ssem_exclude_incompatible_continue ge sp st rs m is1 is2: is1.(icontinue) = true -> is2.(icontinue) = false -> - simatch ge sp st rs m is1 -> - simatch ge sp st rs m is2 -> + ssem ge sp st rs m is1 -> + ssem ge sp st rs m is2 -> False. Proof. Local Hint Resolve all_fallthrough_noexit: core. - unfold simatch. + unfold ssem. 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 simatch_determ_continue ge sp st rs m is1 is2: - simatch ge sp st rs m is1 -> - simatch ge sp st rs m is2 -> +Lemma ssem_determ_continue ge sp st rs m is1 is2: + ssem ge sp st rs m is1 -> + ssem ge sp st rs m is2 -> is1.(icontinue) = is2.(icontinue). Proof. - Local Hint Resolve simatch_exclude_incompatible_continue: core. + Local Hint Resolve ssem_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 simatch_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2: - simatch_local ge sp st rs0 m0 rs1 m1 -> - simatch_local ge sp st rs0 m0 rs2 m2 -> +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 simatch_local. intuition try congruence. + unfold ssem_local. intuition try congruence. generalize (H5 r); rewrite H4; congruence. Qed. @@ -253,47 +291,59 @@ Proof. try congruence. Qed. -Lemma simatch_exits_determ ge sp lx rs0 m0 rs1 m1 pc1 rs2 m2 pc2 : - simatch_exits ge sp lx rs0 m0 rs1 m1 pc1 -> - simatch_exits ge sp lx rs0 m0 rs2 m2 pc2 -> +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. - destruct 1 as [ext1 lx1 TAIL1 NYE1 EVAL1 SEM1 PC1]; subst. - destruct 1 as [ext2 lx2 TAIL2 NYE2 EVAL2 SEM2 PC2]; subst. - assert (X:ext1=ext2). - { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2::lx2) lx) as [TAIL|TAIL]; eauto. } - subst. destruct (simatch_local_determ ge sp (si_elocal ext2) rs0 m0 rs1 m1 rs2 m2); auto. + 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. + +(* Actually this lemma isn't used *) +Remark is_tail_cons_eq {A: Type} (a a': A) l l': + is_tail (a::l) (a'::l') -> + a = a' \/ (In a l'). +Proof. + intros. inv H. + - left. reflexivity. + - right. eapply is_tail_in; eauto. Qed. -Lemma simatch_determ ge sp st rs m is1 is2: - simatch ge sp st rs m is1 -> - simatch ge sp st rs m is2 -> +Lemma ssem_determ ge sp st rs m is1 is2: + ssem ge sp st rs m is1 -> + ssem ge sp st rs m is2 -> istate_eq is1 is2. Proof. unfold istate_eq. intros SEM1 SEM2. - exploit (simatch_determ_continue ge sp st rs m is1 is2); eauto. - intros CONTEQ. unfold simatch in * |-. rewrite CONTEQ in * |- *. + exploit (ssem_determ_continue ge sp st rs m is1 is2); eauto. + intros CONTEQ. unfold ssem in * |-. rewrite CONTEQ in * |- *. destruct (icontinue is2). - - destruct (simatch_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); + - destruct (ssem_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2)); intuition (try congruence). - - destruct (simatch_exits_determ ge sp (si_exits st) rs m (irs is1) (imem is1) (ipc is1) (irs is2) (imem is2) (ipc is2)); - intuition. + - 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 simatch_local_exclude_sabort_local ge sp st rs m rs' m': - simatch_local ge sp (si_local st) rs m rs' m' -> - all_fallthrough ge sp (si_exits st) rs m -> - sabort_local ge sp (si_local st) rs m -> +Lemma ssem_local_exclude_sabort_local ge sp loc rs m rs' m': + ssem_local ge sp loc rs m rs' m' -> +(* all_fallthrough ge sp (si_exits st) rs m -> *) + sabort_local ge sp loc rs m -> False. Proof. - intros SIML ALLF ABORT. inv SIML. destruct H0 as (H0 & H0'). + intros SIML (* ALLF *) ABORT. inv SIML. destruct H0 as (H0 & H0'). inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. Qed. -Lemma simatch_local_exclude_sabort_exits ge sp st rs m rs' m': - simatch_local ge sp (si_local st) rs m rs' m' -> +(* Lemma ssem_local_exclude_sabort_exits 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_exits ge sp (si_exits st) rs m -> False. @@ -301,22 +351,51 @@ Proof. intros SIML ALLF ABORT. inv ABORT. - exploit ALLF; eauto. congruence. - exploit ALLF; eauto. congruence. -Qed. +Qed. *) -Lemma simatch_local_exclude_sabort ge sp st rs m rs' m': - simatch_local ge sp (si_local st) rs m rs' m' -> +(* Lemma all_fallthrough_exclude_sabort_exit ge sp ext lx rs m: + In ext lx -> + all_fallthrough ge sp lx rs m -> + sabort_exit ge sp ext rs m -> + False. +Proof. + intros INL ALLF ABORT. inv ABORT. + - unfold all_fallthrough in ALLF. eapply ALLF in INL. congruence. + - unfold all_fallthrough in ALLF. eapply ALLF in INL. unfold sabort_local in H. congruence. + intros SIML ALLF ABORT. inv ABORT. + - exploit ALLF; eauto. congruence. + - exploit ALLF; eauto. congruence. +Qed. *) + +Lemma ssem_local_exclude_sabort_exit ge sp st ext lx rs m rs' m': + ssem_local ge sp (si_local st) rs m rs' m' -> + all_fallthrough ge sp (si_exits st) rs m -> + is_tail (ext :: lx) (si_exits st) -> + sabort_exit ge sp ext rs m -> + False. +Proof. + intros SSEML ALLF TAIL ABORT. + inv ABORT. + - exploit ALLF; eauto. congruence. + - (* FIXME Problem : if we have a ssem_local, this means we ONLY evaluated the conditions, + but we NEVER actually evaluated the si_elocal from the sistate_exit ! So we cannot prove + a lack of abort on the si_elocal.. We must change the definitions *) +Admitted. + +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 simatch_local_exclude_sabort_local; eauto. - - eapply simatch_local_exclude_sabort_exits; eauto. + - intuition; eapply ssem_local_exclude_sabort_local; eauto. + - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; eauto. Qed. -Lemma simatch_exits_exclude_sabort_local ge sp st rs m rs' m' pc': - simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> +(* Lemma ssem_exit_exclude_sabort_local ge sp st rs m rs' m' pc': + ssem_exit ge sp (si_exit st) rs m rs' m' pc' -> all_fallthrough ge sp (si_exits st) rs m -> sabort_local ge sp (si_local st) rs m -> False. @@ -325,51 +404,88 @@ Proof. inv SIML. exploit ALLF; eauto. congruence. -Qed. +Qed. *) -Lemma simatch_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': - simatch_exits ge sp exits rs m rs' m' pc' -> +(* Lemma ssem_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': + ssem_exits ge sp exits rs m rs' m' pc' -> sabort_exits ge sp exits rs m -> False. Proof. intros SIML ABORT. inv SIML. inv ABORT. -Admitted. +Admitted. *) -Lemma simatch_exits_exclude_sabort ge sp st rs m rs' m' pc': - simatch_exits ge sp (si_exits st) rs m rs' m' pc' -> +(* Lemma ssem_exits_exclude_sabort ge sp st rs m rs' m' pc': + ssem_exits ge sp (si_exits st) rs m rs' m' pc' -> sabort ge sp st rs m -> False. Proof. intros SIME ABORT. inv ABORT. - - intuition. eapply simatch_exits_exclude_sabort_local; eauto. - - eapply simatch_exits_exclude_sabort_exits; eauto. + - intuition. eapply ssem_exits_exclude_sabort_local; eauto. + - eapply ssem_exits_exclude_sabort_exits; eauto. +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. +Admitted. + +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. + - eapply ssem_local_exclude_sabort_local; eauto. Qed. -Lemma simatch_exclude_sabort ge sp st rs m is: +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. + - admit. (* FIXME - we can have a sabort_local ; we don't have any ssem_local on si_local st ! *) + - 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_cons_eq ext' ext lx' lx); eauto. intro. inv H. + + eapply ssem_exit_exclude_sabort_exit; eauto. + + admit. + (* we can prove we cannot abort on the seval_condition because ext' is in lx, and all_fallthrough lx + However, we run into the same issue as above for si_elocal *) +Admitted. + +Lemma ssem_exclude_sabort ge sp st rs m is: sabort ge sp st rs m -> - simatch ge sp st rs m is -> False. + ssem ge sp st rs m is -> False. Proof. intros ABORT SEM. - unfold simatch in SEM. destruct icontinue. + unfold ssem in SEM. destruct icontinue. - destruct SEM as (SEM1 & SEM2 & SEM3). - eapply simatch_local_exclude_sabort; eauto. - - eapply simatch_exits_exclude_sabort; eauto. + 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 simatch_opt_determ ge sp st rs m ois is: - simatch_opt ge sp st rs m ois -> - simatch ge sp st rs m is -> +Lemma ssem_opt_determ ge sp st rs m ois is: + ssem_opt ge sp st rs m ois -> + ssem ge sp st rs m is -> istate_eq_opt is ois. Proof. destruct ois as [is1|]; simpl; eauto. - - intros; eexists; intuition; eapply simatch_determ; eauto. - - intros; exploit simatch_exclude_sabort; eauto. destruct 1. + - intros; eexists; intuition; eapply ssem_determ; eauto. + - intros; exploit ssem_exclude_sabort; eauto. destruct 1. Qed. (** * Symbolic execution of one internal step *) @@ -427,13 +543,15 @@ Proof. try_simplify_someHyps. Qed. -Lemma sistep_preserves_simatch_exits i ge sp st rs0 m0 rs m pc st': - simatch_exits ge sp st.(si_exits) rs0 m0 rs m pc -> +(* TODO - continue updating the rest *) + +Lemma sistep_preserves_ssem_exits i ge sp st rs0 m0 rs m pc st': + ssem_exits ge sp st.(si_exits) rs0 m0 rs m pc -> sistep i st = Some st' -> - simatch_exits ge sp st'.(si_exits) rs0 m0 rs m pc. + ssem_exits ge sp st'.(si_exits) rs0 m0 rs m pc. Proof. intros H. - destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps. + destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps. destruct H as [ext lx TAIL NYE COND INV PC]. econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. Qed. @@ -461,7 +579,7 @@ Lemma sistep_preserves_sabort i ge sp rs0 m0 st st': Proof. Local Hint Resolve slocal_set_sreg_preserves_sabort_local slocal_set_smem_preserves_sabort_local: core. unfold sabort. - destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps; intuition eauto. + destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps; intuition eauto. (* COND *) right. unfold sabort_exits in * |- *. destruct H as (ext & H1 & H2). @@ -482,17 +600,17 @@ Proof. Qed. Lemma sistep_correct ge sp i st rs0 m0 rs m: - simatch_local ge sp st.(si_local) rs0 m0 rs m -> + ssem_local ge sp st.(si_local) rs0 m0 rs m -> all_fallthrough ge sp st.(si_exits) rs0 m0 -> - simatch_opt2 ge sp (sistep i st) rs0 m0 (istep ge i sp rs m). + ssem_opt2 ge sp (sistep i st) rs0 m0 (istep ge i sp rs m). Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. + (* Nop *) - unfold simatch, simatch_local; simpl; try_simplify_someHyps. + unfold ssem, ssem_local; simpl; try_simplify_someHyps. + (* Op *) inversion_SOME v; intros OP; simpl. - - unfold simatch, simatch_local; simpl; intuition. + - unfold ssem, ssem_local; simpl; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl; auto. @@ -506,7 +624,7 @@ Proof. + (* LOAD *) admit. (* FIXME *) (* inversion_SOME a0; intros ADD. { inversion_SOME v; intros LOAD; simpl. - - explore_destruct; unfold simatch, simatch_local; simpl; intuition. + - explore_destruct; unfold ssem, ssem_local; simpl; intuition. * exploit REG. try_simplify_someHyps. * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. @@ -525,7 +643,7 @@ Proof. + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. - - unfold simatch, simatch_local; simpl; intuition. + - unfold ssem, ssem_local; simpl; intuition. * congruence. * erewrite seval_list_sval_inj; simpl; auto. erewrite REG. @@ -541,9 +659,9 @@ Proof. erewrite ADD; simpl; auto. } + (* COND *) Local Hint Resolve is_tail_refl: core. - Local Hint Unfold simatch_local: core. + Local Hint Unfold ssem_local: core. inversion_SOME b; intros COND. - { destruct b; simpl; unfold simatch, simatch_local; simpl. + { destruct b; simpl; unfold ssem, ssem_local; simpl. - intros; econstructor; eauto; simpl. unfold seval_condition. erewrite seval_list_sval_inj; simpl; auto. @@ -561,12 +679,12 @@ Proof. Admitted. Lemma sistep_correct_None ge sp i st rs0 m0 rs m: - simatch_local ge sp (st.(si_local)) rs0 m0 rs m -> + ssem_local ge sp (st.(si_local)) rs0 m0 rs m -> sistep i st = None -> istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). - destruct i; simpl; unfold sist_set_local, simatch, simatch_local; simpl; try_simplify_someHyps. + destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps. Qed. (** * Symbolic execution of the internal steps of a path *) @@ -581,12 +699,12 @@ Fixpoint sisteps (path:nat) (f: function) (st: sistate): option sistate := Lemma sisteps_correct_false ge sp path f rs0 m0 st' is: is.(icontinue)=false -> - forall st, simatch ge sp st rs0 m0 is -> + forall st, ssem ge sp st rs0 m0 is -> sisteps path f st = Some st' -> - simatch ge sp st' rs0 m0 is. + ssem ge sp st' rs0 m0 is. Proof. - Local Hint Resolve sistep_preserves_simatch_exits: core. - intros CONT; unfold simatch; rewrite CONT. + Local Hint Resolve sistep_preserves_ssem_exits: core. + intros CONT; unfold ssem; rewrite CONT. induction path; simpl. + unfold sist_set_local; try_simplify_someHyps. + intros st; inversion_SOME i. @@ -627,19 +745,19 @@ Qed. Lemma sisteps_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(icontinue)=true -> - simatch ge sp st rs0 m0 is -> + ssem ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(si_pc) <> None -> - simatch_opt2 ge sp (sisteps path f st) rs0 m0 + ssem_opt2 ge sp (sisteps path f st) rs0 m0 (isteps ge path f sp is.(irs) is.(imem) is.(ipc)) . Proof. Local Hint Resolve sisteps_correct_false sisteps_preserves_sabort sisteps_WF: core. induction path; simpl. + intros st is CONT INV WF; - unfold simatch, sist_set_local in * |- *; + unfold ssem, sist_set_local in * |- *; try_simplify_someHyps. simpl. destruct is; simpl in * |- *; subst; intuition auto. - + intros st is CONT; unfold simatch at 1; rewrite CONT. + + intros st is CONT; unfold ssem at 1; rewrite CONT. intros (LOCAL & PC & NYE) WF. rewrite <- PC. inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. @@ -751,11 +869,11 @@ Record sstate := { internal:> sistate; final: sfval }. Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | smatch_early is: is.(icontinue) = false -> - simatch ge sp st rs0 m0 is -> + ssem ge sp st rs0 m0 is -> smatch pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) | smatch_normal is t s: is.(icontinue) = true -> - simatch ge sp st rs0 m0 is -> + ssem ge sp st rs0 m0 is -> sfmatch pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> smatch pge ge sp st stack f rs0 m0 t s . @@ -776,7 +894,7 @@ Definition sfstep (i: instruction) (prev: sistate_local): sfval := Lemma sfstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> - simatch_local ge sp (si_local st) rs0 m0 rs m -> + ssem_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> sistep i st = None -> sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s. @@ -797,7 +915,7 @@ Admitted. Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> - simatch_local ge sp (si_local st) rs0 m0 rs m -> + ssem_local ge sp (si_local st) rs0 m0 rs m -> sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s -> sistep i st = None -> path_last_step ge pge stack f sp pc rs m t s. @@ -826,9 +944,9 @@ Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}. -Lemma init_simatch ge sp pc rs m: simatch ge sp (init_sistate pc) rs m (mk_istate true pc rs m). +Lemma init_ssem ge sp pc rs m: ssem ge sp (init_sistate pc) rs m (mk_istate true pc rs m). Proof. - unfold simatch, simatch_local, all_fallthrough; simpl. intuition. + unfold ssem, ssem_local, all_fallthrough; simpl. intuition. Qed. Definition sstep (f: function) (pc:node): option sstate := @@ -869,7 +987,7 @@ Theorem sstep_correct f pc pge ge sp path stack rs m t s: path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. - Local Hint Resolve init_simatch sistep_preserves_simatch_exits: core. + Local Hint Resolve init_ssem sistep_preserves_ssem_exits: core. intros PATH STEP; unfold sstep; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. exploit (sisteps_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. @@ -877,9 +995,9 @@ Proof. (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 simatch_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); + (rewrite STEPS; unfold ssem_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); (* intro SEM *) - (simpl; unfold simatch; simpl; rewrite CONT; intro SEM); + (simpl; unfold ssem; simpl; rewrite CONT; intro SEM); (* intro Hi' *) ( assert (Hi': (fn_code f) ! (si_pc st) = Some i); [ unfold nth_default_succ_inst in Hi; @@ -890,7 +1008,7 @@ Proof. (eexists; constructor; eauto). - (* early *) eapply smatch_early; eauto. - unfold simatch; simpl; rewrite CONT. + unfold ssem; simpl; rewrite CONT. destruct (sistep i st) as [st'|] eqn: Hst'; simpl; eauto. - destruct SEM as (SEM & PC & HNYE). destruct (sistep i st) as [st'|] eqn: Hst'; simpl. @@ -904,7 +1022,7 @@ Proof. intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT. { (* icontinue mk_istate = true *) eapply smatch_normal; simpl; eauto. - unfold simatch in SEM. + unfold ssem in SEM. rewrite CONT in SEM. destruct SEM as (SEM & PC & HNYE). rewrite <- PC. @@ -912,7 +1030,7 @@ Proof. { eapply smatch_early; eauto. } + (* normal non-Snone instruction *) eapply smatch_normal; eauto. - * unfold simatch; simpl; rewrite CONT; intuition. + * unfold ssem; simpl; rewrite CONT; intuition. * simpl. eapply sfstep_correct; eauto. rewrite PC; auto. Qed. @@ -1005,7 +1123,7 @@ Theorem sstep_exact f pc pge ge sp path stack st 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_simatch: core. + Local Hint Resolve init_ssem: core. unfold sstep; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. exploit (sisteps_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. @@ -1026,16 +1144,16 @@ Proof. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. - exploit simatch_opt_determ; eauto. + exploit ssem_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 simatch_opt_determ; eauto. + exploit ssem_opt_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - unfold simatch in SEM1. + unfold ssem in SEM1. rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). exploit sfmatch_equiv; eauto. clear SEM2; destruct 1 as (s' & Ms' & SEM2). @@ -1045,7 +1163,7 @@ Proof. eapply exec_normal_exit; eauto. eapply sfstep_complete; eauto. * congruence. - * unfold simatch_local in * |- *. intuition try congruence. + * unfold ssem_local in * |- *. intuition try congruence. Admitted. (** * Simulation of RTLpath code w.r.t symbolic execution *) @@ -1131,8 +1249,8 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> - forall rs m is1, simatch ge1 sp st1 rs m is1 -> - exists is2, simatch ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. + forall rs m is1, ssem ge1 sp st1 rs m is1 -> + exists is2, ssem ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := -- cgit From 9a658fc6c1aca57e5cfbabc3ef3c1d0ee4667051 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 28 Apr 2020 16:58:51 +0200 Subject: Proof of ssem_exit_exclude_sabort - on avance ! --- mppa_k1c/lib/RTLpathSE_theory.v | 52 +++++++++-------------------------------- 1 file changed, 11 insertions(+), 41 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 9978b3f3..fae3c557 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -136,9 +136,11 @@ Definition ssem_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) ( /\ 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 := - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = None - \/ sabort_local ge sp ext.(si_elocal) rs m. + 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 }. @@ -394,39 +396,6 @@ Proof. - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; eauto. Qed. -(* Lemma ssem_exit_exclude_sabort_local ge sp st rs m rs' m' pc': - ssem_exit ge sp (si_exit st) rs m rs' m' pc' -> - all_fallthrough ge sp (si_exits st) rs m -> - sabort_local ge sp (si_local st) rs m -> - False. -Proof. - intros SIML ALLF ABORT. - inv SIML. - exploit ALLF; eauto. - congruence. -Qed. *) - -(* Lemma ssem_exits_exclude_sabort_exits ge sp exits rs m rs' m' pc': - ssem_exits ge sp exits rs m rs' m' pc' -> - sabort_exits ge sp exits rs m -> - False. -Proof. - intros SIML ABORT. - inv SIML. - inv ABORT. -Admitted. *) - -(* Lemma ssem_exits_exclude_sabort ge sp st rs m rs' m' pc': - ssem_exits ge sp (si_exits st) rs m rs' m' pc' -> - sabort ge sp st rs m -> - False. -Proof. - intros SIME ABORT. - inv ABORT. - - intuition. eapply ssem_exits_exclude_sabort_local; eauto. - - eapply ssem_exits_exclude_sabort_exits; eauto. -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 -> @@ -442,7 +411,7 @@ Lemma ssem_exit_exclude_sabort_exit ge sp ext rs m rs' m' pc': Proof. intros A B. destruct A as (A & A' & A''). inv B. - congruence. - - eapply ssem_local_exclude_sabort_local; eauto. + - 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': @@ -453,16 +422,17 @@ Lemma ssem_exit_exclude_sabort ge sp ext st lx rs m rs' m' pc': Proof. intros SSEM ALLFU ABORT. inv ABORT. - - admit. (* FIXME - we can have a sabort_local ; we don't have any ssem_local on si_local st ! *) + - 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_cons_eq ext' ext lx' lx); eauto. intro. inv H. + eapply ssem_exit_exclude_sabort_exit; eauto. - + admit. - (* we can prove we cannot abort on the seval_condition because ext' is in lx, and all_fallthrough lx - However, we run into the same issue as above for si_elocal *) -Admitted. + + eapply ALLFU2 in H0. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence. +Qed. Lemma ssem_exclude_sabort ge sp st rs m is: sabort ge sp st rs m -> -- cgit From 33c5133df84c9476d928076465bc2d30ef2d13e4 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 28 Apr 2020 18:58:39 +0200 Subject: Proving admitted ssem_exit_fallthrough_upto_exit -> need to prove is_tail_exists --- mppa_k1c/lib/RTLpathSE_theory.v | 68 +++++++++++++++++++++++------------------ 1 file changed, 38 insertions(+), 30 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index fae3c557..07cdc65d 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -304,13 +304,15 @@ Proof. Qed. (* Actually this lemma isn't used *) -Remark is_tail_cons_eq {A: Type} (a a': A) l l': +Remark is_tail_inv_left {A: Type} (a a': A) l l': is_tail (a::l) (a'::l') -> - a = a' \/ (In a l'). + (a = a' /\ l = l') \/ (In a l' /\ is_tail l (a'::l')). Proof. intros. inv H. - - left. reflexivity. - - right. eapply is_tail_in; eauto. + - left. eauto. + - right. econstructor. + + eapply is_tail_in; eauto. + + eapply is_tail_cons_left; eauto. Qed. Lemma ssem_determ ge sp st rs m is1 is2: @@ -344,31 +346,6 @@ Proof. inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. Qed. -(* Lemma ssem_local_exclude_sabort_exits 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_exits ge sp (si_exits st) rs m -> - False. -Proof. - intros SIML ALLF ABORT. inv ABORT. - - exploit ALLF; eauto. congruence. - - exploit ALLF; eauto. congruence. -Qed. *) - -(* Lemma all_fallthrough_exclude_sabort_exit ge sp ext lx rs m: - In ext lx -> - all_fallthrough ge sp lx rs m -> - sabort_exit ge sp ext rs m -> - False. -Proof. - intros INL ALLF ABORT. inv ABORT. - - unfold all_fallthrough in ALLF. eapply ALLF in INL. congruence. - - unfold all_fallthrough in ALLF. eapply ALLF in INL. unfold sabort_local in H. congruence. - intros SIML ALLF ABORT. inv ABORT. - - exploit ALLF; eauto. congruence. - - exploit ALLF; eauto. congruence. -Qed. *) - Lemma ssem_local_exclude_sabort_exit ge sp st ext lx rs m rs' m': ssem_local ge sp (si_local st) rs m rs' m' -> all_fallthrough ge sp (si_exits st) rs m -> @@ -396,13 +373,44 @@ Proof. - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; eauto. Qed. +Lemma is_tail_exists {A: Type}: + forall l l' (a' a: A), + In a' l -> + is_tail (a'::l') (a::l) -> + exists x, is_tail (x::a'::l') (a::l). +Proof. +Admitted. (* should be provable by inference on l ? hopefully *) + + +Remark 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 -> + (ext = ext' /\ lx = lx') \/ exists i, is_tail (i::ext'::lx') (ext::lx). +Proof. + intros EXIT ALLFU ALLFU'. + destruct EXIT as (COND & _ & _). + destruct ALLFU as (TAIL & ALLFU). + destruct ALLFU' as (TAIL' & ALLFU'). + destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto. + - destruct (is_tail_inv_left ext ext' lx lx') as [EQ|TAIL2]; eauto. + destruct TAIL2 as (INEXT & _). eapply ALLFU' in INEXT. congruence. + - destruct (is_tail_inv_left ext' ext lx' lx) as [EQ|TAIL2]; eauto. + + left. intuition eauto. + + right. destruct TAIL2 as (INN & TAIL2). eapply is_tail_exists; eauto. +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. -Admitted. + intros. exploit ssem_exit_fallthrough_upto_exit'; eauto. intro. + destruct H2 as [(A & B) | (i & TAIL)]. + - subst. econstructor; eauto. + - eapply is_tail_cons_left; eauto. +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' -> -- cgit From ad2fe7658be2f4e3d64ef56f26eccaaedd51b1a3 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 29 Apr 2020 13:08:53 +0200 Subject: The lemmas i was basing my theorems on are actually wrong.. --- mppa_k1c/lib/RTLpathSE_theory.v | 87 ++++++++++++++++++----------------------- 1 file changed, 39 insertions(+), 48 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 07cdc65d..495b544e 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -106,26 +106,7 @@ 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. -(* TOO WEAK DEFINITION ! -Definition sabort_exits ge sp (lx: list sistate_exit) rs0 m0: Prop := - exists ext, List.In ext lx - /\ seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = None. -*) - -(* Inductive sabort_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) : Prop := - sabort_exits_on_cond: forall ext lx', - is_tail (ext::lx') lx -> - all_fallthrough ge sp lx' rs m -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = None -> - sabort_exits ge sp lx rs m - | sabort_exits_after_cond: forall ext lx', - is_tail (ext::lx') lx -> - all_fallthrough ge sp lx' rs m -> - seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true -> - sabort_local ge sp (si_elocal ext) rs m -> - sabort_exits ge sp lx rs m. *) - -(* Semantic of an exit in pseudo code: +(** Semantic of an exit in pseudo code: if si_cond (si_condargs) si_elocal; goto if_so else () @@ -148,7 +129,7 @@ Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_ 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: +(** Semantic of a sistate in pseudo code: si_exit1; si_exit2; ...; si_exitn; si_local; goto si_pc *) @@ -164,25 +145,12 @@ Definition ssem (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: 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. -(* Inductive ssem_exits (ge: RTL.genv) (sp: val) (lx: list sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop := - ssem_exits_intro: forall ext lx', - is_tail (ext::lx') lx -> - all_fallthrough ge sp lx' rs m -> - 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' -> - ssem_exits ge sp lx rs m rs' m' pc'. *) - 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 sabort (ge: RTL.genv) (sp:val) (st: sistate) (rs0: regset) (m0: mem): Prop := - (all_fallthrough ge sp st.(si_exits) rs0 m0 /\ sabort_local ge sp st.(si_local) rs0 m0) - \/ sabort_exits ge sp st.(si_exits) rs0 m0. *) - Definition ssem_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with | Some is => ssem ge sp st rs0 m0 is @@ -207,15 +175,6 @@ Definition istate_eq ist1 ist2 := (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\ ist1.(imem) = ist2.(imem). -(* TODO: is it useful -Lemma has_not_yet_exit_cons_split ge sp ext lx rs0 m0: - all_fallthrough ge sp (ext::lx) rs0 m0 <-> - (seval_condition ge sp ext.(the_cond) ext.(cond_args) ext.(exit_ist).(si_smem) rs0 m0 = Some false /\ all_fallthrough ge sp lx rs0 m0). -Proof. - unfold all_fallthrough; simpl; intuition (subst; auto). -Qed. -*) - 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 -> @@ -303,7 +262,6 @@ Proof. destruct (ssem_local_determ ge sp (si_elocal ext) rs0 m0 rs1 m1 rs2 m2); auto. Qed. -(* Actually this lemma isn't used *) 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')). @@ -373,14 +331,43 @@ Proof. - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; eauto. Qed. +Lemma is_tail_revcons_left {A: Type}: + forall l l' (a: A), + ~ (In a l') -> + is_tail l' (a::l) -> + is_tail (a::l') (a::l). +Proof. + induction l. + - intros l' a NOTIN TAIL. inv TAIL. + + contradict NOTIN. econstructor; eauto. + + inv H1. econstructor; eauto. + - intros l' b NOTIN TAIL. inv TAIL. + + contradict NOTIN. repeat (econstructor; eauto). + + +Abort. + Lemma is_tail_exists {A: Type}: forall l l' (a' a: A), In a' l -> is_tail (a'::l') (a::l) -> - exists x, is_tail (x::a'::l') (a::l). + (a' = a /\ l' = l) + \/ (exists x, In x (a::l) /\ is_tail (x::a'::l') (a::l)). Proof. -Admitted. (* should be provable by inference on l ? hopefully *) - + induction l. + - intros. inv H. + - intros l' a' b INH TAIL. destruct INH as [EQA|INH]; subst. + + inv TAIL. + ++ left. econstructor; eauto. + ++ admit. + + inv TAIL. + ++ left. repeat (econstructor; eauto). + ++ eapply IHl in INH; eauto. inv INH. + +++ inv H. right. exists b. constructor. + * repeat (econstructor; eauto). + * econstructor; eauto. + +++ right. destruct H as (x & INX & TAIL). exists x. + constructor. eapply in_cons; eauto. eapply is_tail_cons; eauto. +Abort. Remark 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' -> @@ -397,7 +384,11 @@ Proof. destruct TAIL2 as (INEXT & _). eapply ALLFU' in INEXT. congruence. - destruct (is_tail_inv_left ext' ext lx' lx) as [EQ|TAIL2]; eauto. + left. intuition eauto. - + right. destruct TAIL2 as (INN & TAIL2). eapply is_tail_exists; eauto. + + right. destruct TAIL2 as (INN & TAIL2). + exploit (is_tail_exists lx); eauto. + intros H'. inv H'. + * inv H0. eapply ALLFU in INN. congruence. (* absurd case *) + * destruct H0 as (x & _ & TAIL3). eauto. Qed. Lemma ssem_exit_fallthrough_upto_exit ge sp ext ext' lx lx' exits rs m rs' m' pc': -- cgit From 0e0e169bb8c78da4965c3c301f1d5f5eb58a87b6 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 29 Apr 2020 13:44:34 +0200 Subject: ssem_exit_fallthrough_upto_exit finally done - it was very simple actually --- mppa_k1c/lib/RTLpathSE_theory.v | 71 ++++------------------------------------- 1 file changed, 7 insertions(+), 64 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 495b544e..8ea4f9a7 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -331,76 +331,19 @@ Proof. - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; eauto. Qed. -Lemma is_tail_revcons_left {A: Type}: - forall l l' (a: A), - ~ (In a l') -> - is_tail l' (a::l) -> - is_tail (a::l') (a::l). -Proof. - induction l. - - intros l' a NOTIN TAIL. inv TAIL. - + contradict NOTIN. econstructor; eauto. - + inv H1. econstructor; eauto. - - intros l' b NOTIN TAIL. inv TAIL. - + contradict NOTIN. repeat (econstructor; eauto). - + -Abort. - -Lemma is_tail_exists {A: Type}: - forall l l' (a' a: A), - In a' l -> - is_tail (a'::l') (a::l) -> - (a' = a /\ l' = l) - \/ (exists x, In x (a::l) /\ is_tail (x::a'::l') (a::l)). -Proof. - induction l. - - intros. inv H. - - intros l' a' b INH TAIL. destruct INH as [EQA|INH]; subst. - + inv TAIL. - ++ left. econstructor; eauto. - ++ admit. - + inv TAIL. - ++ left. repeat (econstructor; eauto). - ++ eapply IHl in INH; eauto. inv INH. - +++ inv H. right. exists b. constructor. - * repeat (econstructor; eauto). - * econstructor; eauto. - +++ right. destruct H as (x & INX & TAIL). exists x. - constructor. eapply in_cons; eauto. eapply is_tail_cons; eauto. -Abort. - -Remark 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 -> - (ext = ext' /\ lx = lx') \/ exists i, is_tail (i::ext'::lx') (ext::lx). -Proof. - intros EXIT ALLFU ALLFU'. - destruct EXIT as (COND & _ & _). - destruct ALLFU as (TAIL & ALLFU). - destruct ALLFU' as (TAIL' & ALLFU'). - destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto. - - destruct (is_tail_inv_left ext ext' lx lx') as [EQ|TAIL2]; eauto. - destruct TAIL2 as (INEXT & _). eapply ALLFU' in INEXT. congruence. - - destruct (is_tail_inv_left ext' ext lx' lx) as [EQ|TAIL2]; eauto. - + left. intuition eauto. - + right. destruct TAIL2 as (INN & TAIL2). - exploit (is_tail_exists lx); eauto. - intros H'. inv H'. - * inv H0. eapply ALLFU in INN. congruence. (* absurd case *) - * destruct H0 as (x & _ & TAIL3). eauto. -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. exploit ssem_exit_fallthrough_upto_exit'; eauto. intro. - destruct H2 as [(A & B) | (i & TAIL)]. - - subst. econstructor; eauto. - - eapply is_tail_cons_left; eauto. + 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': -- cgit From 0fdf888d455526cf23374c26a961d6e06a582c4d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 29 Apr 2020 15:02:12 +0200 Subject: Complete proof of ssem_exclude_sabort --- mppa_k1c/lib/RTLpathSE_theory.v | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 8ea4f9a7..f5a19b27 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -317,7 +317,7 @@ Proof. - (* FIXME Problem : if we have a ssem_local, this means we ONLY evaluated the conditions, but we NEVER actually evaluated the si_elocal from the sistate_exit ! So we cannot prove a lack of abort on the si_elocal.. We must change the definitions *) -Admitted. +Abort. Lemma ssem_local_exclude_sabort ge sp st rs m rs' m': ssem_local ge sp (si_local st) rs m rs' m' -> @@ -328,7 +328,10 @@ Proof. intros SIML ALLF ABORT. inv ABORT. - intuition; eapply ssem_local_exclude_sabort_local; eauto. - - destruct H as (ext & lx & ALLFU & SABORT). inv ALLFU. eapply ssem_local_exclude_sabort_exit; 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': @@ -371,9 +374,9 @@ Proof. - 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_cons_eq ext' ext lx' lx); eauto. intro. inv H. - + eapply ssem_exit_exclude_sabort_exit; eauto. - + eapply ALLFU2 in H0. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence. + 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_exclude_sabort ge sp st rs m is: -- cgit From 88cc35dd7c522d68bd1c5d6a35da7f238191a964 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 29 Apr 2020 17:45:03 +0200 Subject: Proof of sistep_preserves_sabort --- mppa_k1c/lib/RTLpathSE_theory.v | 82 +++++++++++++++++++++++++++++++---------- 1 file changed, 63 insertions(+), 19 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index f5a19b27..ab3f2d3c 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -460,17 +460,6 @@ Qed. (* TODO - continue updating the rest *) -Lemma sistep_preserves_ssem_exits i ge sp st rs0 m0 rs m pc st': - ssem_exits ge sp st.(si_exits) rs0 m0 rs m pc -> - sistep i st = Some st' -> - ssem_exits ge sp st'.(si_exits) rs0 m0 rs m pc. -Proof. - intros H. - destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps. - destruct H as [ext lx TAIL NYE COND INV PC]. - econstructor; try (eapply is_tail_cons; eapply TAIL); eauto. -Qed. - 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. @@ -488,17 +477,72 @@ Proof. unfold sabort_local. simpl; intuition. Qed. -Lemma sistep_preserves_sabort i ge sp rs0 m0 st st': +Lemma sist_set_local_preserves_sabort ge sp rs0 m0 st st' loc' pc': + sist_set_local st pc' loc' = Some st' -> + sabort ge sp st rs0 m0 -> + sabort ge sp st' rs0 m0. +Proof. + intros SSL ABORT. inv SSL. unfold sabort. simpl. + destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. + - +Abort. + +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 sistep_preserves_sabort i ge sp rs m st st': sistep i st = Some st' -> - sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0. + sabort ge sp st rs m -> sabort ge sp st' rs m. Proof. - Local Hint Resolve slocal_set_sreg_preserves_sabort_local slocal_set_smem_preserves_sabort_local: core. - unfold sabort. - destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps; intuition eauto. + 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 *) - right. unfold sabort_exits in * |- *. - destruct H as (ext & H1 & H2). - simpl; exists ext; intuition. + * 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 sistep_WF i st: -- cgit From bb64e419eb857882a6f0d102db50ca93b181ad11 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 30 Apr 2020 12:54:53 +0200 Subject: Back to a proven sistep_correct (except for the LOAD case) --- mppa_k1c/lib/RTLpathSE_theory.v | 45 ++++++++++++++++++++++++----------------- 1 file changed, 26 insertions(+), 19 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index ab3f2d3c..ce5e09d7 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -95,7 +95,8 @@ Definition sabort_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset \/ exists (r: reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = None. (* Syntax and semantics of symbolic exit states *) -Record sistate_exit := { si_cond: condition; si_scondargs: list_sval; si_elocal: sistate_local; si_ifso: node }. +Record sistate_exit := mk_sistate_exit + { si_cond: condition; si_scondargs: list_sval; si_elocal: sistate_local; si_ifso: node }. Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool := SOME args <- seval_list_sval ge sp lsv rs0 m0 IN @@ -502,7 +503,7 @@ Proof. intros. unfold all_fallthrough in *. intros. inv H1; eauto. Qed. - + Lemma sistep_preserves_sabort i ge sp rs m st st': sistep i st = Some st' -> sabort ge sp st rs m -> sabort ge sp st' rs m. @@ -566,18 +567,21 @@ Proof. intros (PRE & MEM & REG) NYE. destruct i; simpl; auto. + (* Nop *) - unfold ssem, ssem_local; simpl; try_simplify_someHyps. + constructor; [|constructor]; simpl; auto. + constructor; auto. + (* Op *) inversion_SOME v; intros OP; simpl. - - unfold ssem, ssem_local; simpl; intuition. - * exploit REG. try_simplify_someHyps. - * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl; auto. + - 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 seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - - unfold sabort, sabort_local; simpl. - left. right. right. - exists r. destruct (Pos.eq_dec r r); try congruence. + - left. constructor; simpl; auto. + unfold sabort_local. right. right. + simpl. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. + (* LOAD *) admit. (* FIXME *) @@ -608,12 +612,12 @@ Proof. erewrite REG. try_simplify_someHyps. - unfold sabort, sabort_local; simpl. - left. right. left. + left. constructor; auto. right. left. erewrite seval_list_sval_inj; simpl; auto. erewrite REG. try_simplify_someHyps. } { unfold sabort, sabort_local; simpl. - left. right. left. + left. constructor; auto. right. left. erewrite seval_list_sval_inj; simpl; auto. erewrite ADD; simpl; auto. } + (* COND *) @@ -621,8 +625,9 @@ Proof. Local Hint Unfold ssem_local: core. inversion_SOME b; intros COND. { destruct b; simpl; unfold ssem, ssem_local; simpl. - - intros; econstructor; eauto; simpl. - unfold seval_condition. + - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st). + constructor; constructor; subst; simpl; auto. + unfold seval_condition. subst; simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - intuition. unfold all_fallthrough in * |- *. simpl. @@ -630,11 +635,13 @@ Proof. unfold seval_condition. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } - { unfold sabort, sabort_exits; simpl. - right. eexists; intuition eauto. simpl. - unfold seval_condition. - erewrite seval_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 seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. } Admitted. Lemma sistep_correct_None ge sp i st rs0 m0 rs m: -- cgit From 0dc240cc53b3603dc76419bebfbbadffa3958988 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 30 Apr 2020 13:25:56 +0200 Subject: Progression on LOAD sistep_correct --- mppa_k1c/lib/RTLpathSE_theory.v | 33 ++++++++++++++++++++++++++++++--- 1 file changed, 30 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index ce5e09d7..4d3373a9 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -584,7 +584,34 @@ Proof. simpl. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - + (* LOAD *) admit. (* FIXME *) + + (* LOAD *) + inversion_SOME a0; intro ADD. + { inversion_SOME v; intros LOAD; simpl. + - explore_destruct; unfold ssem, ssem_local; simpl; intuition. + * unfold ssem. 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 seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + * unfold ssem. 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 seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. + - 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 seval_list_sval_inj; simpl; auto. + rewrite ADD; simpl; auto. try_simplify_someHyps. + * unfold ssem. 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 seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. admit. (* FIXME - need to add default_notrap_load_value *) + } { admit. } (* inversion_SOME a0; intros ADD. { inversion_SOME v; intros LOAD; simpl. - explore_destruct; unfold ssem, ssem_local; simpl; intuition. @@ -603,7 +630,7 @@ Proof. exists r. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. rewrite ADD; simpl; auto. } *) - + (* STORE *) + + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. - unfold ssem, ssem_local; simpl; intuition. @@ -620,7 +647,7 @@ Proof. left. constructor; auto. right. left. erewrite seval_list_sval_inj; simpl; auto. erewrite ADD; simpl; auto. } - + (* COND *) + + (* COND *) Local Hint Resolve is_tail_refl: core. Local Hint Unfold ssem_local: core. inversion_SOME b; intros COND. -- cgit From 0a853a484575b60a828307198c5b796475b24d74 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 30 Apr 2020 16:11:13 +0200 Subject: Progress on load_notrap_value in symbolic execution --- mppa_k1c/lib/RTLpathSE_theory.v | 39 +++++++++++++++++++++++++++------------ 1 file changed, 27 insertions(+), 12 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 4d3373a9..b5af6358 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -25,7 +25,7 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ Inductive sval := | Sinput (r: reg) | Sop (op:operation) (lsv: list_sval) (sm: smem) (* TODO for the implementation: add an additionnal case without dependency on sm*) - | Sload (sm: smem) (chunk:memory_chunk) (addr:addressing) (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) @@ -46,7 +46,6 @@ Fixpoint list_sval_inj (l: list sval): list_sval := Local Open Scope option_monad_scope. -(* FIXME - add non trapping load eval *) Fixpoint seval_sval (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val := match sv with | Sinput r => Some (rs0#r) @@ -54,11 +53,25 @@ Fixpoint seval_sval (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): SOME args <- seval_list_sval ge sp l rs0 m0 IN SOME m <- seval_smem ge sp sm rs0 m0 IN eval_operation ge sp op args m - | Sload sm chunk addr lsv => - SOME args <- seval_list_sval ge sp lsv rs0 m0 IN - SOME a <- eval_addressing ge sp addr args IN - SOME m <- seval_smem ge sp sm rs0 m0 IN - Mem.loadv chunk m a + | Sload sm trap chunk addr lsv => + match trap with + | TRAP => + SOME args <- seval_list_sval ge sp lsv rs0 m0 IN + SOME a <- eval_addressing ge sp addr args IN + SOME m <- seval_smem ge sp sm rs0 m0 IN + Mem.loadv chunk m a + | NOTRAP => + SOME args <- seval_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 <- seval_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 seval_list_sval (ge: RTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) := match lsv with @@ -429,10 +442,10 @@ Definition sistep (i: instruction) (st: sistate): option sistate := 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 sist_set_local st pc' next - | Iload _ chunk addr args dst pc' => + | 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) chunk addr vargs) in + let next := slocal_set_sreg prev dst (Sload prev.(si_smem) trap chunk addr vargs) in sist_set_local st pc' next | Istore chunk addr args src pc' => let prev := st.(si_local) in @@ -598,8 +611,10 @@ Proof. constructor; constructor; simpl; auto. congruence. intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. - erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. + inversion_SOME args; intros ARGS. 2: admit. (* TODO - continue *) + admit. +(* erewrite seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. *) - 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. @@ -610,7 +625,7 @@ Proof. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. admit. (* FIXME - need to add default_notrap_load_value *) + try_simplify_someHyps. admit. } { admit. } (* inversion_SOME a0; intros ADD. { inversion_SOME v; intros LOAD; simpl. -- cgit From 8bf7464e0293f514b6a8b4cb7a0d038be6520d1c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 19 May 2020 12:24:47 +0200 Subject: Avancement on sistep_correct for load --- mppa_k1c/lib/RTLpathSE_theory.v | 26 ++++++++++++++++++++++---- 1 file changed, 22 insertions(+), 4 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index b5af6358..7fe212ab 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -572,6 +572,23 @@ Proof. intro H; inversion_clear H; simpl; auto. Qed. + +Lemma seval_list_sval_inj_not_none ge sp st rs0 m0 rs: forall l, + (forall r, List.In r l -> seval_sval ge sp (si_sreg (si_local st) r) rs0 m0 = Some rs # r) -> + seval_list_sval ge sp (list_sval_inj (map (si_sreg (si_local st)) l)) rs0 m0 = None -> + False. +Proof. + induction l. + - intros. simpl in H0. discriminate. + - intros ALLR. simpl. + inversion_SOME v. + + intro SVAL. inversion_SOME lv; [discriminate|]. + assert (forall r : reg, In r l -> seval_sval ge sp (si_sreg (si_local st) r) rs0 m0 = Some rs # r). + { intros r INR. eapply ALLR. right. assumption. } + intro SVALLIST. intro. eapply IHl; eauto. + + intros. exploit (ALLR a); [constructor; eauto|]. congruence. +Qed. + Lemma sistep_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 -> @@ -611,10 +628,10 @@ Proof. 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: admit. (* TODO - continue *) - admit. -(* erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. *) + inversion_SOME args; intros ARGS; [|exploit seval_list_sval_inj_not_none; eauto; contradiction]. + exploit seval_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. @@ -686,6 +703,7 @@ Proof. try_simplify_someHyps. } Admitted. + Lemma sistep_correct_None ge sp i st rs0 m0 rs m: ssem_local ge sp (st.(si_local)) rs0 m0 rs m -> sistep i st = None -> -- cgit From 1df2759e91263fb494133e668b6db6f881e2784e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 19 May 2020 15:32:02 +0200 Subject: sistep_correct proved completely --- mppa_k1c/lib/RTLpathSE_theory.v | 37 ++++++++++++++++--------------------- 1 file changed, 16 insertions(+), 21 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 7fe212ab..2a8c4e3f 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -642,26 +642,21 @@ Proof. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. subst; rewrite Regmap.gss; simpl. erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. admit. - } { admit. } -(* inversion_SOME a0; intros ADD. - { inversion_SOME v; intros LOAD; simpl. - - explore_destruct; unfold ssem, ssem_local; simpl; intuition. - * exploit REG. try_simplify_someHyps. - * destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto]. - subst; rewrite Regmap.gss; simpl. - erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. - - unfold sabort, sabort_local; simpl. - left. right. right. - exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite seval_list_sval_inj; simpl; auto. - try_simplify_someHyps. } - { unfold sabort, sabort_local; simpl. - left. right. right. - exists r. destruct (Pos.eq_dec r r); try congruence. - simpl. erewrite seval_list_sval_inj; simpl; auto. - rewrite ADD; simpl; auto. } *) + try_simplify_someHyps. intros SMEM SEVAL. + rewrite LOAD. reflexivity. + } { 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 seval_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 seval_list_sval_inj; simpl; auto. + try_simplify_someHyps. intro SMEM. rewrite ADD. reflexivity. + } + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. @@ -701,7 +696,7 @@ Proof. unfold seval_condition. erewrite seval_list_sval_inj; simpl; auto. try_simplify_someHyps. } -Admitted. +Qed. Lemma sistep_correct_None ge sp i st rs0 m0 rs m: -- cgit From 11e2b9918f66fe1dd46669b6ece6ee29b53a627e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 19 May 2020 17:26:19 +0200 Subject: Proof of sisteps_correct_false --- mppa_k1c/lib/RTLpathSE_theory.v | 39 +++++++++++++++++++++++++++++++-------- 1 file changed, 31 insertions(+), 8 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 2a8c4e3f..481c43da 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -432,7 +432,7 @@ Definition slocal_set_smem (st:sistate_local) (sm:smem) := Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option sistate := Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. -(* FIXME - add notrap *) + Definition sistep (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => @@ -718,18 +718,41 @@ Fixpoint sisteps (path:nat) (f: function) (st: sistate): option sistate := sisteps p f st1 end. -Lemma sisteps_correct_false ge sp path f rs0 m0 st' is: +Lemma sist_set_local_preserves_si_exits st pc loc st': + sist_set_local st pc loc = Some st' -> + si_exits st' = si_exits st. +Proof. + intros. inv H. simpl. reflexivity. +Qed. + +Lemma sistep_preserves_allfu ge sp ext lx rs0 m0 st st' i: + all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs0 m0 -> + sistep 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; try discriminate; simpl in SISTEP; try (erewrite sist_set_local_preserves_si_exits; eauto; fail). + inv SISTEP. simpl. constructor. assumption. +Qed. + +Lemma sisteps_correct_false ge sp f rs0 m0 st' is: + forall path, is.(icontinue)=false -> - forall st, ssem ge sp st rs0 m0 is -> + forall st, ssem ge sp st rs0 m0 is -> sisteps path f st = Some st' -> ssem ge sp st' rs0 m0 is. Proof. - Local Hint Resolve sistep_preserves_ssem_exits: core. - intros CONT; unfold ssem; rewrite CONT. induction path; simpl. - + unfold sist_set_local; try_simplify_someHyps. - + intros st; inversion_SOME i. - inversion_SOME st1; eauto. + - intros. congruence. + - intros ICF st SSEM STEQ'. + destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate]. + destruct (sistep _ _) eqn:SISTEP; [|discriminate]. + eapply IHpath. 3: eapply STEQ'. eauto. + unfold ssem in SSEM. rewrite ICF in SSEM. + destruct SSEM as (ext & lx & SEXIT & ALLFU). + unfold ssem. rewrite ICF. exists ext, lx. + constructor; auto. eapply sistep_preserves_allfu; eauto. Qed. Lemma sisteps_preserves_sabort ge sp path f rs0 m0 st': forall st, -- cgit From 24062db148f868e68406414f151983c34919a6d9 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 12:04:35 +0200 Subject: Stailcall --- mppa_k1c/lib/RTLpathSE_theory.v | 26 +++++++++++++++++++++----- 1 file changed, 21 insertions(+), 5 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index 481c43da..b5c3b26f 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -856,8 +856,8 @@ Inductive sfval := | Snone | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:node) (* NB: [res] the return register is hard-wired ! Is it restrictive ? *) -(* | Stailcall: signature -> sval + ident -> list_sval -> sfval +(* | Sbuiltin: external_function -> list (builtin_arg sval) -> builtin_res sval -> sfval | Sjumptable: sval -> list node -> sfval *) @@ -879,14 +879,22 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac seval_list_sval ge sp lsv rs0 m0 = Some args -> sfmatch pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m) -(* TODO: - | exec_Itailcall stk pc rs m sig ros args fd m': + | exec_Stailcall stk rs m sig svos args fd m' lsv: + sfind_function pge ge sp svos rs0 m0 = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + seval_list_sval ge sp lsv rs0 m0 = Some args -> + sfmatch pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m + E0 (Callstate stack fd args m') +(* + (fn_code f)!pc = Some(Itailcall sig ros args) -> find_function pge ros rs = Some fd -> funsig fd = sig -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m - E0 (Callstate stack fd rs##args m') + E0 (Callstate stack fd rs##args m') *) +(* TODO: | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> eval_builtin_args ge (fun r => rs#r) sp m args vargs -> @@ -929,6 +937,10 @@ Definition sfstep (i: instruction) (prev: sistate_local): sfval := let svos := sum_left_map prev.(si_sreg) ros in let sargs := list_sval_inj (List.map prev.(si_sreg) args) in Scall sig svos sargs res pc + | Itailcall sig ros args => + let svos := sum_left_map prev.(si_sreg) ros in + let sargs := list_sval_inj (List.map prev.(si_sreg) args) in + Stailcall sig svos sargs | Ireturn or => let sor := SOME r <- or IN Some (prev.(si_sreg) r) in Sreturn sor @@ -949,7 +961,11 @@ Proof. - destruct ros; simpl in * |- *; auto. rewrite REG; auto. - erewrite seval_list_sval_inj; simpl; auto. - + (* Itaillcall *) admit. + + (* Itailcall *) intros. eapply exec_Stailcall; auto. + - destruct ros; simpl in * |- *; auto. + rewrite REG; auto. + - eauto. + - erewrite seval_list_sval_inj; simpl; auto. + (* Ibuiltin *) admit. + (* Ijumptable *) admit. + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto. -- cgit From 993a699510328fbfca6a18820f1f008691eb80db Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 15:09:06 +0200 Subject: Admitting the rest --- mppa_k1c/lib/RTLpathSE_theory.v | 10 +++++----- mppa_k1c/lib/RTLpathScheduler.v | 4 ++-- 2 files changed, 7 insertions(+), 7 deletions(-) diff --git a/mppa_k1c/lib/RTLpathSE_theory.v b/mppa_k1c/lib/RTLpathSE_theory.v index b5c3b26f..77a39ee2 100644 --- a/mppa_k1c/lib/RTLpathSE_theory.v +++ b/mppa_k1c/lib/RTLpathSE_theory.v @@ -1047,7 +1047,7 @@ Theorem sstep_correct f pc pge ge sp path stack rs m t s: path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. - Local Hint Resolve init_ssem sistep_preserves_ssem_exits: core. +(* Local Hint Resolve init_ssem sistep_preserves_ssem_exits: core. intros PATH STEP; unfold sstep; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. exploit (sisteps_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. @@ -1092,8 +1092,8 @@ Proof. eapply smatch_normal; eauto. * unfold ssem; simpl; rewrite CONT; intuition. * simpl. eapply sfstep_correct; eauto. - rewrite PC; auto. -Qed. + rewrite PC; auto. *) +Admitted. (* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *) Inductive equiv_stackframe: stackframe -> stackframe -> Prop := @@ -1170,8 +1170,8 @@ Proof. repeat (constructor; auto). - (* Sreturn *) intros; eexists; econstructor. + eapply equiv_state_refl; eauto. - + eapply exec_Sreturn; eauto. -Qed. + + (* eapply exec_Sreturn; eauto. *) admit. +Admitted. (* NB: each execution of a symbolic state (produced from [sstep]) represents a concrete execution (sstep is exact). diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 4478ee72..9cdb6f0f 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -311,11 +311,11 @@ Proof. Qed. Lemma sistate_simu f dupmap sp st st' rs m is: - simatch ge sp st rs m is -> + ssem ge sp st rs m is -> liveness_ok_function f -> sistate_simu f dupmap st st' -> exists is', - simatch tge sp st' rs m is' /\ istate_simu f dupmap is is'. + ssem tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. intros SEM LIVE X; eapply X; eauto. Qed. -- cgit From f7aa5aa7ded628e010a21a2a7a78489868a24f02 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 15:25:27 +0200 Subject: No need to change default_succ : we keep ifnot as default, and adjust the path number if we want to stop earlier --- mppa_k1c/lib/RTLpath.v | 2 -- 1 file changed, 2 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 7134de10..228bc95b 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -43,7 +43,6 @@ Local Open Scope option_monad_scope. (** Internal instruction = instruction with a default successor in a path. *) -(* FIXME - change successor based on prediction *) Definition default_succ (i: instruction): option node := match i with | Inop s => Some s @@ -54,7 +53,6 @@ Definition default_succ (i: instruction): option node := | _ => None (* TODO: we could choose a successor for jumptable ? *) end. -(* FIXME - change early_exit based on prediction *) Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *) match i with | Icond cond args ifso ifnot _ => Some ifso -- cgit From b6a6e7b783cc59f2efd2fc2470f6676d88d3f58f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 16:02:54 +0200 Subject: Modifying RTLpath in order to allow for stopping at Icond --- mppa_k1c/lib/RTLpath.v | 34 +++++++++++++++++++++++----------- 1 file changed, 23 insertions(+), 11 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 228bc95b..4ade0ad1 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -49,7 +49,12 @@ Definition default_succ (i: instruction): option node := | Iop op args res s => Some s | Iload _ chunk addr args dst s => Some s | Istore chunk addr args src s => Some s - | Icond cond args ifso ifnot _ => Some ifnot + | Icond cond args ifso ifnot pred => + match pred with + | None => None + | Some true => None (* TODO rework RTL and Duplicate so that pred is a boolean ? *) + | Some false => Some ifnot + end | _ => None (* TODO: we could choose a successor for jumptable ? *) end. @@ -141,7 +146,6 @@ Coercion program_RTL: program >-> RTL.program. Record istate := mk_istate { icontinue: bool; ipc: node; irs: regset; imem: mem }. -(* FIXME - prediction *) (* Internal step through the path *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with @@ -166,9 +170,14 @@ Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem) SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN Some (mk_istate true pc' rs m') - | Icond cond args ifso ifnot _ => + | Icond cond args ifso ifnot pred => SOME b <- eval_condition cond rs##args m IN - Some (mk_istate (negb b) (if b then ifso else ifnot) rs m) + let nc := match pred with + | None => false + | Some true => false + | Some false => negb b + end + in Some (mk_istate nc (if b then ifso else ifnot) rs m) | _ => None (* TODO jumptable ? *) end. @@ -711,7 +720,7 @@ Proof. Local Hint Resolve Icond_early_exit: core. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. all: explore_destruct; simplify_SOME b; try discriminate. -Qed. +Abort. (* this lemma isn't true anymore *) Section COMPLETENESS. @@ -872,7 +881,9 @@ Proof. econstructor; eauto. * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - - (* Ijumptable *) + - (* Icond *) + admit. + - (* Ijumptable *) intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. @@ -881,11 +892,11 @@ Proof. econstructor; eauto. * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - - (* Ireturn *) + - (* Ireturn *) intros; exploit exec_Ireturn; eauto. eexists; intuition eauto. eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. -Qed. +Admitted. (* TODO - prove Icond *) Lemma path_step_complete stack f sp rs m pc t s1' idx path st: isteps ge (path.(psize)-idx) f sp rs m pc = Some st -> @@ -922,7 +933,8 @@ Proof. { clear RSTEP. exploit isteps_inversion_early; eauto. intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). - exploit istep_early_exit; eauto. + admit. +(* exploit istep_early_exit; eauto. intros (X1 & X2 & EARLY_EXIT). destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. exploit (internal_node_path path0); omega || eauto. @@ -931,7 +943,7 @@ Proof. erewrite isteps_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; - destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. + destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. *) } destruct HPATH0 as (path1 & Hpath1). destruct (path1.(psize)) as [|ps] eqn:Hpath1size. @@ -953,7 +965,7 @@ Proof. - simpl; eauto. - intuition subst. repeat eexists; intuition eauto. -Qed. +Admitted. Lemma step_noninst_complete s1 t s1' s2: is_inst s1 = false -> -- cgit From a0db2354c891c6f27015bd7b05054b33223b41c1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 16:20:59 +0200 Subject: Revert "Modifying RTLpath in order to allow for stopping at Icond" This reverts commit b6a6e7b783cc59f2efd2fc2470f6676d88d3f58f. --- mppa_k1c/lib/RTLpath.v | 34 +++++++++++----------------------- 1 file changed, 11 insertions(+), 23 deletions(-) diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 4ade0ad1..228bc95b 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -49,12 +49,7 @@ Definition default_succ (i: instruction): option node := | Iop op args res s => Some s | Iload _ chunk addr args dst s => Some s | Istore chunk addr args src s => Some s - | Icond cond args ifso ifnot pred => - match pred with - | None => None - | Some true => None (* TODO rework RTL and Duplicate so that pred is a boolean ? *) - | Some false => Some ifnot - end + | Icond cond args ifso ifnot _ => Some ifnot | _ => None (* TODO: we could choose a successor for jumptable ? *) end. @@ -146,6 +141,7 @@ Coercion program_RTL: program >-> RTL.program. Record istate := mk_istate { icontinue: bool; ipc: node; irs: regset; imem: mem }. +(* FIXME - prediction *) (* Internal step through the path *) Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate := match i with @@ -170,14 +166,9 @@ Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem) SOME a <- eval_addressing ge sp addr rs##args IN SOME m' <- Mem.storev chunk m a rs#src IN Some (mk_istate true pc' rs m') - | Icond cond args ifso ifnot pred => + | Icond cond args ifso ifnot _ => SOME b <- eval_condition cond rs##args m IN - let nc := match pred with - | None => false - | Some true => false - | Some false => negb b - end - in Some (mk_istate nc (if b then ifso else ifnot) rs m) + Some (mk_istate (negb b) (if b then ifso else ifnot) rs m) | _ => None (* TODO jumptable ? *) end. @@ -720,7 +711,7 @@ Proof. Local Hint Resolve Icond_early_exit: core. destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence. all: explore_destruct; simplify_SOME b; try discriminate. -Abort. (* this lemma isn't true anymore *) +Qed. Section COMPLETENESS. @@ -881,9 +872,7 @@ Proof. econstructor; eauto. * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - - (* Icond *) - admit. - - (* Ijumptable *) + - (* Ijumptable *) intros; exploit exec_Ijumptable; eauto. eexists; intuition eauto. unfold match_states, match_inst_states; simpl. @@ -892,11 +881,11 @@ Proof. econstructor; eauto. * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega. * simpl; eauto. - - (* Ireturn *) + - (* Ireturn *) intros; exploit exec_Ireturn; eauto. eexists; intuition eauto. eexists O; unfold match_states, match_inst_states; simpl; intuition eauto. -Admitted. (* TODO - prove Icond *) +Qed. Lemma path_step_complete stack f sp rs m pc t s1' idx path st: isteps ge (path.(psize)-idx) f sp rs m pc = Some st -> @@ -933,8 +922,7 @@ Proof. { clear RSTEP. exploit isteps_inversion_early; eauto. intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0). - admit. -(* exploit istep_early_exit; eauto. + exploit istep_early_exit; eauto. intros (X1 & X2 & EARLY_EXIT). destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst. exploit (internal_node_path path0); omega || eauto. @@ -943,7 +931,7 @@ Proof. erewrite isteps_normal_exit in Hi'; eauto. clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros. destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY; - destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. *) + destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto. } destruct HPATH0 as (path1 & Hpath1). destruct (path1.(psize)) as [|ps] eqn:Hpath1size. @@ -965,7 +953,7 @@ Proof. - simpl; eauto. - intuition subst. repeat eexists; intuition eauto. -Admitted. +Qed. Lemma step_noninst_complete s1 t s1' s2: is_inst s1 = false -> -- cgit From 9e1268915a43c9668b0e6dda5426ba47d6b4870b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 May 2020 17:57:54 +0200 Subject: RTL -> RTLpath -> RTL --- driver/Compiler.v | 18 +++++++++++-- mppa_k1c/lib/RTLpath.v | 71 +++++++++++++++++++++++++++++++++++++++++++++++++- 2 files changed, 86 insertions(+), 3 deletions(-) diff --git a/driver/Compiler.v b/driver/Compiler.v index 499feff2..37bb54f7 100644 --- a/driver/Compiler.v +++ b/driver/Compiler.v @@ -39,6 +39,7 @@ Require Tailcall. Require Inlining. Require Renumber. Require Duplicate. +Require RTLpath RTLpathLivegen. Require Constprop. Require CSE. Require ForwardMoves. @@ -152,6 +153,9 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program := @@ print (print_RTL 11) @@@ time "Unused globals" Unusedglob.transform_program @@ print (print_RTL 12) + @@@ time "RTLpath generation" RTLpathLivegen.transf_program + @@ time "RTL projection" RTLpath.program_RTL + @@ print (print_RTL 13) @@@ time "Register allocation" Allocation.transf_program @@ print print_LTL @@ time "Branch tunneling" Tunneling.tunnel_program @@ -263,6 +267,8 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_redundancy Deadcodeproof.match_prog) ::: mkpass (match_if Compopts.all_loads_nontrap Allnontrapproof.match_prog) ::: mkpass Unusedglobproof.match_prog + ::: mkpass RTLpathLivegen.match_prog + ::: mkpass RTLpath.match_program_RTL ::: mkpass Allocproof.match_prog ::: mkpass Tunnelingproof.match_prog ::: mkpass Linearizeproof.match_prog @@ -310,7 +316,9 @@ Proof. destruct (partial_if optim_redundancy Deadcode.transf_program p13ter) as [p14|e] eqn:P14; simpl in T; try discriminate. set (p14bis := total_if all_loads_nontrap Allnontrap.transf_program p14) in *. destruct (Unusedglob.transform_program p14bis) as [p15|e] eqn:P15; simpl in T; try discriminate. - destruct (Allocation.transf_program p15) as [p16|e] eqn:P16; simpl in T; try discriminate. + destruct (RTLpathLivegen.transf_program p15) as [p15_1|e] eqn:P15_1; simpl in T; try discriminate. + set (p15_2 := RTLpath.program_RTL p15_1) in *. + destruct (Allocation.transf_program p15_2) as [p16|e] eqn:P16; simpl in T; try discriminate. set (p17 := Tunneling.tunnel_program p16) in *. destruct (Linearize.transf_program p17) as [p18|e] eqn:P18; simpl in T; try discriminate. set (p19 := CleanupLabels.transf_program p18) in *. @@ -335,6 +343,8 @@ Proof. exists p14; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match. exists p14bis; split. eapply total_if_match; eauto. apply Allnontrapproof.transf_program_match. exists p15; split. apply Unusedglobproof.transf_program_match; auto. + exists p15_1; split. apply RTLpathLivegen.transf_program_match; auto. + exists p15_2; split. apply RTLpath.transf_program_match; auto. exists p16; split. apply Allocproof.transf_program_match; auto. exists p17; split. apply Tunnelingproof.transf_program_match. exists p18; split. apply Linearizeproof.transf_program_match; auto. @@ -392,7 +402,7 @@ Ltac DestructM := destruct H as (p & M & MM); clear H end. repeat DestructM. subst tp. - assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p25)). + assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p27)). { eapply compose_forward_simulations. eapply SimplExprproof.transl_program_correct; eassumption. @@ -429,6 +439,10 @@ Ltac DestructM := eapply match_if_simulation. eassumption. exact Allnontrapproof.transf_program_correct. eapply compose_forward_simulations. eapply Unusedglobproof.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply RTLpathLivegen.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply RTLpath.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply Allocproof.transf_program_correct; eassumption. eapply compose_forward_simulations. diff --git a/mppa_k1c/lib/RTLpath.v b/mppa_k1c/lib/RTLpath.v index 228bc95b..f170e0c3 100644 --- a/mppa_k1c/lib/RTLpath.v +++ b/mppa_k1c/lib/RTLpath.v @@ -135,6 +135,15 @@ Coercion fundef_RTL: fundef >-> RTL.fundef. Definition program_RTL (p: program) : RTL.program := transform_program fundef_RTL p. Coercion program_RTL: program >-> RTL.program. +Definition match_program_RTL (p: RTLpath.program) (tp: RTL.program) := + match_program (fun ctx f tf => tf = fundef_RTL f) eq p tp. + +Lemma transf_program_match: + forall p, match_program_RTL p (program_RTL p). +Proof. + intros. eapply match_transform_program; eauto. +Qed. + (** * Path-step semantics of RTLpath programs *) (* Semantics of internal instructions (mimicking RTL semantics) *) @@ -330,7 +339,6 @@ Definition final_state (st: state) (i:int): Prop Definition semantics (p: program) := Semantics (step (Genv.globalenv (program_RTL p))) (initial_state p) final_state (Genv.globalenv p). - (** * Proving the bisimulation between (semantics p) and (RTL.semantics p). *) (** ** Preliminaries: simple tactics for option-monad *) @@ -516,6 +524,67 @@ Qed. End CORRECTNESS. +Lemma program_equals {A B: Type} : forall (p1 p2: AST.program A B), + prog_defs p1 = prog_defs p2 -> + prog_public p1 = prog_public p2 -> + prog_main p1 = prog_main p2 -> + p1 = p2. +Proof. + intros. destruct p1. destruct p2. simpl in *. subst. auto. +Qed. + +Lemma cons_extract {A: Type} : forall (l: list A) a b, a = b -> a::l = b::l. +Proof. + intros. congruence. +Qed. + +(* Definition transf_program : RTLpath.program -> RTL.program := transform_program fundef_RTL. + +Lemma transf_program_proj: forall p, program_RTL (transf_program p) = p. +Proof. + intros p. destruct p as [defs pub main]. unfold program_proj. simpl. + apply program_equals; simpl; auto. + induction defs. + - simpl; auto. + - simpl. rewrite IHdefs. + destruct a as [id gd]; simpl. + destruct gd as [f|v]; simpl; auto. + rewrite transf_fundef_proj. auto. +Qed. *) + +Lemma match_program_transf: + forall p tp, match_program_RTL p tp -> program_RTL p = tp. +Proof. + intros p tp H. inversion_clear H. inv H1. + destruct p as [defs pub main]. destruct tp as [tdefs tpub tmain]. simpl in *. + subst. unfold program_RTL. unfold transform_program. simpl. + apply program_equals; simpl; auto. + induction H0; simpl; auto. + rewrite IHlist_forall2. apply cons_extract. + destruct a1 as [ida gda]. destruct b1 as [idb gdb]. + simpl in *. + inv H. inv H2. + - simpl in *. subst. auto. + - simpl in *. subst. inv H. auto. +Qed. + + +Section PRESERVATION. + +Variable prog: RTLpath.program. +Variable tprog: RTL.program. +Hypothesis TRANSF: match_program_RTL prog tprog. +Let ge := Genv.globalenv prog. +Let tge := Genv.globalenv tprog. + +Theorem transf_program_correct: + forward_simulation (RTLpath.semantics prog) (RTL.semantics tprog). +Proof. + pose proof (match_program_transf prog tprog TRANSF) as TR. subst. + eapply RTLpath_correct. +Qed. + +End PRESERVATION. -- cgit From 8c1e56414da2326711efdda18fde7578cc78a373 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 25 May 2020 11:11:06 +0200 Subject: Adding stub RTLpathLivegenaux.ml --- Makefile.extr | 3 ++- mppa_k1c/lib/RTLpathLivegen.v | 2 ++ mppa_k1c/lib/RTLpathLivegenaux.ml | 4 ++++ 3 files changed, 8 insertions(+), 1 deletion(-) create mode 100644 mppa_k1c/lib/RTLpathLivegenaux.ml diff --git a/Makefile.extr b/Makefile.extr index f2d06def..b131e84b 100644 --- a/Makefile.extr +++ b/Makefile.extr @@ -43,7 +43,8 @@ cparser/pre_parser_messages.ml: DIRS=extraction \ lib common $(ARCH) backend cfrontend cparser driver \ - exportclight debug mppa_k1c/unittest mppa_k1c/abstractbb/Impure/ocaml + exportclight debug mppa_k1c/unittest mppa_k1c/abstractbb/Impure/ocaml \ + mppa_k1c/lib INCLUDES=$(patsubst %,-I %, $(DIRS)) diff --git a/mppa_k1c/lib/RTLpathLivegen.v b/mppa_k1c/lib/RTLpathLivegen.v index 80df6e7c..4acd646f 100644 --- a/mppa_k1c/lib/RTLpathLivegen.v +++ b/mppa_k1c/lib/RTLpathLivegen.v @@ -18,6 +18,8 @@ Local Open Scope option_monad_scope. Axiom build_path_map: RTL.function -> path_map. +Extract Constant build_path_map => "RTLpathLivegenaux.build_path_map". + Fixpoint list_mem (rl: list reg) (alive: Regset.t) {struct rl}: bool := match rl with | nil => true diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml new file mode 100644 index 00000000..84ba345a --- /dev/null +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -0,0 +1,4 @@ +(* open RTLpath *) +open Maps + +let build_path_map f = PTree.empty -- cgit From 2a6cd67fd463e19097316962adf5558fefe98eeb Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 25 May 2020 17:39:39 +0200 Subject: First try for build_path_map --- mppa_k1c/lib/RTLpathLivegenaux.ml | 168 +++++++++++++++++++++++++++++++++++++- 1 file changed, 166 insertions(+), 2 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index 84ba345a..fa0d23d6 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -1,4 +1,168 @@ -(* open RTLpath *) +open RTL +open RTLpath +open Registers open Maps +open Camlcoq +open Datatypes -let build_path_map f = PTree.empty +let debug_flag = ref false + +let dprintf fmt = let open Printf in + match !debug_flag with + | true -> printf fmt + | false -> ifprintf stdout fmt + +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) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> Some n +| 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) + +let rec list_to_regset = function + | [] -> Regset.empty + | r::l -> Regset.add r (list_to_regset l) + +let get_input_regs i = + let empty = Regset.empty in + match i with + | Inop _ -> empty + | Iop (_,lr,_,_) | Iload (_,_,_,lr,_,_) | Icond (_,lr,_,_,_) -> list_to_regset lr + | Istore (_,_,lr,r,_) -> Regset.add r (list_to_regset lr) + | Icall (_, ri, lr, _, _) | Itailcall (_, ri, lr) -> begin + let rs = list_to_regset lr in + match ri with + | Coq_inr _ -> rs + | Coq_inl r -> Regset.add r rs + end + | Ibuiltin (_, lbr, _, _) -> list_to_regset @@ AST.params_of_builtin_args lbr + | Ijumptable (r, _) -> Regset.add r empty + | Ireturn opr -> (match opr with Some r -> Regset.add r empty | None -> empty) + +let get_output_reg i = + match i with + | Inop _ | Istore _ | Icond _ | Itailcall _ | Ijumptable _ | Ireturn _ -> None + | 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 + +let get_path_map code entry join_points = + let visited = ref (PTree.map (fun n i -> false) code) in + let path_map = ref PTree.empty in + let rec dig_path e = + let input_regs = ref (Regset.empty) in + let alive_regs = ref (Regset.empty) in + let psize = ref (-1) in + let path_successors = ref [] in + let rec dig_path_rec n : (path_info * node list) option = + if not (get_some @@ PTree.get n !visited) then + let inst = get_some @@ PTree.get n code in + begin + visited := PTree.set n true !visited; + psize := !psize + 1; + input_regs := Regset.union !input_regs (Regset.diff (get_input_regs inst) !alive_regs); + (match (get_output_reg inst) with None -> () | Some r -> alive_regs := Regset.add r !alive_regs); + let successor = match predicted_successor inst 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; + dig_path_rec n' + end + | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); input_regs = !input_regs }, !path_successors @ successors_inst inst) + end + else None + in match dig_path_rec e with + | None -> () + | Some ret -> + let (path_info, succs) = ret in + begin + path_map := PTree.set e path_info !path_map; + List.iter dig_path succs + end + in begin + dig_path entry; + !path_map + end + +let print_true_nodes booltree = begin + dprintf "["; + List.iter (fun (n,b) -> + if b then dprintf "%d " (P.to_int n) + ) (PTree.elements booltree); + dprintf "]"; +end + +let print_regset rs = begin + dprintf "["; + List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs); + dprintf "]" +end + +let print_path_info pi = begin + dprintf "(psize=%d; " (Camlcoq.Nat.to_int pi.psize); + dprintf "input_regs="; + print_regset pi.input_regs; + dprintf ")" +end + +let print_path_map path_map = begin + dprintf "["; + List.iter (fun (n,pi) -> + dprintf "\n\t"; + dprintf "%d: " (P.to_int n); + print_path_info pi + ) (PTree.elements path_map); + dprintf "]" +end + +let build_path_map f = + let code = f.fn_code in + let entry = f.fn_entrypoint in + let join_points = get_join_points code entry in + let path_map = get_path_map code entry join_points in + begin + debug_flag := true; + dprintf "Join points: "; + print_true_nodes join_points; + dprintf "\nPath map: "; + print_path_map path_map; + dprintf "\n"; + debug_flag := false; + path_map + end -- cgit From 053f796a74e818eab58bd0ae0f5642779ad90345 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 26 May 2020 12:54:27 +0200 Subject: The backend/Liveness.v is not enough --- mppa_k1c/lib/RTLpathLivegenaux.ml | 23 +++++++++++++++++------ 1 file changed, 17 insertions(+), 6 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index fa0d23d6..4b063623 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -80,12 +80,11 @@ let get_join_points code entry = | pc :: l -> traverse pc; traverse_succs l in traverse entry; !reached_twice +(* Does not set the input_regs field *) let get_path_map code entry join_points = let visited = ref (PTree.map (fun n i -> false) code) in let path_map = ref PTree.empty in let rec dig_path e = - let input_regs = ref (Regset.empty) in - let alive_regs = ref (Regset.empty) in let psize = ref (-1) in let path_successors = ref [] in let rec dig_path_rec n : (path_info * node list) option = @@ -94,8 +93,6 @@ let get_path_map code entry join_points = begin visited := PTree.set n true !visited; psize := !psize + 1; - input_regs := Regset.union !input_regs (Regset.diff (get_input_regs inst) !alive_regs); - (match (get_output_reg inst) with None -> () | Some r -> alive_regs := Regset.add r !alive_regs); let successor = match predicted_successor inst with | None -> None | Some n' -> if get_some @@ PTree.get n' join_points then None else Some n' @@ -104,7 +101,7 @@ let get_path_map code entry join_points = path_successors := !path_successors @ non_predicted_successors inst; dig_path_rec n' end - | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); input_regs = !input_regs }, !path_successors @ successors_inst inst) + | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); input_regs = Regset.empty }, !path_successors @ successors_inst inst) end else None in match dig_path_rec e with @@ -120,6 +117,20 @@ let get_path_map code entry join_points = !path_map end +let set_pathmap_liveness f pm = + let liveness = Liveness.analyze f in + let new_pm = ref PTree.empty in + begin + match liveness with + | None -> failwith "RTLpathLivegenaux.set_pathmap_liveness: Liveness analysis failed" + | Some pmap -> + List.iter (fun (n, pi) -> + let rs = PMap.get n pmap in + new_pm := PTree.set n {psize=pi.psize; input_regs=rs} !new_pm + ) (PTree.elements pm); + !new_pm + end + let print_true_nodes booltree = begin dprintf "["; List.iter (fun (n,b) -> @@ -155,7 +166,7 @@ let build_path_map f = let code = f.fn_code in let entry = f.fn_entrypoint in let join_points = get_join_points code entry in - let path_map = get_path_map code entry join_points in + let path_map = set_pathmap_liveness f @@ get_path_map code entry join_points in begin debug_flag := true; dprintf "Join points: "; -- cgit From 2ee7ae8a6206bc953e18d3d73e6c69c1270f6ada Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 26 May 2020 13:16:18 +0200 Subject: Attempt to modify the transfer function --- mppa_k1c/lib/RTLpathLivegenaux.ml | 40 ++++++++++++++++++++++++++++++++++++++- 1 file changed, 39 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index 4b063623..b031665e 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -4,6 +4,8 @@ open Registers open Maps open Camlcoq open Datatypes +open Lattice +open Kildall let debug_flag = ref false @@ -117,8 +119,44 @@ let get_path_map code entry join_points = !path_map end +(* Adapted from backend/Liveness.v *) +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) + +(** val analyze : coq_function -> Regset.t PMap.t option **) + +let analyze f = + DS.fixpoint f.fn_code successors_instr (transfer f) + let set_pathmap_liveness f pm = - let liveness = Liveness.analyze f in + let liveness = analyze f in let new_pm = ref PTree.empty in begin match liveness with -- cgit From 77c9e3ecec1dbdbecc958d6ed40e1fd53c1b3556 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 26 May 2020 17:37:05 +0200 Subject: Making our own Kildall variant --- mppa_k1c/lib/RTLpathLivegenaux.ml | 132 ++++++++++++++++++++++++-------------- 1 file changed, 85 insertions(+), 47 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index b031665e..e2609a60 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -4,8 +4,6 @@ open Registers open Maps open Camlcoq open Datatypes -open Lattice -open Kildall let debug_flag = ref false @@ -120,52 +118,98 @@ let get_path_map code entry join_points = end (* Adapted from backend/Liveness.v *) -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) - -(** val analyze : coq_function -> Regset.t PMap.t option **) +let transfer after = let open Liveness in function + | 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 + +let get_last_nodes f = + let visited = ref (PTree.map (fun n i -> false) f.fn_code) in + let rec step n = + let inst = get_some @@ PTree.get n f.fn_code in + let successors = successors_inst inst in + if get_some @@ PTree.get n !visited then [] + else begin + visited := PTree.set n true !visited; + match successors with + | [] -> [n] + | _ -> List.concat (List.map step successors) + end + in step f.fn_entrypoint + +let print_regset rs = begin + dprintf "["; + List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs); + dprintf "]" +end + +let print_ptree_regset pt = begin + dprintf "["; + List.iter (fun (n, rs) -> + dprintf "\n\t"; + dprintf "%d: " (P.to_int n); + print_regset rs + ) (PTree.elements pt); + dprintf "]" +end let analyze f = - DS.fixpoint f.fn_code successors_instr (transfer f) + let liveness = ref (PTree.map (fun n i -> None) f.fn_code) in + let predecessors = Duplicateaux.get_predecessors_rtl f.fn_code in + let last_nodes = get_last_nodes f in + let rec step liveout n = (* liveout is the input_regs from the successor *) + let inst = get_some @@ PTree.get n f.fn_code in + let continue = ref true in + let alive = match get_some @@ PTree.get n !liveness with + | None -> transfer liveout inst + | Some pre_alive -> begin + let union = Regset.union pre_alive liveout in + let new_alive = transfer union inst in + (if Regset.equal pre_alive new_alive then continue := false); + new_alive + end + in begin + liveness := PTree.set n (Some alive) !liveness; + if !continue then + let preds = get_some @@ PTree.get n predecessors in + List.iter (step alive) preds + end + in begin + List.iter (step Regset.empty) last_nodes; + let liveness_noopt = PTree.map (fun n i -> get_some i) !liveness in + begin + debug_flag := true; + dprintf "Liveness: "; print_ptree_regset liveness_noopt; dprintf "\n"; + debug_flag := false; + liveness_noopt + end + end let set_pathmap_liveness f pm = let liveness = analyze f in let new_pm = ref PTree.empty in begin - match liveness with - | None -> failwith "RTLpathLivegenaux.set_pathmap_liveness: Liveness analysis failed" - | Some pmap -> - List.iter (fun (n, pi) -> - let rs = PMap.get n pmap in - new_pm := PTree.set n {psize=pi.psize; input_regs=rs} !new_pm - ) (PTree.elements pm); + List.iter (fun (n, pi) -> + let rs = get_some @@ PTree.get n liveness in + new_pm := PTree.set n {psize=pi.psize; input_regs=rs} !new_pm + ) (PTree.elements pm); !new_pm end @@ -177,12 +221,6 @@ let print_true_nodes booltree = begin dprintf "]"; end -let print_regset rs = begin - dprintf "["; - List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs); - dprintf "]" -end - let print_path_info pi = begin dprintf "(psize=%d; " (Camlcoq.Nat.to_int pi.psize); dprintf "input_regs="; -- cgit From 822dffca97430dadb4933d1cca1cb09672387769 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 26 May 2020 18:08:57 +0200 Subject: Using Sylvain's suggestion instead --- mppa_k1c/lib/RTLpathLivegenaux.ml | 78 +++++++++++++++++++++++++++------------ 1 file changed, 55 insertions(+), 23 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index e2609a60..f5a121ac 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -4,6 +4,8 @@ open Registers open Maps open Camlcoq open Datatypes +open Kildall +open Lattice let debug_flag = ref false @@ -117,7 +119,58 @@ let get_path_map code entry join_points = !path_map end -(* Adapted from backend/Liveness.v *) +let print_regset rs = begin + dprintf "["; + List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs); + dprintf "]" +end + +let print_ptree_regset pt = begin + dprintf "["; + List.iter (fun (n, rs) -> + dprintf "\n\t"; + dprintf "%d: " (P.to_int n); + print_regset rs + ) (PTree.elements pt); + dprintf "]" +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 | Inop _ -> after | Iop (_, args, res, _) -> @@ -147,28 +200,6 @@ let get_last_nodes f = let successors = successors_inst inst in if get_some @@ PTree.get n !visited then [] else begin - visited := PTree.set n true !visited; - match successors with - | [] -> [n] - | _ -> List.concat (List.map step successors) - end - in step f.fn_entrypoint - -let print_regset rs = begin - dprintf "["; - List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs); - dprintf "]" -end - -let print_ptree_regset pt = begin - dprintf "["; - List.iter (fun (n, rs) -> - dprintf "\n\t"; - dprintf "%d: " (P.to_int n); - print_regset rs - ) (PTree.elements pt); - dprintf "]" -end let analyze f = let liveness = ref (PTree.map (fun n i -> None) f.fn_code) in @@ -201,6 +232,7 @@ let analyze f = liveness_noopt end end +*) let set_pathmap_liveness f pm = let liveness = analyze f in -- cgit From b15ddf756d6edad5a6bbe55722dcb2356e3a9a85 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 27 May 2020 16:17:04 +0200 Subject: FIX: was making too big paths (including call and builtins in paths) --- mppa_k1c/lib/RTLpathLivegenaux.ml | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index f5a121ac..bfa32cdf 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -25,7 +25,8 @@ let successors_inst = function | Itailcall _ | Ireturn _ -> [] let predicted_successor = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> Some n +| 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 @@ -238,6 +239,8 @@ let set_pathmap_liveness f pm = let liveness = analyze f in let new_pm = ref PTree.empty in begin + debug_flag := true; + dprintf "Liveness: "; print_ptree_regset liveness; dprintf "\n"; List.iter (fun (n, pi) -> let rs = get_some @@ PTree.get n liveness in new_pm := PTree.set n {psize=pi.psize; input_regs=rs} !new_pm -- cgit From c2f5e529e8b1af261301f6bac89eecad6e523347 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 27 May 2020 16:39:43 +0200 Subject: Removing debug --- mppa_k1c/lib/RTLpathLivegenaux.ml | 3 --- 1 file changed, 3 deletions(-) diff --git a/mppa_k1c/lib/RTLpathLivegenaux.ml b/mppa_k1c/lib/RTLpathLivegenaux.ml index bfa32cdf..09684f22 100644 --- a/mppa_k1c/lib/RTLpathLivegenaux.ml +++ b/mppa_k1c/lib/RTLpathLivegenaux.ml @@ -239,7 +239,6 @@ let set_pathmap_liveness f pm = let liveness = analyze f in let new_pm = ref PTree.empty in begin - debug_flag := true; dprintf "Liveness: "; print_ptree_regset liveness; dprintf "\n"; List.iter (fun (n, pi) -> let rs = get_some @@ PTree.get n liveness in @@ -279,12 +278,10 @@ let build_path_map f = let join_points = get_join_points code entry in let path_map = set_pathmap_liveness f @@ get_path_map code entry join_points in begin - debug_flag := true; dprintf "Join points: "; print_true_nodes join_points; dprintf "\nPath map: "; print_path_map path_map; dprintf "\n"; - debug_flag := false; path_map end -- cgit From d46e96ef6c0287d6892bfc7d2272b7473f5e4979 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 28 May 2020 16:36:11 +0200 Subject: Branching stub scheduler --- driver/Compiler.v | 16 ++++++++----- mppa_k1c/lib/RTLpathScheduler.v | 45 +++++++++++++++++++++++++++++-------- mppa_k1c/lib/RTLpathScheduleraux.ml | 8 +++++++ 3 files changed, 55 insertions(+), 14 deletions(-) create mode 100644 mppa_k1c/lib/RTLpathScheduleraux.ml diff --git a/driver/Compiler.v b/driver/Compiler.v index 37bb54f7..07889ae5 100644 --- a/driver/Compiler.v +++ b/driver/Compiler.v @@ -39,7 +39,7 @@ Require Tailcall. Require Inlining. Require Renumber. Require Duplicate. -Require RTLpath RTLpathLivegen. +Require RTLpath RTLpathLivegen RTLpathScheduler. Require Constprop. Require CSE. Require ForwardMoves. @@ -154,6 +154,7 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program := @@@ time "Unused globals" Unusedglob.transform_program @@ print (print_RTL 12) @@@ time "RTLpath generation" RTLpathLivegen.transf_program + @@@ time "RTLpath scheduling" RTLpathScheduler.transf_program @@ time "RTL projection" RTLpath.program_RTL @@ print (print_RTL 13) @@@ time "Register allocation" Allocation.transf_program @@ -268,6 +269,7 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.all_loads_nontrap Allnontrapproof.match_prog) ::: mkpass Unusedglobproof.match_prog ::: mkpass RTLpathLivegen.match_prog + ::: mkpass RTLpathScheduler.match_prog ::: mkpass RTLpath.match_program_RTL ::: mkpass Allocproof.match_prog ::: mkpass Tunnelingproof.match_prog @@ -317,8 +319,9 @@ Proof. set (p14bis := total_if all_loads_nontrap Allnontrap.transf_program p14) in *. destruct (Unusedglob.transform_program p14bis) as [p15|e] eqn:P15; simpl in T; try discriminate. destruct (RTLpathLivegen.transf_program p15) as [p15_1|e] eqn:P15_1; simpl in T; try discriminate. - set (p15_2 := RTLpath.program_RTL p15_1) in *. - destruct (Allocation.transf_program p15_2) as [p16|e] eqn:P16; simpl in T; try discriminate. + destruct (RTLpathScheduler.transf_program p15_1) as [p15_2|e] eqn:P15_2; simpl in T; try discriminate. + set (p15_3 := RTLpath.program_RTL p15_2) in *. + destruct (Allocation.transf_program p15_3) as [p16|e] eqn:P16; simpl in T; try discriminate. set (p17 := Tunneling.tunnel_program p16) in *. destruct (Linearize.transf_program p17) as [p18|e] eqn:P18; simpl in T; try discriminate. set (p19 := CleanupLabels.transf_program p18) in *. @@ -344,7 +347,8 @@ Proof. exists p14bis; split. eapply total_if_match; eauto. apply Allnontrapproof.transf_program_match. exists p15; split. apply Unusedglobproof.transf_program_match; auto. exists p15_1; split. apply RTLpathLivegen.transf_program_match; auto. - exists p15_2; split. apply RTLpath.transf_program_match; auto. + exists p15_2; split. apply RTLpathScheduler.transf_program_match; auto. + exists p15_3; split. apply RTLpath.transf_program_match; auto. exists p16; split. apply Allocproof.transf_program_match; auto. exists p17; split. apply Tunnelingproof.transf_program_match. exists p18; split. apply Linearizeproof.transf_program_match; auto. @@ -402,7 +406,7 @@ Ltac DestructM := destruct H as (p & M & MM); clear H end. repeat DestructM. subst tp. - assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p27)). + assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p28)). { eapply compose_forward_simulations. eapply SimplExprproof.transl_program_correct; eassumption. @@ -441,6 +445,8 @@ Ltac DestructM := eapply Unusedglobproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply RTLpathLivegen.transf_program_correct; eassumption. + eapply compose_forward_simulations. + eapply RTLpathScheduler.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply RTLpath.transf_program_correct; eassumption. eapply compose_forward_simulations. diff --git a/mppa_k1c/lib/RTLpathScheduler.v b/mppa_k1c/lib/RTLpathScheduler.v index 9cdb6f0f..55ee02c6 100644 --- a/mppa_k1c/lib/RTLpathScheduler.v +++ b/mppa_k1c/lib/RTLpathScheduler.v @@ -20,11 +20,11 @@ NB: the new RTLpath function is generated from the fn_code, the fn_entrypoint an It requires to check that the path structure is wf ! *) -Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node). +(* Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node). *) -(* -Extract Constant untrusted_scheduler => "TODO". -*) +Axiom untrusted_scheduler: RTLpath.function -> code * (PTree.t node). + +Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := { preserv_fnsig: fn_sig f1 = fn_sig f2; @@ -36,8 +36,30 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 }. -(* TODO: remove these two stub parameters *) -Parameter transf_function: RTLpath.function -> res RTLpath.function. +(* TODO - perform appropriate checks on tc and dupmap *) +Definition verified_scheduler (f: RTLpath.function) : res (code * (PTree.t node)) := + let (tc, dupmap) := untrusted_scheduler f in OK (tc, dupmap). + +Lemma verified_scheduler_wellformed_pm f tc pm dm: + fn_path f = pm -> + verified_scheduler f = OK (tc, dm) -> + wellformed_path_map tc pm. +Proof. +Admitted. + +Program Definition transf_function (f: RTLpath.function): res RTLpath.function := + match (verified_scheduler f) with + | Error e => Error e + | OK (tc, dupmap) => + let rtl_tf := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc (fn_entrypoint f) in + OK {| fn_RTL := rtl_tf; fn_path := (fn_path f) |} + end. +Next Obligation. + destruct f as [rtl_f pm EP_WF PATH_WF]. assumption. +Qed. Next Obligation. + destruct f as [rtl_f pm EP_WF PATH_WF]. simpl. + eapply verified_scheduler_wellformed_pm; eauto. simpl. reflexivity. +Qed. Parameter transf_function_correct: forall f f', transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. @@ -58,7 +80,6 @@ Definition transf_fundef (f: fundef) : res fundef := Definition transf_program (p: program) : res program := transform_partial_program transf_fundef p. - (** * Preservation proof *) Local Notation ext alive := (fun r => Regset.In r alive). @@ -186,8 +207,11 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. -Hypothesis all_fundef_liveness_ok: forall b fd, +(* Was a Hypothesis *) +Theorem all_fundef_liveness_ok: forall b fd, Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd. +Proof. +Admitted. Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s. Proof. @@ -250,6 +274,7 @@ Proof. - constructor; eauto. + constructor. + apply transf_fundef_correct; auto. + + eapply all_fundef_liveness_ok; eauto. Qed. Theorem transf_final_states s1 s2 r: @@ -305,9 +330,11 @@ Proof. intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. eapply find_funct_liveness_ok; eauto. + intros. eapply all_fundef_liveness_ok; eauto. + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. intros; exploit function_ptr_preserved; eauto. intros (fd' & X). eexists. intuition eauto. + intros. eapply all_fundef_liveness_ok; eauto. Qed. Lemma sistate_simu f dupmap sp st st' rs m is: @@ -421,7 +448,7 @@ Proof. destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS. (* exec_path *) - try_simplify_someHyps. intros. - exploit path_step_eqlive; eauto. + exploit path_step_eqlive; eauto. { intros. eapply all_fundef_liveness_ok; eauto. } clear STEP EQUIV rs; intros (s2 & STEP & EQLIVE). exploit exec_path_simulation; eauto. clear STEP; intros (path' & s' & PATH' & STEP' & MATCH'). diff --git a/mppa_k1c/lib/RTLpathScheduleraux.ml b/mppa_k1c/lib/RTLpathScheduleraux.ml new file mode 100644 index 00000000..1017cf63 --- /dev/null +++ b/mppa_k1c/lib/RTLpathScheduleraux.ml @@ -0,0 +1,8 @@ +open RTLpath +open RTL +open Maps + +let scheduler f = + let code = f.fn_RTL.fn_code in + let id_ptree = PTree.map (fun n i -> n) code in + (code, id_ptree) -- cgit From 8f246ae75c1f7aa6cb84d3ebdff57d2350877b4a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 29 May 2020 15:50:55 +0200 Subject: Scheduler: Separating the code in superblocks --- kvx/lib/RTLpathScheduleraux.ml | 74 +++++++++++++++++++++++++++++++++++++++++- 1 file changed, 73 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 1017cf63..433ba405 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -1,8 +1,80 @@ open RTLpath open RTL open Maps +open RTLpathLivegenaux +open Registers +open Camlcoq + +type superblock = { + instructions: instruction array; + (* each predicted Pcb has its attached liveins *) + (* This is indexed by the array index (not the pc value) *) + liveins: Regset.t PTree.t; +} + +let print_instructions insts = + if (!debug_flag) then begin + dprintf "[ "; List.iter (PrintRTL.print_instruction stdout) insts; dprintf "]" + end + +let print_superblock sb = + let insts = sb.instructions in + let li = sb.liveins in + begin + dprintf "{ instructions = "; print_instructions (List.map (fun i -> (0, i)) @@ Array.to_list insts); dprintf "\n"; + dprintf " liveins = "; print_ptree_regset li; dprintf "}" + end + +let print_superblocks lsb = + let rec f = function + | [] -> () + | sb :: lsb -> (print_superblock sb; dprintf ",\n"; f lsb) + in begin + dprintf "Superblocks: [\n"; + f lsb; + dprintf "]" + end + +let get_superblocks code entry pm = + let visited = ref (PTree.map (fun n i -> false) code) in + let rec get_superblocks_rec pc = + let liveins = ref (PTree.empty) in + let rec follow pc n (index: int) = + let inst = get_some @@ PTree.get pc code in + if (n == 0) then ([inst], successors_inst inst) + else + begin + (match (non_predicted_successors inst) with + | [pcout] -> + let live = (get_some @@ PTree.get pcout pm).input_regs in + liveins := PTree.set (P.of_int index) live !liveins + | [] -> () + | _ -> failwith "Having more than one non_predicted_successor is not handled"); + match (predicted_successor inst) with + | None -> failwith "Incorrect path" + | Some succ -> + let (insts, nexts) = follow succ (n-1) (index+1) in + (inst :: insts, nexts) + end + in if (get_some @@ PTree.get pc !visited) then [] + else begin + visited := PTree.set pc true !visited; + let n = (get_some @@ PTree.get pc pm).psize in + let (insts, nexts) = follow pc (Camlcoq.Nat.to_int n) 0 in + let superblock = { instructions = Array.of_list insts; liveins = !liveins } in + superblock :: (List.concat @@ List.map get_superblocks_rec nexts) + end + in get_superblocks_rec entry let scheduler f = let code = f.fn_RTL.fn_code in let id_ptree = PTree.map (fun n i -> n) code in - (code, id_ptree) + let entry = f.fn_RTL.fn_entrypoint in + let pm = f.fn_path in + let lsb = get_superblocks code entry pm in + begin + debug_flag := true; + print_superblocks lsb; + debug_flag := false; + (code, id_ptree) + end -- cgit From d82a2138e3b429fc3a8d6ed9f578dd8dcfb6b1ef Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 29 May 2020 16:52:43 +0200 Subject: Scheduler removing debug flag --- kvx/lib/RTLpathScheduleraux.ml | 2 -- 1 file changed, 2 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 433ba405..d890243e 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -73,8 +73,6 @@ let scheduler f = let pm = f.fn_path in let lsb = get_superblocks code entry pm in begin - debug_flag := true; print_superblocks lsb; - debug_flag := false; (code, id_ptree) end -- cgit From 7bbb9afba4e8826e6d23dc2cc71ecc26b2a2fc50 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 29 May 2020 16:52:57 +0200 Subject: Using kvx-elf-ar instead of ar --- runtime/Makefile | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/runtime/Makefile b/runtime/Makefile index ea3c914f..ed4c1d39 100644 --- a/runtime/Makefile +++ b/runtime/Makefile @@ -38,6 +38,12 @@ OBJS=i64_dtos.o i64_dtou.o i64_sar.o i64_sdiv.o i64_shl.o \ vararg.o endif +ifeq ($(ARCH),kvx) + AR=kvx-elf-ar +else + AR=ar +endif + OBJS+=write_profiling_table.o LIB=libcompcert.a @@ -59,7 +65,7 @@ endif $(LIB): $(OBJS) rm -f $(LIB) - ar rcs $(LIB) $(OBJS) + $(AR) rcs $(LIB) $(OBJS) %.o: %.s $(CASMRUNTIME) -o $@ $^ -- cgit From 584a5f28288cee52b2f0eb180f3efcf8e5a2e554 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 1 Jun 2020 16:12:56 +0200 Subject: Putting back path_map and node in untrusted_scheduler --- kvx/lib/RTLpathScheduler.v | 39 +++++++++++++++++++++++++-------------- 1 file changed, 25 insertions(+), 14 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 57ae6c7f..660715a1 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -20,9 +20,11 @@ NB: the new RTLpath function is generated from the fn_code, the fn_entrypoint an It requires to check that the path structure is wf ! *) -(* Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node). *) -Axiom untrusted_scheduler: RTLpath.function -> code * (PTree.t node). +(* Returns: new code, new entrypoint, new pathmap, revmap + * Indeed, the entrypoint might not be the same if the entrypoint node is moved further down + * a path ; same reasoning for the pathmap *) +Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node). Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". @@ -37,28 +39,37 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := }. (* TODO - perform appropriate checks on tc and dupmap *) -Definition verified_scheduler (f: RTLpath.function) : res (code * (PTree.t node)) := - let (tc, dupmap) := untrusted_scheduler f in OK (tc, dupmap). +Definition verified_scheduler (f: RTLpath.function) : res (code * node * path_map * (PTree.t node)) := + let (tctetpm, dupmap) := untrusted_scheduler f in + let (tcte, tpm) := tctetpm in + let (tc, te) := tcte in + OK (tc, te, tpm, dupmap). + +Lemma verified_scheduler_wellformed_pm f tc te tpm dm: + verified_scheduler f = OK (tc, te, tpm, dm) -> + wellformed_path_map tc tpm. +Proof. +Admitted. -Lemma verified_scheduler_wellformed_pm f tc pm dm: - fn_path f = pm -> - verified_scheduler f = OK (tc, dm) -> - wellformed_path_map tc pm. +Lemma verified_scheduler_path_entry f tc te tpm dm: + verified_scheduler f = OK (tc, te, tpm, dm) -> + path_entry tpm te. Proof. Admitted. Program Definition transf_function (f: RTLpath.function): res RTLpath.function := match (verified_scheduler f) with | Error e => Error e - | OK (tc, dupmap) => - let rtl_tf := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc (fn_entrypoint f) in - OK {| fn_RTL := rtl_tf; fn_path := (fn_path f) |} + | OK (tc, te, tpm, dupmap) => + let rtl_tf := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in + OK {| fn_RTL := rtl_tf; fn_path := tpm |} end. Next Obligation. - destruct f as [rtl_f pm EP_WF PATH_WF]. assumption. + destruct f as [rtl_f pm EP_WF PATH_WF]. + eapply verified_scheduler_path_entry; eauto. Qed. Next Obligation. - destruct f as [rtl_f pm EP_WF PATH_WF]. simpl. - eapply verified_scheduler_wellformed_pm; eauto. simpl. reflexivity. + destruct f as [rtl_f pm EP_WF PATH_WF]. + eapply verified_scheduler_wellformed_pm; eauto. Qed. Parameter transf_function_correct: -- cgit From 8c1aee035fcae2e64e639bb2afd2d96a416242af Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 1 Jun 2020 16:50:16 +0200 Subject: Adapting untrusted scheduler --- kvx/lib/RTLpathScheduleraux.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index d890243e..e88aa4ed 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -74,5 +74,5 @@ let scheduler f = let lsb = get_superblocks code entry pm in begin print_superblocks lsb; - (code, id_ptree) + (((code, entry), pm), id_ptree) end -- cgit From 160289d2300f98ff94a89bf8da683334d5e499ba Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 1 Jun 2020 17:29:58 +0200 Subject: Treating instructions as array of positive instead of array of instructions --- kvx/lib/RTLpathScheduleraux.ml | 50 ++++++++++++++++++++++-------------------- 1 file changed, 26 insertions(+), 24 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index e88aa4ed..f656b1cf 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -6,29 +6,32 @@ open Registers open Camlcoq type superblock = { - instructions: instruction array; + instructions: P.t array; (* pointers to code instructions *) (* each predicted Pcb has its attached liveins *) - (* This is indexed by the array index (not the pc value) *) + (* This is indexed by the pc value *) liveins: Regset.t PTree.t; } -let print_instructions insts = +let print_instructions insts code = if (!debug_flag) then begin - dprintf "[ "; List.iter (PrintRTL.print_instruction stdout) insts; dprintf "]" + dprintf "[ "; + List.iter ( + fun n -> (PrintRTL.print_instruction stdout (P.to_int n, get_some @@ PTree.get n code)) + ) insts; dprintf "]" end -let print_superblock sb = +let print_superblock sb code = let insts = sb.instructions in let li = sb.liveins in begin - dprintf "{ instructions = "; print_instructions (List.map (fun i -> (0, i)) @@ Array.to_list insts); dprintf "\n"; + dprintf "{ instructions = "; print_instructions (Array.to_list insts) code; dprintf "\n"; dprintf " liveins = "; print_ptree_regset li; dprintf "}" end -let print_superblocks lsb = +let print_superblocks lsb code = let rec f = function | [] -> () - | sb :: lsb -> (print_superblock sb; dprintf ",\n"; f lsb) + | sb :: lsb -> (print_superblock sb code; dprintf ",\n"; f lsb) in begin dprintf "Superblocks: [\n"; f lsb; @@ -39,28 +42,27 @@ let get_superblocks code entry pm = let visited = ref (PTree.map (fun n i -> false) code) in let rec get_superblocks_rec pc = let liveins = ref (PTree.empty) in - let rec follow pc n (index: int) = + let rec follow pc n = let inst = get_some @@ PTree.get pc code in - if (n == 0) then ([inst], successors_inst inst) + if (n == 0) then ([pc], successors_inst inst) else - begin - (match (non_predicted_successors inst) with + let nexts_from_exit = match (non_predicted_successors inst) with | [pcout] -> - let live = (get_some @@ PTree.get pcout pm).input_regs in - liveins := PTree.set (P.of_int index) live !liveins - | [] -> () - | _ -> failwith "Having more than one non_predicted_successor is not handled"); - match (predicted_successor inst) with - | None -> failwith "Incorrect path" - | Some succ -> - let (insts, nexts) = follow succ (n-1) (index+1) in - (inst :: insts, nexts) - end + let live = (get_some @@ PTree.get pcout pm).input_regs in begin + liveins := PTree.set pc live !liveins; + [pcout] + end + | [] -> [] + | _ -> failwith "Having more than one non_predicted_successor is not handled" + in match (predicted_successor inst) with + | None -> failwith "Incorrect path" + | Some succ -> + let (insts, nexts) = follow succ (n-1) in (pc :: insts, nexts_from_exit @ nexts) in if (get_some @@ PTree.get pc !visited) then [] else begin visited := PTree.set pc true !visited; let n = (get_some @@ PTree.get pc pm).psize in - let (insts, nexts) = follow pc (Camlcoq.Nat.to_int n) 0 in + let (insts, nexts) = follow pc (Camlcoq.Nat.to_int n) in let superblock = { instructions = Array.of_list insts; liveins = !liveins } in superblock :: (List.concat @@ List.map get_superblocks_rec nexts) end @@ -73,6 +75,6 @@ let scheduler f = let pm = f.fn_path in let lsb = get_superblocks code entry pm in begin - print_superblocks lsb; + print_superblocks lsb code; dprintf "\n"; (((code, entry), pm), id_ptree) end -- cgit From ac6d4a6d8d18f65b258d5a588d0d050a1e9aaa77 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 2 Jun 2020 12:56:16 +0200 Subject: Scheduler ready to implement --- kvx/lib/RTLpathScheduleraux.ml | 128 +++++++++++++++++++++++++++++++++++++++-- 1 file changed, 122 insertions(+), 6 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index f656b1cf..effd4970 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -33,11 +33,31 @@ let print_superblocks lsb code = | [] -> () | sb :: lsb -> (print_superblock sb code; dprintf ",\n"; f lsb) in begin - dprintf "Superblocks: [\n"; + dprintf "[\n"; f lsb; dprintf "]" end +(* Adapted from backend/PrintRTL.ml: print_function *) +let print_code code = let open PrintRTL in let open Printf in + if (!debug_flag) then begin + fprintf stdout "{\n"; + let instrs = + List.sort + (fun (pc1, _) (pc2, _) -> compare pc2 pc1) + (List.rev_map + (fun (pc, i) -> (P.to_int pc, i)) + (PTree.elements code)) in + List.iter (print_instruction stdout) instrs; + fprintf stdout "}" + end + +let print_arrayp arr = begin + dprintf "[| "; + Array.iter (fun n -> dprintf "%d, " (P.to_int n)) arr; + dprintf "|]" +end + let get_superblocks code entry pm = let visited = ref (PTree.map (fun n i -> false) code) in let rec get_superblocks_rec pc = @@ -66,7 +86,105 @@ let get_superblocks code entry pm = let superblock = { instructions = Array.of_list insts; liveins = !liveins } in superblock :: (List.concat @@ List.map get_superblocks_rec nexts) end - in get_superblocks_rec entry + in let lsb = get_superblocks_rec entry in begin + (* debug_flag := true; *) + dprintf "Superblocks identified:"; print_superblocks lsb code; dprintf "\n"; + (* debug_flag := false; *) + lsb +end + +(* TODO David *) +let schedule_superblock sb code = + (* stub2: reverse function *) + (* + let reversed = Array.of_list @@ List.rev @@ Array.to_list (sb.instructions) in + let tmp = reversed.(0) in + let last_index = Array.length reversed - 1 in + begin + reversed.(0) <- reversed.(last_index); + reversed.(last_index) <- tmp; + reversed + end *) + (* stub: identity function *) + sb.instructions + +let change_successors i = function + | [] -> ( + match i with + | Itailcall _ | Ireturn _ -> i + | _ -> failwith "Wrong instruction (1)") + | [s] -> ( + match i with + | Inop n -> Inop s + | Iop (a,b,c,n) -> Iop (a,b,c,s) + | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) + | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s) + | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) + | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) + | Ijumptable (a,[n]) -> Ijumptable (a,[s]) + | _ -> failwith "Wrong instruction (2)") + | [s1; s2] -> ( + match i with + | Icond (a,b,n1,n2,p) -> Icond (a,b,s1,s2,p) + | Ijumptable (a, [n1; n2]) -> Ijumptable (a, [s1; s2]) + | _ -> failwith "Wrong instruction (3)") + | ls -> ( + match i with + | Ijumptable (a, ln) -> begin + assert ((List.length ln) == (List.length ls)); + Ijumptable (a, ls) + end + | _ -> failwith "Wrong instruction (4)") + +let change_predicted_successor i s = match i with + | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5)" + | Ijumptable _ -> failwith "Wrong instruction (6) (shouldn't be predicted successor for jumptable)" + | Inop n -> Inop s + | Iop (a,b,c,n) -> Iop (a,b,c,s) + | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) + | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s) + | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) + | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) + | Icond (a,b,n1,n2,p) -> ( + match p with + | Some true -> Icond (a,b,s,n2,p) + | Some false -> Icond (a,b,n1,s,p) + | None -> failwith "Predicted a successor for an Icond with p=None" + ) + +let apply_schedule code sb new_order = + let tc = ref code in + let old_order = sb.instructions in + let last_node = Array.get old_order (Array.length old_order - 1) in + let last_successors = successors_inst @@ get_some @@ PTree.get last_node code in + begin + assert ((Array.length old_order) == (Array.length new_order)); + Array.iteri (fun i n' -> + let inst' = get_some @@ PTree.get n' code in + let new_inst = + if (i == (Array.length old_order - 1)) then + change_successors inst' last_successors + else + change_predicted_successor inst' (Array.get old_order (i+1)) + in tc := PTree.set (Array.get old_order i) new_inst !tc + ) new_order; + !tc + end + +let rec do_schedule code = function + | [] -> code + | sb :: lsb -> + let schedule = schedule_superblock sb code in + let new_code = apply_schedule code sb schedule in + begin + (* debug_flag := true; *) + dprintf "Old Code: "; print_code code; + dprintf "\nSchedule to apply: "; print_arrayp schedule; + dprintf "\nNew Code: "; print_code new_code; + dprintf "\n"; + (* debug_flag := false; *) + do_schedule new_code lsb + end let scheduler f = let code = f.fn_RTL.fn_code in @@ -74,7 +192,5 @@ let scheduler f = let entry = f.fn_RTL.fn_entrypoint in let pm = f.fn_path in let lsb = get_superblocks code entry pm in - begin - print_superblocks lsb code; dprintf "\n"; - (((code, entry), pm), id_ptree) - end + let tc = do_schedule code lsb in + (((tc, entry), pm), id_ptree) -- cgit From 5a8d362c5edbcef2f98a9879213a7e159527ddba Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 4 Jun 2020 18:11:39 +0200 Subject: Progression on sfval builtin --- kvx/lib/RTLpathSE_theory.v | 118 +++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 113 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index b0dcbe80..6d3c50f5 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -857,8 +857,8 @@ Inductive sfval := | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:node) (* 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:node) (* - | Sbuiltin: external_function -> list (builtin_arg sval) -> builtin_res sval -> sfval | Sjumptable: sval -> list node -> sfval *) | Sreturn: option sval -> sfval @@ -870,6 +870,63 @@ Definition sfind_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : s | inr symb => SOME b <- Genv.find_symbol pge symb IN Genv.find_funct_ptr pge b end. +Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *) + +Variable ge: RTL.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, + seval_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. + Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: sfmatch pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(si_pc) rs m) @@ -886,14 +943,17 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac seval_list_sval ge sp lsv rs0 m0 = Some args -> sfmatch pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m E0 (Callstate stack fd args m') -(* - - (fn_code f)!pc = Some(Itailcall sig ros args) -> +(* (fn_code f)!pc = Some(Itailcall sig ros args) -> find_function pge ros rs = Some fd -> funsig fd = sig -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m E0 (Callstate stack fd rs##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' -> + sfmatch pge ge sp st stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m + t (State stack f sp pc (regmap_setres res vres rs) m') (* TODO: | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> @@ -930,6 +990,49 @@ Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) smatch pge ge sp st stack f rs0 m0 t s . +Fixpoint builtin_arg_inj (f: reg -> sval) (arg: builtin_arg reg) : builtin_arg sval := + 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_inj f ba1) (builtin_arg_inj f ba2) + | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_inj f ba1) (builtin_arg_inj f ba2) + end. + +Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg varg: forall arg, + (forall r, seval_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_inj sreg arg) varg. +Proof. + induction arg. + all: try (intros SEVAL BARG; inv BARG; constructor; congruence). + - intros SEVAL BARG. inv BARG. simpl. constructor. (* eapply IHarg1. + (* FIXME - not really clear whether my above definitions are correct or not *) + + intros SEVAL BARG. + destruct arg. + all: try (simpl; inv BARG; constructor; congruence). *) + +Admitted. + +Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs: + (forall r, seval_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_inj sreg) args) vargs. +Proof. + induction 2. + - constructor. + - simpl. constructor; [| assumption]. + eapply seval_builtin_arg_correct; eauto. +Qed. + + (** * Symbolic execution of final step *) Definition sfstep (i: instruction) (prev: sistate_local): sfval := match i with @@ -941,6 +1044,9 @@ Definition sfstep (i: instruction) (prev: sistate_local): sfval := let svos := sum_left_map prev.(si_sreg) ros in let sargs := list_sval_inj (List.map prev.(si_sreg) args) in Stailcall sig svos sargs + | Ibuiltin ef args res pc => + let sargs := List.map (builtin_arg_inj prev.(si_sreg)) args in + Sbuiltin ef sargs res pc | Ireturn or => let sor := SOME r <- or IN Some (prev.(si_sreg) r) in Sreturn sor @@ -966,12 +1072,14 @@ Proof. rewrite REG; auto. - eauto. - erewrite seval_list_sval_inj; simpl; auto. - + (* Ibuiltin *) admit. + + (* Ibuiltin *) intros. eapply exec_Sbuiltin; eauto. + eapply seval_builtin_args_correct; eauto. + (* Ijumptable *) admit. + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto. destruct or; simpl; auto. Admitted. + Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> -- cgit From 57c653f6e5ffed83a03e9f95ef2735b17de67ff1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 15:14:57 +0200 Subject: seval_builtin_correct proved --- kvx/lib/RTLpathSE_theory.v | 17 +++++++---------- 1 file changed, 7 insertions(+), 10 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 6d3c50f5..36838862 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1005,21 +1005,18 @@ Fixpoint builtin_arg_inj (f: reg -> sval) (arg: builtin_arg reg) : builtin_arg s | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_inj f ba1) (builtin_arg_inj f ba2) end. -Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg varg: forall arg, +Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg: forall arg varg, (forall r, seval_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_inj sreg arg) varg. Proof. induction arg. - all: try (intros SEVAL BARG; inv BARG; constructor; congruence). - - intros SEVAL BARG. inv BARG. simpl. constructor. (* eapply IHarg1. - (* FIXME - not really clear whether my above definitions are correct or not *) - - intros SEVAL BARG. - destruct arg. - all: try (simpl; inv BARG; constructor; congruence). *) - -Admitted. + 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, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> -- cgit From f8db0c2ca3a0040f3edf3341e644cd805400e722 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 15:24:12 +0200 Subject: sfstep_correct proved --- kvx/lib/RTLpathSE_theory.v | 33 +++++++++------------------------ 1 file changed, 9 insertions(+), 24 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 36838862..d1133af6 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -858,9 +858,7 @@ Inductive sfval := (* 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:node) -(* - | Sjumptable: sval -> list node -> sfval -*) + | Sjumptable (sv: sval) (tbl: list node) | Sreturn: option sval -> sfval . @@ -943,31 +941,16 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac seval_list_sval ge sp lsv rs0 m0 = Some args -> sfmatch pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m E0 (Callstate stack fd args m') -(* (fn_code f)!pc = Some(Itailcall sig ros args) -> - find_function pge ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m - E0 (Callstate stack fd rs##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' -> sfmatch pge ge sp st stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m t (State stack f sp pc (regmap_setres res vres rs) m') -(* TODO: - | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m': - (fn_code f)!pc = Some(Ibuiltin ef args res pc') -> - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - path_last_step ge pge stack f sp pc rs m - t (State stack f sp pc' (regmap_setres res vres rs) m') - | exec_Ijumptable sp pc rs m arg tbl n pc': (* TODO remove jumptable from here ? *) - (fn_code f)!pc = Some(Ijumptable arg tbl) -> - rs#arg = Vint n -> + | exec_Sjumptable sv tbl pc' n rs m: + seval_sval ge sp sv rs0 m0 = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - path_last_step ge pge stack f sp pc rs m + sfmatch pge ge sp st stack f rs0 m0 (Sjumptable sv tbl) 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' -> @@ -1047,6 +1030,9 @@ Definition sfstep (i: instruction) (prev: sistate_local): sfval := | Ireturn or => let sor := SOME r <- or IN Some (prev.(si_sreg) r) in Sreturn sor + | Ijumptable reg tbl => + let sv := prev.(si_sreg) reg in + Sjumptable sv tbl | _ => Snone end. @@ -1071,11 +1057,10 @@ Proof. - erewrite seval_list_sval_inj; simpl; auto. + (* Ibuiltin *) intros. eapply exec_Sbuiltin; eauto. eapply seval_builtin_args_correct; eauto. - + (* Ijumptable *) admit. + + (* Ijumptable *) intros. eapply exec_Sjumptable; eauto. congruence. + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto. destruct or; simpl; auto. -Admitted. - +Qed. Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> -- cgit From ab7b5e278c276644ccb12c2252dce62848943c17 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 15:41:28 +0200 Subject: sfstep_complete: proof for Itailcall --- kvx/lib/RTLpathSE_theory.v | 21 ++++++++++++++++----- 1 file changed, 16 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index d1133af6..d4d72247 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -937,6 +937,7 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac | exec_Stailcall stk rs m sig svos args fd m' lsv: sfind_function pge 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' -> seval_list_sval ge sp lsv rs0 m0 = Some args -> sfmatch pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m @@ -1053,7 +1054,6 @@ Proof. + (* Itailcall *) intros. eapply exec_Stailcall; auto. - destruct ros; simpl in * |- *; auto. rewrite REG; auto. - - eauto. - erewrite seval_list_sval_inj; simpl; auto. + (* Ibuiltin *) intros. eapply exec_Sbuiltin; eauto. eapply seval_builtin_args_correct; eauto. @@ -1070,15 +1070,26 @@ Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: sistep i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. - intros PC1 PC2 (PRE&MEM®) LAST Hi. - destruct i as [ | | | |(*Icall*)sig ros args res pc'| | | | |(*Ireturn*)or]; - subst; try_simplify_someHyps; try (unfold sist_set_local in Hi; try congruence); inversion LAST; subst; clear LAST; simpl in * |- *. + intros PC1 PC2 (PRE&MEM®) LAST HSIS. + destruct i as [ (* Inop *) | (* Iop *) | (* Iload *) | (* Istore *) + | (* Icall *) sig ros args res pc' + | (* Itailcall *) sig ros args + | (* Ibuiltin *) ef bargs br pc' + | (* Icond *) + | (* Ijumptable *) jr tbl + | (*Ireturn*) or]; + subst; try_simplify_someHyps; try (unfold sist_set_local in HSIS; try congruence); + inversion LAST; subst; clear LAST; simpl in * |- *. + (* Icall *) erewrite seval_list_sval_inj in * |- ; simpl; try_simplify_someHyps; auto. intros; eapply exec_Icall; eauto. destruct ros; simpl in * |- *; auto. rewrite REG in * |- ; auto. - + (* Itaillcall *) admit. + + (* Itailcall *) + intros HPC SMEM. erewrite seval_list_sval_inj in H10; auto. inv H10. + eapply exec_Itailcall; eauto. + destruct ros; simpl in * |- *; auto. + rewrite REG in * |- ; auto. + (* Ibuiltin *) admit. + (* Ijumptable *) admit. + (* Ireturn *) -- cgit From 047d7758ea85d99b0e6a332e53939421cf1b8068 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 15:55:49 +0200 Subject: sfstep_complete : Ibuiltin proved --- kvx/lib/RTLpathSE_theory.v | 29 ++++++++++++++++++++++++++++- 1 file changed, 28 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index d4d72247..b096fcbf 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1013,6 +1013,31 @@ 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, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + seval_builtin_arg ge sp m rs0 m0 (builtin_arg_inj 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, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> + seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_inj 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. (** * Symbolic execution of final step *) Definition sfstep (i: instruction) (prev: sistate_local): sfval := @@ -1090,7 +1115,9 @@ Proof. eapply exec_Itailcall; eauto. destruct ros; simpl in * |- *; auto. rewrite REG in * |- ; auto. - + (* Ibuiltin *) admit. + + (* Ibuiltin *) intros HPC SMEM. + eapply exec_Ibuiltin; eauto. + eapply seval_builtin_args_complete; eauto. + (* Ijumptable *) admit. + (* Ireturn *) intros; subst. enough (v=regmap_optget or Vundef rs) as ->. -- cgit From d2b5f69a433a0dd2b83923d46225316596155e58 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 15:57:30 +0200 Subject: sfstep_complete proved --- kvx/lib/RTLpathSE_theory.v | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index b096fcbf..0ca64ab8 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1118,12 +1118,14 @@ Proof. + (* Ibuiltin *) intros HPC SMEM. eapply exec_Ibuiltin; eauto. eapply seval_builtin_args_complete; eauto. - + (* Ijumptable *) admit. + + (* Ijumptable *) intros HPC SMEM. + eapply exec_Ijumptable; eauto. + congruence. + (* Ireturn *) intros; subst. enough (v=regmap_optget or Vundef rs) as ->. * eapply exec_Ireturn; eauto. * intros; destruct or; simpl; congruence. -Admitted. +Qed. (** * Main function of the symbolic execution *) -- cgit From 9a28e40b9a842f7eeb52e9b26e2357e326919b5c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 5 Jun 2020 17:02:19 +0200 Subject: Proof of sfmatch_equiv --- kvx/lib/RTLpathSE_theory.v | 127 ++++++++++++++++++++++++++------------------- 1 file changed, 74 insertions(+), 53 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 0ca64ab8..d8e19295 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1127,7 +1127,6 @@ Proof. * intros; destruct or; simpl; congruence. Qed. - (** * Main function of the symbolic execution *) Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}. @@ -1177,53 +1176,55 @@ Theorem sstep_correct f pc pge ge sp path stack rs m t s: path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. -(* Local Hint Resolve init_ssem sistep_preserves_ssem_exits: core. - intros PATH STEP; unfold sstep; rewrite PATH; simpl. - lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (sisteps_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_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); - (* intro SEM *) - (simpl; unfold ssem; 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 sisteps_default_succ; eauto; simpl; - intros DEF; rewrite DEF in Hi; auto - | clear Hi; rewrite Hi' ]); - (* eexists *) - (eexists; constructor; eauto). - - (* early *) - eapply smatch_early; eauto. - unfold ssem; simpl; rewrite CONT. - destruct (sistep i st) as [st'|] eqn: Hst'; simpl; eauto. - - destruct SEM as (SEM & PC & HNYE). - destruct (sistep 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 sistep_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 smatch_normal; simpl; eauto. - unfold ssem in SEM. - rewrite CONT in SEM. - destruct SEM as (SEM & PC & HNYE). - rewrite <- PC. - eapply exec_Snone. } - { eapply smatch_early; eauto. } - + (* normal non-Snone instruction *) - eapply smatch_normal; eauto. - * unfold ssem; simpl; rewrite CONT; intuition. - * simpl. eapply sfstep_correct; eauto. - rewrite PC; auto. *) -Admitted. + Local Hint Resolve init_ssem: core. + intros PATH STEP; unfold sstep; rewrite PATH; simpl. + lapply (final_node_path_simpl f path pc); eauto. intro WF. + exploit (sisteps_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_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); + (* intro SEM *) + (simpl; unfold ssem; 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 sisteps_default_succ; eauto; simpl; + intros DEF; rewrite DEF in Hi; auto + | clear Hi; rewrite Hi' ]); + (* eexists *) + (eexists; constructor; eauto). + - (* early *) + eapply smatch_early; eauto. + unfold ssem; simpl; rewrite CONT. + destruct (sistep i st) as [st'|] eqn: Hst'; simpl; eauto. + destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx. + constructor; auto. eapply sistep_preserves_allfu; eauto. + - destruct SEM as (SEM & PC & HNYE). + destruct (sistep 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 sistep_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 smatch_normal; simpl; eauto. + unfold ssem in SEM. + rewrite CONT in SEM. + destruct SEM as (SEM & PC & HNYE). + rewrite <- PC. + eapply exec_Snone. } + { eapply smatch_early; eauto. } + + (* normal non-Snone instruction *) + eapply smatch_normal; eauto. + * unfold ssem; simpl; rewrite CONT; intuition. + * simpl. eapply sfstep_correct; eauto. + rewrite PC; auto. +Qed. (* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *) Inductive equiv_stackframe: stackframe -> stackframe -> Prop := @@ -1284,6 +1285,16 @@ Proof. 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 sfmatch_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: sfmatch pge ge sp st stack f rs0 m0 sv rs1 m t s -> (forall r, rs1#r = rs2#r) -> @@ -1294,14 +1305,24 @@ Proof. - (* Snone *) intros; eexists; econstructor. + eapply State_equiv; eauto. + eapply exec_Snone. - - (* Scall *) intros; eexists; econstructor. + - (* Scall *) + intros; eexists; econstructor. 2: { eapply exec_Scall; eauto. } apply Call_equiv; auto. repeat (constructor; auto). - - (* Sreturn *) intros; eexists; econstructor. - + eapply equiv_state_refl; eauto. - + (* eapply exec_Sreturn; eauto. *) admit. -Admitted. + - (* 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. (* NB: each execution of a symbolic state (produced from [sstep]) represents a concrete execution (sstep is exact). -- cgit From adf742534fd234013909be6b5c9e8b0f390d348a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 8 Jun 2020 14:52:13 +0200 Subject: Branching function_checker from RTLpathLivegen in the scheduler --- kvx/lib/RTLpathScheduler.v | 50 +++++++++++++++++++++++++++++++--------------- 1 file changed, 34 insertions(+), 16 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 660715a1..cddf3e4c 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -9,10 +9,24 @@ Require Import Coqlib Maps Events Errors Op. Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory. +Notation "'ASSERT' A 'WITH' MSG 'IN' B" := (if A then B else Error (msg MSG)) + (at level 200, A at level 100, B at level 200) + : error_monad_scope. + Local Open Scope error_monad_scope. Local Open Scope positive_scope. - +(* Enhanced from kvx/Asmblockgenproof.v *) +Ltac explore := + repeat match goal with + | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in destruct var eqn:EQ1) + | [ H : OK _ = OK _ |- _ ] => monadInv H + | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "EQ" in destruct b eqn:EQ1) + | [ |- context[match ?m with | _ => _ end] ] => destruct m + | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m + | [ H : bind _ _ = OK _ |- _ ] => monadInv H + | [ H : Error _ = OK _ |- _ ] => inversion H + end. (** External oracle returning the new RTLpath function and a mapping of new path_entries to old path_entries @@ -39,36 +53,40 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := }. (* TODO - perform appropriate checks on tc and dupmap *) -Definition verified_scheduler (f: RTLpath.function) : res (code * node * path_map * (PTree.t node)) := +Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_map * (PTree.t node)) := let (tctetpm, dupmap) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in let (tc, te) := tcte in - OK (tc, te, tpm, dupmap). + let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in + ASSERT function_checker tfr tpm WITH "In verified_scheduler: (tfr, tpm) is not wellformed" IN + OK (tfr, tpm, dupmap). -Lemma verified_scheduler_wellformed_pm f tc te tpm dm: - verified_scheduler f = OK (tc, te, tpm, dm) -> - wellformed_path_map tc tpm. +Lemma verified_scheduler_wellformed_pm f tfr tpm dm: + verified_scheduler f = OK (tfr, tpm, dm) -> + wellformed_path_map (fn_code tfr) tpm. Proof. -Admitted. + intros. unfold verified_scheduler in H. explore. + apply function_checker_wellformed_path_map. assumption. +Qed. -Lemma verified_scheduler_path_entry f tc te tpm dm: - verified_scheduler f = OK (tc, te, tpm, dm) -> - path_entry tpm te. +Lemma verified_scheduler_path_entry f tfr tpm dm: + verified_scheduler f = OK (tfr, tpm, dm) -> + path_entry tpm (fn_entrypoint tfr). Proof. -Admitted. + intros. unfold verified_scheduler in H. explore. + apply function_checker_path_entry. assumption. +Qed. Program Definition transf_function (f: RTLpath.function): res RTLpath.function := match (verified_scheduler f) with | Error e => Error e - | OK (tc, te, tpm, dupmap) => - let rtl_tf := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - OK {| fn_RTL := rtl_tf; fn_path := tpm |} + | OK (tfr, tpm, dm) => OK {| fn_RTL := tfr; fn_path := tpm |} end. Next Obligation. - destruct f as [rtl_f pm EP_WF PATH_WF]. + destruct f as [fr pm EP_WF PATH_WF]. eapply verified_scheduler_path_entry; eauto. Qed. Next Obligation. - destruct f as [rtl_f pm EP_WF PATH_WF]. + destruct f as [fr pm EP_WF PATH_WF]. eapply verified_scheduler_wellformed_pm; eauto. Qed. -- cgit From 3284d30617f80d75a86e0a1caaec34b5a39488d3 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 8 Jun 2020 17:52:34 +0200 Subject: Attempt to branch a function equivalence checker --- kvx/lib/RTLpathScheduler.v | 46 +++++++++++++++++++++++++++++----------------- 1 file changed, 29 insertions(+), 17 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index cddf3e4c..358752c4 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -42,24 +42,17 @@ Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". -Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.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_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; - dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 -}. +(* TODO *) +Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : bool := true. -(* TODO - perform appropriate checks on tc and dupmap *) Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_map * (PTree.t node)) := - let (tctetpm, dupmap) := untrusted_scheduler f in + let (tctetpm, dm) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in let (tc, te) := tcte in let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - ASSERT function_checker tfr tpm WITH "In verified_scheduler: (tfr, tpm) is not wellformed" IN - OK (tfr, tpm, dupmap). + ASSERT RTLpathLivegen.function_checker tfr tpm WITH "In verified_scheduler: (tfr, tpm) is not wellformed" IN + ASSERT function_equiv_checker dm f tfr tpm WITH "In verified_scheduler: f and tfr are not equivalent" IN + OK (tfr, tpm, dm). Lemma verified_scheduler_wellformed_pm f tfr tpm dm: verified_scheduler f = OK (tfr, tpm, dm) -> @@ -77,6 +70,17 @@ Proof. apply function_checker_path_entry. assumption. Qed. +Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.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_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; + dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 +}. + +(* Program Definition transf_function (f: RTLpath.function): { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := *) Program Definition transf_function (f: RTLpath.function): res RTLpath.function := match (verified_scheduler f) with | Error e => Error e @@ -85,13 +89,21 @@ Program Definition transf_function (f: RTLpath.function): res RTLpath.function : Next Obligation. destruct f as [fr pm EP_WF PATH_WF]. eapply verified_scheduler_path_entry; eauto. -Qed. Next Obligation. +Defined. Next Obligation. destruct f as [fr pm EP_WF PATH_WF]. eapply verified_scheduler_wellformed_pm; eauto. -Qed. +Defined. (* Next Obligation. +Admitted. *) + +(* Theorem function_equiv_checker_correct: *) -Parameter transf_function_correct: - forall f f', transf_function f = OK f' -> exists dupmap, match_function dupmap f f'. +Theorem transf_function_correct f f': + transf_function f = OK f' -> + exists dupmap, match_function dupmap f f'. +Proof. + intros TF. destruct (verified_scheduler f) eqn:VERIF. + - unfold transf_function in TF. (* rewrite VERIF in TF. *) +Admitted. Theorem transf_function_preserves: forall f f', -- cgit From 7fd913b3e2377104f0fb2aa6cbbcd4f49f53a1f6 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 9 Jun 2020 16:26:20 +0200 Subject: Proof down to verified_scheduler_correct and all_fundef_liveness_ok --- kvx/lib/RTLpathScheduler.v | 54 +++++++++++++++++++++++------------------ kvx/lib/RTLpathSchedulerproof.v | 16 +++++++++--- 2 files changed, 42 insertions(+), 28 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 358752c4..aaf3137d 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -80,8 +80,21 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 }. -(* Program Definition transf_function (f: RTLpath.function): { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := *) -Program Definition transf_function (f: RTLpath.function): res RTLpath.function := +Theorem verified_scheduler_correct f tfr tf tpm dm: + verified_scheduler f = OK (tfr, tpm, dm) -> + (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> + fn_sig f = fn_sig tfr + /\ fn_params f = fn_params tfr + /\ fn_stacksize f = fn_stacksize tfr + /\ dm ! (fn_entrypoint tfr) = Some (fn_entrypoint f) + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1) + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry tpm pc2) + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2). +Proof. +Admitted. + +Program Definition transf_function (f: RTLpath.function): + { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := match (verified_scheduler f) with | Error e => Error e | OK (tfr, tpm, dm) => OK {| fn_RTL := tfr; fn_path := tpm |} @@ -92,31 +105,23 @@ Next Obligation. Defined. Next Obligation. destruct f as [fr pm EP_WF PATH_WF]. eapply verified_scheduler_wellformed_pm; eauto. -Defined. (* Next Obligation. -Admitted. *) - -(* Theorem function_equiv_checker_correct: *) - -Theorem transf_function_correct f f': - transf_function f = OK f' -> - exists dupmap, match_function dupmap f f'. -Proof. - intros TF. destruct (verified_scheduler f) eqn:VERIF. - - unfold transf_function in TF. (* rewrite VERIF in TF. *) -Admitted. - -Theorem transf_function_preserves: - forall f f', - transf_function f = OK f' -> - fn_sig f = fn_sig f' /\ fn_params f = fn_params f' /\ fn_stacksize f = fn_stacksize f'. +Defined. Next Obligation. + exploit verified_scheduler_correct; eauto. + intros (A & B & C & D & E & F & G). + exists dm. econstructor; eauto. +Defined. + +Theorem match_function_preserves f f' dm: + match_function dm f f' -> + fn_sig f = fn_sig f' /\ fn_params f = fn_params f' /\ fn_stacksize f = fn_stacksize f'. Proof. - intros. exploit transf_function_correct; eauto. - destruct 1 as (dupmap & [SIG PARAM SIZE ENTRY CORRECT]). + intros. + destruct H as [SIG PARAM SIZE ENTRY CORRECT]. intuition. Qed. Definition transf_fundef (f: fundef) : res fundef := - transf_partial_fundef transf_function f. + transf_partial_fundef (fun f => proj1_sig (transf_function f)) f. Definition transf_program (p: program) : res program := transform_partial_program transf_fundef p. @@ -223,8 +228,9 @@ 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). + + destruct (transf_function f) as [res H]; simpl in * |- *; auto. + destruct (H _ EQ). + intuition subst; auto. eapply match_Internal; eauto. + eapply match_External. Qed. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index db1f6be7..a1a67b2f 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -65,9 +65,13 @@ Lemma function_sig_preserved: forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. Proof. intros. destruct f. - - simpl in H. monadInv H. simpl. symmetry. - eapply transf_function_preserves. - assumption. + - simpl in H. monadInv H. + destruct (transf_function f) as [res H]; simpl in * |- *; auto. + destruct (H _ EQ). + intuition subst; auto. + symmetry. + eapply match_function_preserves. + eassumption. - simpl in H. monadInv H. reflexivity. Qed. @@ -84,7 +88,11 @@ Proof. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. symmetry. eapply match_program_main. eauto. + destruct f. - * monadInv TRANSF. rewrite <- H3. symmetry; eapply transf_function_preserves. assumption. + * monadInv TRANSF. rewrite <- H3. + destruct (transf_function f) as [res H]; simpl in * |- *; auto. + destruct (H _ EQ). + intuition subst; auto. + symmetry; eapply match_function_preserves. eassumption. * monadInv TRANSF. assumption. - constructor; eauto. + constructor. -- cgit From d5f3ee00bb93b07f6f1a5859750cac345ed261aa Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 9 Jun 2020 17:32:11 +0200 Subject: RTLpathSchedulerproof.all_fundef_liveness_ok is a hypothesis again --- driver/Compiler.vexpand | 2 ++ kvx/lib/RTLpathSchedulerproof.v | 14 +++++--------- 2 files changed, 7 insertions(+), 9 deletions(-) diff --git a/driver/Compiler.vexpand b/driver/Compiler.vexpand index 3285a012..8aaa40c6 100644 --- a/driver/Compiler.vexpand +++ b/driver/Compiler.vexpand @@ -298,6 +298,8 @@ EXPAND_ASM_SEMANTICS EXPAND_RTL_FORWARD_SIMULATIONS eapply compose_forward_simulations. eapply RTLpathLivegenproof.transf_program_correct; eassumption. + pose proof RTLpathLivegenproof.all_fundef_liveness_ok as X. + refine (modusponens _ _ (X _ _ _) _); eauto. intro. eapply compose_forward_simulations. eapply RTLpathSchedulerproof.transf_program_correct; eassumption. eapply compose_forward_simulations. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index a1a67b2f..4d7dfa2f 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -22,11 +22,7 @@ Hypothesis TRANSL: match_prog prog tprog. Let pge := Genv.globalenv prog. Let tpge := Genv.globalenv tprog. -(* Was a Hypothesis *) -Theorem all_fundef_liveness_ok: forall b fd, - Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd. -Proof. -Admitted. +Hypothesis all_fundef_liveness_ok: forall b fd, Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd. Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s. Proof. @@ -97,7 +93,7 @@ Proof. - constructor; eauto. + constructor. + apply transf_fundef_correct; auto. - + eapply all_fundef_liveness_ok; eauto. +(* + eapply all_fundef_liveness_ok; eauto. *) Qed. Theorem transf_final_states s1 s2 r: @@ -153,11 +149,11 @@ Proof. intros; exploit functions_preserved; eauto. intros (fd' & cunit & X). eexists. intuition eauto. eapply find_funct_liveness_ok; eauto. - intros. eapply all_fundef_liveness_ok; eauto. +(* intros. eapply all_fundef_liveness_ok; eauto. *) + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. intros; exploit function_ptr_preserved; eauto. intros (fd' & X). eexists. intuition eauto. - intros. eapply all_fundef_liveness_ok; eauto. +(* intros. eapply all_fundef_liveness_ok; eauto. *) Qed. Lemma sistate_simu f dupmap sp st st' rs m is: @@ -271,7 +267,7 @@ Proof. destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS. (* exec_path *) - try_simplify_someHyps. intros. - exploit path_step_eqlive; eauto. { intros. eapply all_fundef_liveness_ok; eauto. } + exploit path_step_eqlive; eauto. (* { intros. eapply all_fundef_liveness_ok; eauto. } *) clear STEP EQUIV rs; intros (s2 & STEP & EQLIVE). exploit exec_path_simulation; eauto. clear STEP; intros (path' & s' & PATH' & STEP' & MATCH'). -- cgit From d7ee5bde4423ae046a0c7d417cb974fae17b2751 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 9 Jun 2020 17:58:07 +0200 Subject: Cutting into more lemmas --- kvx/lib/RTLpathScheduler.v | 56 +++++++++++++++++++++++++++++++++++++--------- 1 file changed, 45 insertions(+), 11 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index aaf3137d..f3b43894 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -45,6 +45,34 @@ Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". (* TODO *) Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : bool := true. +Lemma function_equiv_checker_entrypoint f tfr tpm dm: + function_equiv_checker dm f tfr tpm = true -> + dm ! (fn_entrypoint tfr) = Some (fn_entrypoint f). +Proof. +Admitted. + +Lemma function_equiv_checker_pathentry1 f tfr tpm dm: + function_equiv_checker dm f tfr tpm = true -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + path_entry tpm pc2. +Proof. +Admitted. + +Lemma function_equiv_checker_pathentry2 f tfr tpm dm: + function_equiv_checker dm f tfr tpm = true -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + path_entry (fn_path f) pc1. +Proof. +Admitted. + +Lemma function_equiv_checker_correct f tfr tpm dm tf: + function_equiv_checker dm f tfr tpm = true -> + (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + sstep_simu dm f tf pc1 pc2. +Proof. +Admitted. + Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_map * (PTree.t node)) := let (tctetpm, dm) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in @@ -70,16 +98,6 @@ Proof. apply function_checker_path_entry. assumption. Qed. -Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.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_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; - dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 -}. - Theorem verified_scheduler_correct f tfr tf tpm dm: verified_scheduler f = OK (tfr, tpm, dm) -> (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> @@ -91,7 +109,23 @@ Theorem verified_scheduler_correct f tfr tf tpm dm: /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry tpm pc2) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2). Proof. -Admitted. + intros VERIF OBL. unfold verified_scheduler in VERIF. explore. simpl in * |- *. + Local Hint Resolve function_equiv_checker_entrypoint + function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 + function_equiv_checker_correct: core. + repeat (constructor; eauto). + exploit function_equiv_checker_entrypoint. eauto. simpl. auto. +Qed. + +Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.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_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; + dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 +}. Program Definition transf_function (f: RTLpath.function): { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := -- cgit From 8b29b29efe1bbfc7f6c14a88a90e62acaddc5985 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 10 Jun 2020 16:45:47 +0200 Subject: Proof down to simu_check_correct --- kvx/lib/RTLpathScheduler.v | 112 ++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 101 insertions(+), 11 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index f3b43894..ae1f69c6 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -42,36 +42,125 @@ Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". +Definition entrypoint_check (dm: PTree.t node) (fr tfr: RTL.function) : res unit := + match dm ! (fn_entrypoint tfr) with + | None => Error (msg "No mapping for (entrypoint tfr)") + | Some etp => if (Pos.eq_dec (fn_entrypoint fr) etp) then OK tt + else Error (msg "Entrypoints do not match") + end. + +Lemma entrypoint_check_correct fr tfr dm: + entrypoint_check dm fr tfr = OK tt -> + dm ! (fn_entrypoint tfr) = Some (fn_entrypoint fr). +Proof. + unfold entrypoint_check. explore; try discriminate. congruence. +Qed. + +Definition path_entry_check_single (pm tpm: path_map) (m: node * node) := + let (pc2, pc1) := m in + match (tpm ! pc2) with + | None => Error (msg "pc2 isn't an entry of tpm") + | Some _ => + match (pm ! pc1) with + | None => Error (msg "pc1 isn't an entry of pm") + | Some _ => OK tt + end + end. + +Lemma path_entry_check_single_correct pm tpm pc1 pc2: + path_entry_check_single pm tpm (pc2, pc1) = OK tt -> + path_entry tpm pc2 /\ path_entry pm pc1. +Proof. + unfold path_entry_check_single. intro. explore. + constructor; congruence. +Qed. + +(* Inspired from Duplicate.verify_mapping_rec *) +Fixpoint path_entry_check_rec (pm tpm: path_map) lm := + match lm with + | nil => OK tt + | m :: lm => do u1 <- path_entry_check_single pm tpm m; + do u2 <- path_entry_check_rec pm tpm lm; + OK tt + end. + +Lemma path_entry_check_rec_correct pm tpm pc1 pc2: forall lm, + path_entry_check_rec pm tpm lm = OK tt -> + In (pc2, pc1) lm -> + path_entry tpm pc2 /\ path_entry pm pc1. +Proof. + induction lm. + - simpl. intuition. + - simpl. intros. explore. destruct x; destruct x0. destruct H0. + + subst. eapply path_entry_check_single_correct; eauto. + + eapply IHlm; assumption. +Qed. + +Definition path_entry_check (dm: PTree.t node) (pm tpm: path_map) := path_entry_check_rec pm tpm (PTree.elements dm). + +Lemma path_entry_check_correct dm pm tpm: + path_entry_check dm pm tpm = OK tt -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + path_entry tpm pc2 /\ path_entry pm pc1. +Proof. + unfold path_entry_check. intros. eapply PTree.elements_correct in H0. + eapply path_entry_check_rec_correct; eassumption. +Qed. + (* TODO *) -Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : bool := true. +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) := OK tt. + +Lemma simu_check_correct dm f tf tfr tpm: + simu_check dm f tfr tpm = OK tt -> + (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + sstep_simu dm f tf pc1 pc2. +Proof. +Admitted. + +Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : res unit := + let pm := fn_path f in + let fr := fn_RTL f in + do _ <- entrypoint_check dm fr tfr; + do _ <- path_entry_check dm pm tpm; + do _ <- simu_check dm f tfr tpm; + OK tt. Lemma function_equiv_checker_entrypoint f tfr tpm dm: - function_equiv_checker dm f tfr tpm = true -> + function_equiv_checker dm f tfr tpm = OK tt -> dm ! (fn_entrypoint tfr) = Some (fn_entrypoint f). Proof. -Admitted. + unfold function_equiv_checker. intros. explore. destruct x. + eapply entrypoint_check_correct; eauto. +Qed. Lemma function_equiv_checker_pathentry1 f tfr tpm dm: - function_equiv_checker dm f tfr tpm = true -> + function_equiv_checker dm f tfr tpm = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry tpm pc2. Proof. -Admitted. + unfold function_equiv_checker. intros. explore. destruct x0. + exploit path_entry_check_correct. eassumption. all: eauto. intuition. +Qed. Lemma function_equiv_checker_pathentry2 f tfr tpm dm: - function_equiv_checker dm f tfr tpm = true -> + function_equiv_checker dm f tfr tpm = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1. Proof. -Admitted. + unfold function_equiv_checker. intros. explore. destruct x0. + exploit path_entry_check_correct. eassumption. all: eauto. intuition. +Qed. Lemma function_equiv_checker_correct f tfr tpm dm tf: - function_equiv_checker dm f tfr tpm = true -> + function_equiv_checker dm f tfr tpm = OK tt -> (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. -Admitted. + unfold function_equiv_checker. intros. explore. destruct x1. + eapply simu_check_correct; eauto. +Qed. Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_map * (PTree.t node)) := let (tctetpm, dm) := untrusted_scheduler f in @@ -79,7 +168,7 @@ Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_m let (tc, te) := tcte in let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in ASSERT RTLpathLivegen.function_checker tfr tpm WITH "In verified_scheduler: (tfr, tpm) is not wellformed" IN - ASSERT function_equiv_checker dm f tfr tpm WITH "In verified_scheduler: f and tfr are not equivalent" IN + do tt <- function_equiv_checker dm f tfr tpm; OK (tfr, tpm, dm). Lemma verified_scheduler_wellformed_pm f tfr tpm dm: @@ -113,8 +202,9 @@ Proof. Local Hint Resolve function_equiv_checker_entrypoint function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 function_equiv_checker_correct: core. + destruct x. repeat (constructor; eauto). - exploit function_equiv_checker_entrypoint. eauto. simpl. auto. + - exploit function_equiv_checker_entrypoint. eapply EQ3. intuition. Qed. Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := { -- cgit From e7269c4fec8c0b3c422f6791f6e178f38a8d0f63 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 10 Jun 2020 17:01:01 +0200 Subject: Correcting id_ptree in RTLpathScheduleraux.ml --- kvx/lib/RTLpathScheduleraux.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index effd4970..1a13da67 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -188,7 +188,7 @@ let rec do_schedule code = function let scheduler f = let code = f.fn_RTL.fn_code in - let id_ptree = PTree.map (fun n i -> n) code in + let id_ptree = PTree.map (fun n i -> n) (f.fn_path) in let entry = f.fn_RTL.fn_entrypoint in let pm = f.fn_path in let lsb = get_superblocks code entry pm in -- cgit From c8aa5b17ed697b4f2434888a3954a28c16a1ec81 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 11 Jun 2020 12:37:33 +0200 Subject: Proof down to simu_check_single_correct --- kvx/lib/RTLpathScheduler.v | 44 ++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 40 insertions(+), 4 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index ae1f69c6..87206b95 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -108,15 +108,51 @@ Proof. Qed. (* TODO *) -Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) := OK tt. +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) (m: node * node) := OK tt. + +Definition function_defined tf tfr tpm := + exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}. + +(* TODO *) +Lemma simu_check_single_correct dm tf f tfr tpm pc1 pc2: + simu_check_single dm f tfr tpm (pc1, pc2) = OK tt -> + function_defined tf tfr tpm -> + sstep_simu dm f tf pc1 pc2. +Proof. +Admitted. + +Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) lm := + match lm with + | nil => OK tt + | m :: lm => do u1 <- simu_check_single dm f tfr tpm m; + do u2 <- simu_check_rec dm f tfr tpm lm; + OK tt + end. + +Lemma simu_check_rec_correct dm f tf tfr tpm pc1 pc2: forall lm, + simu_check_rec dm f tfr tpm lm = OK tt -> + function_defined tf tfr tpm -> + In (pc2, pc1) lm -> + sstep_simu dm f tf pc1 pc2. +Proof. + induction lm. + - simpl. intuition. + - simpl. intros. explore. destruct x. destruct H1. + + subst. eapply simu_check_single_correct; eauto. + + eapply IHlm; assumption. +Qed. + +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) := simu_check_rec dm f tfr tpm (PTree.elements dm). Lemma simu_check_correct dm f tf tfr tpm: simu_check dm f tfr tpm = OK tt -> - (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> + function_defined tf tfr tpm -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. -Admitted. + unfold simu_check. intros. eapply PTree.elements_correct in H1. + eapply simu_check_rec_correct; eassumption. +Qed. Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : res unit := let pm := fn_path f in @@ -154,7 +190,7 @@ Qed. Lemma function_equiv_checker_correct f tfr tpm dm tf: function_equiv_checker dm f tfr tpm = OK tt -> - (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> + function_defined tf tfr tpm -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. -- cgit From 2da975d495b910d6cf8be0a00c15f9f6bba731f1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 11 Jun 2020 16:04:38 +0200 Subject: Restructuring code to allow tf in simu_check_single --- kvx/lib/RTLpathScheduler.v | 127 ++++++++++++++++++++------------------------- 1 file changed, 56 insertions(+), 71 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 87206b95..d7f70cd3 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -42,6 +42,18 @@ Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler". +Program Definition function_builder (tfr: RTL.function) (tpm: path_map) : + { r : res RTLpath.function | forall f', r = OK f' -> fn_RTL f' = tfr} := + match RTLpathLivegen.function_checker tfr tpm with + | false => Error (msg "In function_builder: (tfr, tpm) is not wellformed") + | true => OK {| fn_RTL := tfr; fn_path := tpm |} + end. +Next Obligation. + apply function_checker_path_entry. auto. +Defined. Next Obligation. + apply function_checker_wellformed_path_map. auto. +Defined. + Definition entrypoint_check (dm: PTree.t node) (fr tfr: RTL.function) : res unit := match dm ! (fn_entrypoint tfr) with | None => Error (msg "No mapping for (entrypoint tfr)") @@ -108,79 +120,75 @@ Proof. Qed. (* TODO *) -Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) (m: node * node) := OK tt. - -Definition function_defined tf tfr tpm := - exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}. +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := OK tt. (* TODO *) -Lemma simu_check_single_correct dm tf f tfr tpm pc1 pc2: - simu_check_single dm f tfr tpm (pc1, pc2) = OK tt -> - function_defined tf tfr tpm -> +Lemma simu_check_single_correct dm tf f pc1 pc2: + simu_check_single dm f tf (pc1, pc2) = OK tt -> sstep_simu dm f tf pc1 pc2. Proof. Admitted. -Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) lm := +Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with | nil => OK tt - | m :: lm => do u1 <- simu_check_single dm f tfr tpm m; - do u2 <- simu_check_rec dm f tfr tpm lm; + | m :: lm => do u1 <- simu_check_single dm f tf m; + do u2 <- simu_check_rec dm f tf lm; OK tt end. -Lemma simu_check_rec_correct dm f tf tfr tpm pc1 pc2: forall lm, - simu_check_rec dm f tfr tpm lm = OK tt -> - function_defined tf tfr tpm -> +Lemma simu_check_rec_correct dm f tf pc1 pc2: forall lm, + simu_check_rec dm f tf lm = OK tt -> In (pc2, pc1) lm -> sstep_simu dm f tf pc1 pc2. Proof. induction lm. - simpl. intuition. - - simpl. intros. explore. destruct x. destruct H1. + - simpl. intros. explore. destruct x. destruct H0. + subst. eapply simu_check_single_correct; eauto. + eapply IHlm; assumption. Qed. -Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) := simu_check_rec dm f tfr tpm (PTree.elements dm). +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). -Lemma simu_check_correct dm f tf tfr tpm: - simu_check dm f tfr tpm = OK tt -> - function_defined tf tfr tpm -> +Lemma simu_check_correct dm f tf: + simu_check dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. - unfold simu_check. intros. eapply PTree.elements_correct in H1. + unfold simu_check. intros. eapply PTree.elements_correct in H0. eapply simu_check_rec_correct; eassumption. Qed. -Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tfr: RTL.function) (tpm: path_map) : res unit := +Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) : res unit := let pm := fn_path f in let fr := fn_RTL f in + let tpm := fn_path tf in + let tfr := fn_RTL tf in do _ <- entrypoint_check dm fr tfr; do _ <- path_entry_check dm pm tpm; - do _ <- simu_check dm f tfr tpm; + do _ <- simu_check dm f tf; OK tt. -Lemma function_equiv_checker_entrypoint f tfr tpm dm: - function_equiv_checker dm f tfr tpm = OK tt -> - dm ! (fn_entrypoint tfr) = Some (fn_entrypoint f). +Lemma function_equiv_checker_entrypoint f tf dm: + function_equiv_checker dm f tf = OK tt -> + dm ! (fn_entrypoint tf) = Some (fn_entrypoint f). Proof. unfold function_equiv_checker. intros. explore. destruct x. eapply entrypoint_check_correct; eauto. Qed. -Lemma function_equiv_checker_pathentry1 f tfr tpm dm: - function_equiv_checker dm f tfr tpm = OK tt -> +Lemma function_equiv_checker_pathentry1 f tf dm: + function_equiv_checker dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> - path_entry tpm pc2. + path_entry (fn_path tf) pc2. Proof. unfold function_equiv_checker. intros. explore. destruct x0. exploit path_entry_check_correct. eassumption. all: eauto. intuition. Qed. -Lemma function_equiv_checker_pathentry2 f tfr tpm dm: - function_equiv_checker dm f tfr tpm = OK tt -> +Lemma function_equiv_checker_pathentry2 f tf dm: + function_equiv_checker dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1. Proof. @@ -188,9 +196,8 @@ Proof. exploit path_entry_check_correct. eassumption. all: eauto. intuition. Qed. -Lemma function_equiv_checker_correct f tfr tpm dm tf: - function_equiv_checker dm f tfr tpm = OK tt -> - function_defined tf tfr tpm -> +Lemma function_equiv_checker_correct f tf dm: + function_equiv_checker dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. @@ -198,49 +205,33 @@ Proof. eapply simu_check_correct; eauto. Qed. -Definition verified_scheduler (f: RTLpath.function) : res (RTL.function * path_map * (PTree.t node)) := +Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (PTree.t node)) := let (tctetpm, dm) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in let (tc, te) := tcte in let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in - ASSERT RTLpathLivegen.function_checker tfr tpm WITH "In verified_scheduler: (tfr, tpm) is not wellformed" IN - do tt <- function_equiv_checker dm f tfr tpm; - OK (tfr, tpm, dm). - -Lemma verified_scheduler_wellformed_pm f tfr tpm dm: - verified_scheduler f = OK (tfr, tpm, dm) -> - wellformed_path_map (fn_code tfr) tpm. -Proof. - intros. unfold verified_scheduler in H. explore. - apply function_checker_wellformed_path_map. assumption. -Qed. - -Lemma verified_scheduler_path_entry f tfr tpm dm: - verified_scheduler f = OK (tfr, tpm, dm) -> - path_entry tpm (fn_entrypoint tfr). -Proof. - intros. unfold verified_scheduler in H. explore. - apply function_checker_path_entry. assumption. -Qed. - -Theorem verified_scheduler_correct f tfr tf tpm dm: - verified_scheduler f = OK (tfr, tpm, dm) -> - (exists OBL1 OBL2, tf = {| fn_RTL := tfr; fn_path := tpm; fn_entry_point_wf := OBL1; fn_path_wf := OBL2 |}) -> - fn_sig f = fn_sig tfr - /\ fn_params f = fn_params tfr - /\ fn_stacksize f = fn_stacksize tfr - /\ dm ! (fn_entrypoint tfr) = Some (fn_entrypoint f) + do tf <- proj1_sig (function_builder tfr tpm); + do tt <- function_equiv_checker dm f tf; + OK (tf, dm). + +Theorem verified_scheduler_correct f tf dm: + verified_scheduler f = OK (tf, dm) -> + fn_sig f = fn_sig tf + /\ fn_params f = fn_params tf + /\ fn_stacksize f = fn_stacksize tf + /\ dm ! (fn_entrypoint tf) = Some (fn_entrypoint f) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1) - /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry tpm pc2) + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2). Proof. - intros VERIF OBL. unfold verified_scheduler in VERIF. explore. simpl in * |- *. + intros VERIF. unfold verified_scheduler in VERIF. explore. Local Hint Resolve function_equiv_checker_entrypoint function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 function_equiv_checker_correct: core. - destruct x. + destruct x0. destruct (function_builder _ _) as [res H]; simpl in * |- *; auto. + apply H in EQ2. rewrite EQ2. simpl. repeat (constructor; eauto). - - exploit function_equiv_checker_entrypoint. eapply EQ3. intuition. + - exploit function_equiv_checker_entrypoint. eapply EQ4. rewrite EQ2. intuition. Qed. Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := { @@ -257,15 +248,9 @@ Program Definition transf_function (f: RTLpath.function): { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} := match (verified_scheduler f) with | Error e => Error e - | OK (tfr, tpm, dm) => OK {| fn_RTL := tfr; fn_path := tpm |} + | OK (tf, dm) => OK tf end. Next Obligation. - destruct f as [fr pm EP_WF PATH_WF]. - eapply verified_scheduler_path_entry; eauto. -Defined. Next Obligation. - destruct f as [fr pm EP_WF PATH_WF]. - eapply verified_scheduler_wellformed_pm; eauto. -Defined. Next Obligation. exploit verified_scheduler_correct; eauto. intros (A & B & C & D & E & F & G). exists dm. econstructor; eauto. -- cgit From 7d42425213b88bf496b866c05bf4ce35605852b5 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 11 Jun 2020 16:33:01 +0200 Subject: Proof down to sstate_simub and sstate_simub_correct --- kvx/lib/RTLpathScheduler.v | 30 +++++++++++++++++++++++++----- 1 file changed, 25 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index d7f70cd3..4f56e0ac 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -120,14 +120,34 @@ Proof. Qed. (* TODO *) -Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := OK tt. +Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := OK tt. + +Lemma sstate_simub_correct dm f st1 st2: + sstate_simub dm f st1 st2 = OK tt -> + sstate_simu f dm st1 st2. +Proof. +Admitted. + +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := + let (pc2, pc1) := m in + match (sstep f pc1) with + | None => Error (msg "simu_check_single: sstep f pc1 failed") + | Some st1 => + match (sstep tf pc2) with + | None => Error (msg "simu_check_single: sstep tf pc2 failed") + | Some st2 => sstate_simub dm f st1 st2 + end + end. -(* TODO *) Lemma simu_check_single_correct dm tf f pc1 pc2: - simu_check_single dm f tf (pc1, pc2) = OK tt -> + simu_check_single dm f tf (pc2, pc1) = OK tt -> sstep_simu dm f tf pc1 pc2. Proof. -Admitted. + unfold simu_check_single. intro. explore. + unfold sstep_simu. intros st1 STEQ. assert (s = st1) by congruence. subst. + exists s0. constructor; auto. + apply sstate_simub_correct. assumption. +Qed. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with @@ -146,7 +166,7 @@ Proof. - simpl. intuition. - simpl. intros. explore. destruct x. destruct H0. + subst. eapply simu_check_single_correct; eauto. - + eapply IHlm; assumption. + + destruct x0. eapply IHlm; assumption. Qed. Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). -- cgit From 0ade02b1bee73d88efd9c078ffdf0c8380669f5f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 12 Jun 2020 12:39:19 +0200 Subject: Beginning of sstate_simub --- kvx/lib/RTLpathSE_theory.v | 52 +++++++++++++++++++++++++++++++++++++++++++--- kvx/lib/RTLpathScheduler.v | 37 +++++++-------------------------- 2 files changed, 57 insertions(+), 32 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index d8e19295..a54bc410 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -8,6 +8,22 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. +Require Import Errors. + +Local Open Scope error_monad_scope. + +(* Enhanced from kvx/Asmblockgenproof.v *) +Ltac explore := + repeat match goal with + | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in destruct var eqn:EQ1) + | [ H : OK _ = OK _ |- _ ] => monadInv H + | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "EQ" in destruct b eqn:EQ1) + | [ |- context[match ?m with | _ => _ end] ] => destruct m + | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m + | [ H : bind _ _ = OK _ |- _ ] => monadInv H + | [ H : Error _ = OK _ |- _ ] => inversion H + | [ x : unit |- _ ] => destruct x + end. (* NOTE for the implementation of the symbolic execution. It is important to remove dependency of Op on memory. @@ -1463,6 +1479,15 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall rs m is1, ssem ge1 sp st1 rs m is1 -> exists is2, ssem ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. +(* TODO *) +Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := OK tt. + +Lemma sistate_simub_correct dm f st1 st2: + sistate_simub dm f st1 st2 = OK tt -> + sistate_simu f dm st1 st2. +Proof. +Admitted. + (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: @@ -1474,14 +1499,35 @@ Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node | Sreturn_simu os: sfval_simu f srce opc1 opc2 (Sreturn os) (Sreturn os). +(* TODO *) +Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := OK tt. + +Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2: + sfval_simub dm f pc1 pc2 fv1 fv2 = OK tt -> + sfval_simu f dm pc1 pc2 fv1 fv2. +Proof. +Admitted. + Definition sstate_simu f srce (s1 s2: sstate): Prop := sistate_simu f srce s1.(internal) s2.(internal) /\ sfval_simu f srce s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). +(* Computable check of sstate_simu *) +Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := + do u1 <- sistate_simub dm f (internal st1) (internal st2); + do u2 <- sfval_simub dm f (si_pc st1) (si_pc st2) (final st1) (final st2); + OK tt. + +Lemma sstate_simub_correct dm f st1 st2: + sstate_simub dm f st1 st2 = OK tt -> + sstate_simu f dm st1 st2. +Proof. + unfold sstate_simub. intro. explore. constructor. + apply sistate_simub_correct; assumption. + apply sfval_simub_correct; assumption. +Qed. + Definition sstep_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, sstep f1 pc1 = Some st1 -> exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu f1 srce st1 st2. -(* IDEA: we will provide an efficient test for checking "sstep_simu" property, by hash-consing. - See usage of [sstep_simu] in [RTLpathScheduler]. -*) diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 4f56e0ac..2c775d18 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -16,18 +16,6 @@ Notation "'ASSERT' A 'WITH' MSG 'IN' B" := (if A then B else Error (msg MSG)) Local Open Scope error_monad_scope. Local Open Scope positive_scope. -(* Enhanced from kvx/Asmblockgenproof.v *) -Ltac explore := - repeat match goal with - | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in destruct var eqn:EQ1) - | [ H : OK _ = OK _ |- _ ] => monadInv H - | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "EQ" in destruct b eqn:EQ1) - | [ |- context[match ?m with | _ => _ end] ] => destruct m - | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m - | [ H : bind _ _ = OK _ |- _ ] => monadInv H - | [ H : Error _ = OK _ |- _ ] => inversion H - end. - (** External oracle returning the new RTLpath function and a mapping of new path_entries to old path_entries NB: the new RTLpath function is generated from the fn_code, the fn_entrypoint and the fn_path @@ -103,7 +91,7 @@ Lemma path_entry_check_rec_correct pm tpm pc1 pc2: forall lm, Proof. induction lm. - simpl. intuition. - - simpl. intros. explore. destruct x; destruct x0. destruct H0. + - simpl. intros. explore. destruct H0. + subst. eapply path_entry_check_single_correct; eauto. + eapply IHlm; assumption. Qed. @@ -119,15 +107,6 @@ Proof. eapply path_entry_check_rec_correct; eassumption. Qed. -(* TODO *) -Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := OK tt. - -Lemma sstate_simub_correct dm f st1 st2: - sstate_simub dm f st1 st2 = OK tt -> - sstate_simu f dm st1 st2. -Proof. -Admitted. - Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in match (sstep f pc1) with @@ -164,9 +143,9 @@ Lemma simu_check_rec_correct dm f tf pc1 pc2: forall lm, Proof. induction lm. - simpl. intuition. - - simpl. intros. explore. destruct x. destruct H0. + - simpl. intros. explore. destruct H0. + subst. eapply simu_check_single_correct; eauto. - + destruct x0. eapply IHlm; assumption. + + eapply IHlm; assumption. Qed. Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). @@ -194,7 +173,7 @@ Lemma function_equiv_checker_entrypoint f tf dm: function_equiv_checker dm f tf = OK tt -> dm ! (fn_entrypoint tf) = Some (fn_entrypoint f). Proof. - unfold function_equiv_checker. intros. explore. destruct x. + unfold function_equiv_checker. intros. explore. eapply entrypoint_check_correct; eauto. Qed. @@ -203,7 +182,7 @@ Lemma function_equiv_checker_pathentry1 f tf dm: forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2. Proof. - unfold function_equiv_checker. intros. explore. destruct x0. + unfold function_equiv_checker. intros. explore. exploit path_entry_check_correct. eassumption. all: eauto. intuition. Qed. @@ -212,7 +191,7 @@ Lemma function_equiv_checker_pathentry2 f tf dm: forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1. Proof. - unfold function_equiv_checker. intros. explore. destruct x0. + unfold function_equiv_checker. intros. explore. exploit path_entry_check_correct. eassumption. all: eauto. intuition. Qed. @@ -221,7 +200,7 @@ Lemma function_equiv_checker_correct f tf dm: forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2. Proof. - unfold function_equiv_checker. intros. explore. destruct x1. + unfold function_equiv_checker. intros. explore. eapply simu_check_correct; eauto. Qed. @@ -248,7 +227,7 @@ Proof. Local Hint Resolve function_equiv_checker_entrypoint function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 function_equiv_checker_correct: core. - destruct x0. destruct (function_builder _ _) as [res H]; simpl in * |- *; auto. + destruct (function_builder _ _) as [res H]; simpl in * |- *; auto. apply H in EQ2. rewrite EQ2. simpl. repeat (constructor; eauto). - exploit function_equiv_checker_entrypoint. eapply EQ4. rewrite EQ2. intuition. -- cgit From fea8b249f3d573bdaf8a3b515555c3847f0a85cb Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 12 Jun 2020 13:26:58 +0200 Subject: Augmenting sfval_simu ; some corrections to do in sfmatch_simu --- kvx/lib/RTLpathSE_theory.v | 43 +++++++++++++++++++++++++++-------------- kvx/lib/RTLpathSchedulerproof.v | 17 +++++++++------- 2 files changed, 39 insertions(+), 21 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index a54bc410..572a43a9 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1479,7 +1479,6 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall rs m is1, ssem ge1 sp st1 rs m is1 -> exists is2, ssem ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. -(* TODO *) Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := OK tt. Lemma sistate_simub_correct dm f st1 st2: @@ -1488,29 +1487,45 @@ Lemma sistate_simub_correct dm f st1 st2: Proof. Admitted. +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. + (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) -Inductive sfval_simu (f: RTLpath.function) (srce: PTree.t node) (opc1 opc2: node): sfval -> sfval -> Prop := +Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: - srce!opc2 = Some opc1 -> - sfval_simu f srce opc1 opc2 Snone Snone + dm!opc2 = Some opc1 -> + sfval_simu dm f opc1 opc2 Snone Snone | Scall_simu sig svos lsv res pc1 pc2: - srce!pc2 = Some pc1 -> - sfval_simu f srce opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) + dm!pc2 = Some pc1 -> + sfval_simu dm f opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) + | Stailcall_simu sig svos lsv: + sfval_simu dm f opc1 opc2 (Stailcall sig svos lsv) (Stailcall sig svos lsv) + | Sbuiltin_simu ef lbs br pc1 pc2: + dm!pc2 = Some pc1 -> + sfval_simu dm f opc1 opc2 (Sbuiltin ef lbs br pc1) (Sbuiltin ef lbs br pc2) + | Sjumptable_simu sv lpc1 lpc2: + ptree_get_list dm lpc2 = Some lpc1 -> + sfval_simu dm f opc1 opc2 (Sjumptable sv lpc1) (Sjumptable sv lpc2) | Sreturn_simu os: - sfval_simu f srce opc1 opc2 (Sreturn os) (Sreturn os). + sfval_simu dm f opc1 opc2 (Sreturn os) (Sreturn os). (* TODO *) Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := OK tt. Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2: sfval_simub dm f pc1 pc2 fv1 fv2 = OK tt -> - sfval_simu f dm pc1 pc2 fv1 fv2. + sfval_simu dm f pc1 pc2 fv1 fv2. Proof. Admitted. -Definition sstate_simu f srce (s1 s2: sstate): Prop := - sistate_simu f srce s1.(internal) s2.(internal) - /\ sfval_simu f srce s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). +Definition sstate_simu dm f (s1 s2: sstate): Prop := + sistate_simu f dm s1.(internal) s2.(internal) + /\ sfval_simu dm f s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). (* Computable check of sstate_simu *) Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := @@ -1520,14 +1535,14 @@ Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstat Lemma sstate_simub_correct dm f st1 st2: sstate_simub dm f st1 st2 = OK tt -> - sstate_simu f dm st1 st2. + sstate_simu dm f st1 st2. Proof. unfold sstate_simub. intro. explore. constructor. apply sistate_simub_correct; assumption. apply sfval_simub_correct; assumption. Qed. -Definition sstep_simu srce (f1 f2: RTLpath.function) pc1 pc2: Prop := +Definition sstep_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, sstep f1 pc1 = Some st1 -> - exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu f1 srce st1 st2. + exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu dm f1 st1 st2. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 4d7dfa2f..6cbe9c14 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -167,12 +167,12 @@ Proof. Qed. -Lemma sfmatch_simu dupmap f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: - match_function dupmap f f' -> +Lemma sfmatch_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: + match_function dm f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) - sfval_simu f dupmap st.(si_pc) st'.(si_pc) sv sv' -> + sfval_simu dm f st.(si_pc) st'.(si_pc) sv sv' -> sfmatch pge ge sp st stk f rs0 m0 sv rs m t s -> exists s', sfmatch tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. @@ -192,20 +192,23 @@ Proof. eexists; split; econstructor; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. + simpl. repeat (econstructor; eauto). + - (* Stailcall *) admit. + - (* Sbuiltin *) admit. + - (* Sjumptable *) admit. - (* Sreturn *) eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. erewrite <- seval_preserved; eauto. -Qed. +Admitted. (* The main theorem on simulation of symbolic states ! *) -Theorem smatch_sstate_simu dupmap f f' stk stk' sp st st' rs m t s: - match_function dupmap f f' -> +Theorem smatch_sstate_simu dm f f' stk stk' sp st st' rs m t s: + match_function dm f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> smatch pge ge sp st stk f rs m t s -> - sstate_simu f dupmap st st' -> + sstate_simu dm f st st' -> exists s', smatch tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. -- cgit From 5a7db5425cb42746a52d3dcb9c1f6ade3ee1b2a2 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 12 Jun 2020 15:52:54 +0200 Subject: sfmatch_simu: Stailcall and Sjumptable --- kvx/lib/RTLpathSE_theory.v | 33 +++++++++++++++++++++++++++++++++ kvx/lib/RTLpathSchedulerproof.v | 19 ++++++++++++++++--- 2 files changed, 49 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 572a43a9..040afebf 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1495,6 +1495,39 @@ Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list po 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; try discriminate. + inv PGL. 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; try discriminate. + inv PGL. 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: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): sfval -> sfval -> Prop := | Snone_simu: diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 6cbe9c14..4462ab21 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -166,7 +166,6 @@ Proof. intros SEM LIVE X; eapply X; eauto. Qed. - Lemma sfmatch_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dm f f' -> liveness_ok_function f -> @@ -192,9 +191,21 @@ Proof. eexists; split; econstructor; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. + simpl. repeat (econstructor; eauto). - - (* Stailcall *) admit. + - (* Stailcall *) + exploit s_find_function_preserved; eauto. + intros (fd' & FIND & TRANSF & LIVE'). + erewrite <- function_sig_preserved; eauto. + eexists; split; econstructor; eauto. + + erewrite <- preserv_fnstacksize; eauto. + + erewrite <- (list_sval_eval_preserved ge tge); auto. - (* Sbuiltin *) admit. - - (* Sjumptable *) admit. + - (* Sjumptable *) + exploit ptree_get_list_nth_rev; eauto. intros (p2 & LNZ & DM). + exploit initialize_path. { eapply dupmap_path_entry1; eauto. } + intros (path & PATH). + eexists; split; econstructor; eauto. + + erewrite <- (seval_preserved ge tge); eauto. + + eapply eqlive_reg_refl. - (* Sreturn *) eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. @@ -202,6 +213,8 @@ Proof. erewrite <- seval_preserved; eauto. Admitted. + + (* The main theorem on simulation of symbolic states ! *) Theorem smatch_sstate_simu dm f f' stk stk' sp st st' rs m t s: match_function dm f f' -> -- cgit From 64e6fc823ee3f79efb031c0b0c7197f7b49b0075 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 16 Jun 2020 13:29:37 +0200 Subject: Sbuiltin in sfmatch_simu --- kvx/lib/RTLpathSE_theory.v | 35 +++++++++++++++++++++++++++++++++++ kvx/lib/RTLpathSchedulerproof.v | 13 +++++++++---- 2 files changed, 44 insertions(+), 4 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 040afebf..beb7c41a 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1401,6 +1401,21 @@ Variable ge ge': RTL.genv. Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s. +Hypothesis senv_preserved_RTL: Senv.equiv ge ge'. + +Lemma senv_find_symbol_preserved id: + Senv.find_symbol ge id = Senv.find_symbol ge' id. +Proof. + destruct senv_preserved_RTL 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 seval_preserved sp sv rs0 m0: seval_sval ge sp sv rs0 m0 = seval_sval ge' sp sv rs0 m0. Proof. @@ -1423,6 +1438,26 @@ Proof. 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 <- seval_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: seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv rs0 m0. Proof. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 4462ab21..8319a8ff 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -198,7 +198,14 @@ Proof. eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. + erewrite <- (list_sval_eval_preserved ge tge); auto. - - (* Sbuiltin *) admit. + - (* Sbuiltin *) + pose senv_preserved_RTL as SRTL. + exploit initialize_path. { eapply dupmap_path_entry1; eauto. } + intros (path & PATH). + eexists; split; econstructor; eauto. + + eapply seval_builtin_args_preserved; eauto. + + eapply external_call_symbols_preserved; eauto. + + eapply eqlive_reg_refl. - (* Sjumptable *) exploit ptree_get_list_nth_rev; eauto. intros (p2 & LNZ & DM). exploit initialize_path. { eapply dupmap_path_entry1; eauto. } @@ -211,9 +218,7 @@ Proof. + erewrite <- preserv_fnstacksize; eauto. + destruct os; auto. erewrite <- seval_preserved; eauto. -Admitted. - - +Qed. (* The main theorem on simulation of symbolic states ! *) Theorem smatch_sstate_simu dm f f' stk stk' sp st st' rs m t s: -- cgit From 354f8a65972fad54bb01ec295cb4303615261b7b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 17 Jun 2020 16:22:55 +0200 Subject: Stuck in sstep_exact --- kvx/lib/RTLpathSE_theory.v | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index beb7c41a..7e462fe1 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1365,9 +1365,13 @@ Proof. intros NDS; rewrite NDS in Hi. rewrite Hi in SSTEP. intros ISTEPS. try_simplify_someHyps. - destruct (sistep i st0) as [st'|] eqn: Hst'; simpl. + destruct (sistep i st0) as [st'|] eqn:Hst'; simpl. + (* normal exit on Snone instruction *) - admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. + - remember (isteps ge (psize path) f sp rs m pc) as oist0. + intros. eexists. constructor 1. constructor. + + - admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. -- cgit From 41356eb6342a06d2a23bfa19a7133e68408c40e0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 17 Jun 2020 23:42:05 +0200 Subject: 1 cas sur 3 (le seul non absurde ?) --- kvx/lib/RTLpathSE_theory.v | 35 ++++++++++++++++++++++++++++++----- 1 file changed, 30 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 7e462fe1..eb0a0cf7 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1367,11 +1367,36 @@ Proof. intros ISTEPS. try_simplify_someHyps. destruct (sistep i st0) as [st'|] eqn:Hst'; simpl. + (* normal exit on Snone instruction *) - destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - - remember (isteps ge (psize path) f sp rs m pc) as oist0. - intros. eexists. constructor 1. constructor. - - - admit. + assert (SEM': t = E0 /\ exists is, ssem 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 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 sistep_correct; eauto. + rewrite Hst'; simpl. + intros; exploit ssem_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. + cutrewrite ((ipc st) = (if icontinue st then si_pc st' else ipc is)). + { rewrite EQ2, EQ4. eapply State_equiv; auto. } + destruct (icontinue st) eqn:ICONT; auto. + exploit sistep_default_succ; eauto. + erewrite istep_normal_exit; eauto. + try_simplify_someHyps. + ** (* should be an absurd case from ISTEPS and SEM ? *) + admit. + * admit. (* should be an absurd case from ISTEPS and SEM ? *) + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. -- cgit From 0c5dd10a93cd8e6796eaf0d99f25c56b58b4d460 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 18 Jun 2020 00:10:29 +0200 Subject: fix comments ? --- kvx/lib/RTLpathSE_theory.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index eb0a0cf7..cedc540d 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1366,7 +1366,7 @@ Proof. rewrite Hi in SSTEP. intros ISTEPS. try_simplify_someHyps. destruct (sistep i st0) as [st'|] eqn:Hst'; simpl. - + (* normal exit on Snone instruction *) + + (* exit on Snone instruction *) assert (SEM': t = E0 /\ exists is, ssem 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). @@ -1394,9 +1394,10 @@ Proof. exploit sistep_default_succ; eauto. erewrite istep_normal_exit; eauto. try_simplify_someHyps. - ** (* should be an absurd case from ISTEPS and SEM ? *) + ** (* The concrete execution has not reached "i" => early exit *) admit. - * admit. (* should be an absurd case from ISTEPS and SEM ? *) + * (* The concrete execution has not reached "i" => abort case *) + admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. -- cgit From 85819a855b923aff21831e2d4c962d12e4c90968 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 18 Jun 2020 11:22:02 +0200 Subject: Renaming sstep -> sexec ; sistep -> siexec_inst ; sisteps -> siexec_path --- kvx/lib/RTLpathSE_theory.v | 158 ++++++++++++++++++++-------------------- kvx/lib/RTLpathScheduler.v | 22 +++--- kvx/lib/RTLpathSchedulerproof.v | 4 +- 3 files changed, 92 insertions(+), 92 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index cedc540d..4d571374 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -449,7 +449,7 @@ Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. -Definition sistep (i: instruction) (st: sistate): option sistate := +Definition siexec_inst (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => sist_set_local st pc' st.(si_local) @@ -533,8 +533,8 @@ Proof. inv H1; eauto. Qed. -Lemma sistep_preserves_sabort i ge sp rs m st st': - sistep i st = Some st' -> +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. @@ -575,14 +575,14 @@ Proof. - right. exists ext0, lx0. constructor; eauto. eapply all_fallthrough_upto_exit_cons; eauto. Qed. -Lemma sistep_WF i st: - sistep i st = None -> default_succ i = None. +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 sistep_default_succ i st st': - sistep i st = Some st' -> default_succ i = Some (st'.(si_pc)). +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. @@ -605,10 +605,10 @@ Proof. + intros. exploit (ALLR a); [constructor; eauto|]. congruence. Qed. -Lemma sistep_correct ge sp i st rs0 m0 rs m: +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_opt2 ge sp (sistep i st) rs0 m0 (istep ge i sp rs m). + ssem_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. @@ -715,9 +715,9 @@ Proof. Qed. -Lemma sistep_correct_None ge sp i st rs0 m0 rs m: +Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m: ssem_local ge sp (st.(si_local)) rs0 m0 rs m -> - sistep i st = None -> + siexec_inst i st = None -> istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). @@ -725,13 +725,13 @@ Proof. Qed. (** * Symbolic execution of the internal steps of a path *) -Fixpoint sisteps (path:nat) (f: function) (st: sistate): option sistate := +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 <- sistep i st IN - sisteps p f st1 + SOME st1 <- siexec_inst i st IN + siexec_path p f st1 end. Lemma sist_set_local_preserves_si_exits st pc loc st': @@ -741,9 +741,9 @@ Proof. intros. inv H. simpl. reflexivity. Qed. -Lemma sistep_preserves_allfu ge sp ext lx rs0 m0 st st' i: +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 -> - sistep i st = Some st' -> + 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). @@ -752,66 +752,66 @@ Proof. inv SISTEP. simpl. constructor. assumption. Qed. -Lemma sisteps_correct_false ge sp f rs0 m0 st' is: +Lemma siexec_path_correct_false ge sp f rs0 m0 st' is: forall path, is.(icontinue)=false -> forall st, ssem ge sp st rs0 m0 is -> - sisteps path f st = Some st' -> + siexec_path path f st = Some st' -> ssem 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 (sistep _ _) eqn:SISTEP; [|discriminate]. + destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate]. eapply IHpath. 3: eapply STEQ'. eauto. unfold ssem in SSEM. rewrite ICF in SSEM. destruct SSEM as (ext & lx & SEXIT & ALLFU). unfold ssem. rewrite ICF. exists ext, lx. - constructor; auto. eapply sistep_preserves_allfu; eauto. + constructor; auto. eapply siexec_inst_preserves_allfu; eauto. Qed. -Lemma sisteps_preserves_sabort ge sp path f rs0 m0 st': forall st, - sisteps path f st = Some st' -> +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 sistep_preserves_sabort: core. + 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 sisteps_WF path f: forall st, - sisteps path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None. +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 (sistep i st) as [st1|] eqn: Hst1; simpl. - - intros; erewrite sistep_default_succ; eauto. - - intros; erewrite sistep_WF; eauto. + 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 sisteps_default_succ path f st': forall st, - sisteps path f st = Some st' -> nth_default_succ (fn_code f) path st.(si_pc) = Some st'.(si_pc). +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 (sistep i st) as [st1|] eqn: Hst1; simpl; try congruence. - intros; erewrite sistep_default_succ; eauto. + destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl; try congruence. + intros; erewrite siexec_inst_default_succ; eauto. Qed. -Lemma sisteps_correct_true ge sp path (f:function) rs0 m0: forall st is, +Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(icontinue)=true -> ssem ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(si_pc) <> None -> - ssem_opt2 ge sp (sisteps path f st) rs0 m0 + ssem_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 sisteps_correct_false sisteps_preserves_sabort sisteps_WF: core. + 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, sist_set_local in * |- *; @@ -821,7 +821,7 @@ Proof. intros (LOCAL & PC & NYE) WF. rewrite <- PC. inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto. - exploit sistep_correct; eauto. + 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. @@ -829,18 +829,18 @@ Proof. intros; eapply IHpath; eauto. destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps. (* icontinue is0 = false -> EARLY EXIT *) - destruct (sisteps path f st1) as [st2|] eqn: Hst2; simpl; eauto. - destruct WF. erewrite sistep_default_succ; eauto. + 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 (sisteps path f st1) as [st2|] eqn: Hst2; simpl; 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 sisteps_right_assoc_decompose f path: forall st st', - sisteps (S path) f st = Some st' -> - exists st0, sisteps path f st = Some st0 /\ sisteps 1%nat f st0 = Some st'. +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'. @@ -849,10 +849,10 @@ Proof. try_simplify_someHyps; eauto. Qed. -Lemma sisteps_right_assoc_compose f path: forall st st0 st', - sisteps path f st = Some st0 -> - sisteps 1%nat f st0 = Some st' -> - sisteps (S path) f st = Some st'. +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. @@ -1083,7 +1083,7 @@ Lemma sfstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> - sistep i st = None -> + siexec_inst i st = None -> sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. @@ -1108,7 +1108,7 @@ Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s -> - sistep i st = None -> + siexec_inst i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST HSIS. @@ -1154,11 +1154,11 @@ Proof. unfold ssem, ssem_local, all_fallthrough; simpl. intuition. Qed. -Definition sstep (f: function) (pc:node): option sstate := +Definition sexec (f: function) (pc:node): option sstate := SOME path <- (fn_path f)!pc IN - SOME st <- sisteps path.(psize) f (init_sistate pc) IN + SOME st <- siexec_path path.(psize) f (init_sistate pc) IN SOME i <- (fn_code f)!(st.(si_pc)) IN - Some (match sistep i st with + Some (match siexec_inst i st with | Some st' => {| internal := st'; final := Snone |} | None => {| internal := st; final := sfstep i st.(si_local) |} end). @@ -1174,7 +1174,7 @@ 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) -> - sistep i st = Some st' -> + 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 @@ -1184,29 +1184,29 @@ 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 [sstep]) -(sstep is a correct over-approximation) +(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec]) +(sexec is a correct over-approximation) *) -Theorem sstep_correct f pc pge ge sp path stack rs m t s: +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, sstep f pc = Some st /\ smatch pge ge sp st stack f rs m t s. + exists st, sexec f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. Local Hint Resolve init_ssem: core. - intros PATH STEP; unfold sstep; rewrite PATH; simpl. + intros PATH STEP; unfold sexec; rewrite PATH; simpl. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (sisteps_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. + 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_opt2; destruct (sisteps _ _ _) as [st|] eqn: Hst; try congruence); + (rewrite STEPS; unfold ssem_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence); (* intro SEM *) (simpl; unfold ssem; 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 sisteps_default_succ; eauto; simpl; + exploit siexec_path_default_succ; eauto; simpl; intros DEF; rewrite DEF in Hi; auto | clear Hi; rewrite Hi' ]); (* eexists *) @@ -1214,16 +1214,16 @@ Proof. - (* early *) eapply smatch_early; eauto. unfold ssem; simpl; rewrite CONT. - destruct (sistep i st) as [st'|] eqn: Hst'; simpl; eauto. + destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl; eauto. destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx. - constructor; auto. eapply sistep_preserves_allfu; eauto. + constructor; auto. eapply siexec_inst_preserves_allfu; eauto. - destruct SEM as (SEM & PC & HNYE). - destruct (sistep i st) as [st'|] eqn: Hst'; simpl. + 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 sistep_correct; eauto. simpl. + 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. @@ -1340,32 +1340,32 @@ Proof. eapply equiv_state_refl; eauto. Qed. -(* NB: each execution of a symbolic state (produced from [sstep]) represents a concrete execution - (sstep is exact). +(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution + (sexec is exact). *) -Theorem sstep_exact f pc pge ge sp path stack st rs m t s1: +Theorem sexec_exact f pc pge ge sp path stack st rs m t s1: (fn_path f)!pc = Some path -> - sstep f pc = Some st -> + sexec f pc = Some st -> smatch 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: core. - unfold sstep; intros PATH SSTEP SEM; rewrite PATH in SSTEP. + unfold sexec; intros PATH SSTEP SEM; rewrite PATH in SSTEP. lapply (final_node_path_simpl f path pc); eauto. intro WF. - exploit (sisteps_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto. + 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 (sisteps path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl. + destruct (siexec_path path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl. 2:{ (* absurd case *) - exploit sisteps_WF; eauto. + exploit siexec_path_WF; eauto. simpl; intros NDS; rewrite NDS in Hi; congruence. } - exploit sisteps_default_succ; eauto; simpl. + exploit siexec_path_default_succ; eauto; simpl. intros NDS; rewrite NDS in Hi. rewrite Hi in SSTEP. intros ISTEPS. try_simplify_someHyps. - destruct (sistep i st0) as [st'|] eqn:Hst'; simpl. + destruct (siexec_inst i st0) as [st'|] eqn:Hst'; simpl. + (* exit on Snone instruction *) assert (SEM': t = E0 /\ exists is, ssem 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 * |- *. @@ -1380,7 +1380,7 @@ Proof. ** (* icontinue is0=true: path_step by normal_exit *) destruct ISTEPS as (SEMis0&H1&H2). rewrite H1 in * |-. - exploit sistep_correct; eauto. + exploit siexec_inst_correct; eauto. rewrite Hst'; simpl. intros; exploit ssem_opt_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). @@ -1391,7 +1391,7 @@ Proof. cutrewrite ((ipc st) = (if icontinue st then si_pc st' else ipc is)). { rewrite EQ2, EQ4. eapply State_equiv; auto. } destruct (icontinue st) eqn:ICONT; auto. - exploit sistep_default_succ; eauto. + exploit siexec_inst_default_succ; eauto. erewrite istep_normal_exit; eauto. try_simplify_someHyps. ** (* The concrete execution has not reached "i" => early exit *) @@ -1640,7 +1640,7 @@ Proof. apply sfval_simub_correct; assumption. Qed. -Definition sstep_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, sstep f1 pc1 = Some st1 -> - exists st2, sstep f2 pc2 = Some st2 /\ sstate_simu dm f1 st1 st2. +Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := + forall st1, sexec f1 pc1 = Some st1 -> + exists st2, sexec f2 pc2 = Some st2 /\ sstate_simu dm f1 st1 st2. diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 2c775d18..8b96db74 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -109,21 +109,21 @@ Qed. Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in - match (sstep f pc1) with - | None => Error (msg "simu_check_single: sstep f pc1 failed") + match (sexec f pc1) with + | None => Error (msg "simu_check_single: sexec f pc1 failed") | Some st1 => - match (sstep tf pc2) with - | None => Error (msg "simu_check_single: sstep tf pc2 failed") + match (sexec tf pc2) with + | None => Error (msg "simu_check_single: sexec tf pc2 failed") | Some st2 => sstate_simub dm f st1 st2 end end. Lemma simu_check_single_correct dm tf f pc1 pc2: simu_check_single dm f tf (pc2, pc1) = OK tt -> - sstep_simu dm f tf pc1 pc2. + sexec_simu dm f tf pc1 pc2. Proof. unfold simu_check_single. intro. explore. - unfold sstep_simu. intros st1 STEQ. assert (s = st1) by congruence. subst. + unfold sexec_simu. intros st1 STEQ. assert (s = st1) by congruence. subst. exists s0. constructor; auto. apply sstate_simub_correct. assumption. Qed. @@ -139,7 +139,7 @@ Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.fu Lemma simu_check_rec_correct dm f tf pc1 pc2: forall lm, simu_check_rec dm f tf lm = OK tt -> In (pc2, pc1) lm -> - sstep_simu dm f tf pc1 pc2. + sexec_simu dm f tf pc1 pc2. Proof. induction lm. - simpl. intuition. @@ -153,7 +153,7 @@ Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.func Lemma simu_check_correct dm f tf: simu_check dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> - sstep_simu dm f tf pc1 pc2. + sexec_simu dm f tf pc1 pc2. Proof. unfold simu_check. intros. eapply PTree.elements_correct in H0. eapply simu_check_rec_correct; eassumption. @@ -198,7 +198,7 @@ Qed. Lemma function_equiv_checker_correct f tf dm: function_equiv_checker dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> - sstep_simu dm f tf pc1 pc2. + sexec_simu dm f tf pc1 pc2. Proof. unfold function_equiv_checker. intros. explore. eapply simu_check_correct; eauto. @@ -221,7 +221,7 @@ Theorem verified_scheduler_correct f tf dm: /\ dm ! (fn_entrypoint tf) = Some (fn_entrypoint f) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2) - /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sstep_simu dm f tf pc1 pc2). + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2). Proof. intros VERIF. unfold verified_scheduler in VERIF. explore. Local Hint Resolve function_equiv_checker_entrypoint @@ -240,7 +240,7 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sstep_simu dupmap f1 f2 pc1 pc2 + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sexec_simu dupmap f1 f2 pc1 pc2 }. Program Definition transf_function (f: RTLpath.function): diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 8319a8ff..bec13681 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -265,13 +265,13 @@ Proof. exploit initialize_path. { eapply dupmap_path_entry2; eauto. } intros (path' & PATH'). exists path'. - exploit (sstep_correct f pc pge ge sp path stk rs m t s); eauto. + exploit (sexec_correct f pc pge ge sp path stk rs m t s); eauto. intros (st & SYMB & SEM); clear STEP. exploit dupmap_correct; eauto. clear SYMB; intros (st' & SYMB & SIMU). exploit smatch_sstate_simu; eauto. intros (s0 & SEM0 & MATCH). - exploit sstep_exact; eauto. + exploit sexec_exact; eauto. intros (s' & STEP' & EQ). exists s'; intuition. eapply match_states_equiv; eauto. -- cgit From 95d44160f22177dd46703634bd868dd0a226ee35 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 18 Jun 2020 11:30:25 +0200 Subject: sfstep -> sexec_final --- kvx/lib/RTLpathSE_theory.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 4d571374..04b43587 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1056,7 +1056,7 @@ Proof. Qed. (** * Symbolic execution of final step *) -Definition sfstep (i: instruction) (prev: sistate_local): sfval := +Definition sexec_final (i: instruction) (prev: sistate_local): sfval := match i with | Icall sig ros args res pc => let svos := sum_left_map prev.(si_sreg) ros in @@ -1078,13 +1078,13 @@ Definition sfstep (i: instruction) (prev: sistate_local): sfval := | _ => Snone end. -Lemma sfstep_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: +Lemma sexec_final_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> siexec_inst i st = None -> - sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s. + sfmatch pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + (* Snone *) destruct i; simpl in Hi |- *; unfold sist_set_local in Hi; try congruence. @@ -1103,11 +1103,11 @@ Proof. destruct or; simpl; auto. Qed. -Lemma sfstep_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: +Lemma sexec_final_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> - sfmatch pge ge sp st stack f rs0 m0 (sfstep i (si_local st)) rs m t s -> + sfmatch pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s -> siexec_inst i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. @@ -1160,7 +1160,7 @@ Definition sexec (f: function) (pc:node): option sstate := 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 := sfstep i st.(si_local) |} + | None => {| internal := st; final := sexec_final i st.(si_local) |} end). Lemma final_node_path_simpl f path pc: @@ -1238,7 +1238,7 @@ Proof. + (* normal non-Snone instruction *) eapply smatch_normal; eauto. * unfold ssem; simpl; rewrite CONT; intuition. - * simpl. eapply sfstep_correct; eauto. + * simpl. eapply sexec_final_correct; eauto. rewrite PC; auto. Qed. @@ -1418,7 +1418,7 @@ Proof. rewrite ! EQ2 in * |-; clear EQ2. exists s'; intuition. eapply exec_normal_exit; eauto. - eapply sfstep_complete; eauto. + eapply sexec_final_complete; eauto. * congruence. * unfold ssem_local in * |- *. intuition try congruence. Admitted. -- cgit From 8d32dc15f91896bec6158b216f985896706dfa8d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 18 Jun 2020 11:45:46 +0200 Subject: ssem --> ssem_internal --- kvx/lib/RTLpathSE_theory.v | 124 ++++++++++++++++++++-------------------- kvx/lib/RTLpathSchedulerproof.v | 4 +- 2 files changed, 64 insertions(+), 64 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 04b43587..ff36bd65 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -165,7 +165,7 @@ Definition all_fallthrough_upto_exit ge sp ext lx' lx rs m : Prop := (* Note: in RTLpath, is.(icontinue) = false iff we took an early exit *) -Definition ssem (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop := +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) @@ -181,21 +181,21 @@ Definition sabort (ge: RTL.genv) (sp: val) (st: sistate) (rs: regset) (m: mem): (* 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_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := +Definition ssem_internal_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop := match ois with - | Some is => ssem ge sp st rs0 m0 is + | Some is => ssem_internal ge sp st rs0 m0 is | None => sabort ge sp st rs0 m0 end. -Definition ssem_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := +Definition ssem_internal_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop := match ost with - | Some st => ssem_opt ge sp st rs0 m0 ois + | 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_opt] is deterministic. + We prove below that the semantics [ssem_internal_opt] is deterministic. *) @@ -218,15 +218,15 @@ Proof. discriminate. Qed. -Lemma ssem_exclude_incompatible_continue ge sp st rs m is1 is2: +Lemma ssem_internal_exclude_incompatible_continue ge sp st rs m is1 is2: is1.(icontinue) = true -> is2.(icontinue) = false -> - ssem ge sp st rs m is1 -> - ssem ge sp st rs m is2 -> + 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. + unfold ssem_internal. intros CONT1 CONT2. rewrite CONT1, CONT2; simpl. intuition eauto. @@ -235,12 +235,12 @@ Proof. eapply all_fallthrough_noexit; eauto. Qed. -Lemma ssem_determ_continue ge sp st rs m is1 is2: - ssem ge sp st rs m is1 -> - ssem ge sp st rs m is2 -> +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_exclude_incompatible_continue: core. + 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. @@ -303,15 +303,15 @@ Proof. + eapply is_tail_cons_left; eauto. Qed. -Lemma ssem_determ ge sp st rs m is1 is2: - ssem ge sp st rs m is1 -> - ssem ge sp st rs m is2 -> +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_determ_continue ge sp st rs m is1 is2); eauto. - intros CONTEQ. unfold ssem in * |-. rewrite CONTEQ in * |- *. + 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). @@ -409,12 +409,12 @@ Proof. + destruct H0 as (INE & TAIL). eapply ALLFU2 in INE. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence. Qed. -Lemma ssem_exclude_sabort ge sp st rs m is: +Lemma ssem_internal_exclude_sabort ge sp st rs m is: sabort ge sp st rs m -> - ssem ge sp st rs m is -> False. + ssem_internal ge sp st rs m is -> False. Proof. intros ABORT SEM. - unfold ssem in SEM. destruct icontinue. + 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. @@ -423,14 +423,14 @@ Qed. Definition istate_eq_opt ist1 oist := exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2. -Lemma ssem_opt_determ ge sp st rs m ois is: - ssem_opt ge sp st rs m ois -> - ssem ge sp st rs m is -> +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_determ; eauto. - - intros; exploit ssem_exclude_sabort; eauto. destruct 1. + - intros; eexists; intuition; eapply ssem_internal_determ; eauto. + - intros; exploit ssem_internal_exclude_sabort; eauto. destruct 1. Qed. (** * Symbolic execution of one internal step *) @@ -608,7 +608,7 @@ 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_opt2 ge sp (siexec_inst i st) rs0 m0 (istep ge i sp rs m). + 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. @@ -633,14 +633,14 @@ Proof. + (* LOAD *) inversion_SOME a0; intro ADD. { inversion_SOME v; intros LOAD; simpl. - - explore_destruct; unfold ssem, ssem_local; simpl; intuition. - * unfold ssem. simpl. constructor; [|constructor]; auto. + - 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 seval_list_sval_inj; simpl; auto. try_simplify_someHyps. - * unfold ssem. simpl. constructor; [|constructor]; auto. + * 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. @@ -653,7 +653,7 @@ Proof. right. right. exists r. simpl. destruct (Pos.eq_dec r r); try congruence. simpl. erewrite seval_list_sval_inj; simpl; auto. rewrite ADD; simpl; auto. try_simplify_someHyps. - * unfold ssem. simpl. constructor; [|constructor]; auto. + * 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. @@ -676,7 +676,7 @@ Proof. + (* STORE *) inversion_SOME a0; intros ADD. { inversion_SOME m'; intros STORE; simpl. - - unfold ssem, ssem_local; simpl; intuition. + - unfold ssem_internal, ssem_local; simpl; intuition. * congruence. * erewrite seval_list_sval_inj; simpl; auto. erewrite REG. @@ -694,7 +694,7 @@ Proof. Local Hint Resolve is_tail_refl: core. Local Hint Unfold ssem_local: core. inversion_SOME b; intros COND. - { destruct b; simpl; unfold ssem, ssem_local; simpl. + { 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. @@ -721,7 +721,7 @@ Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m: istep ge i sp rs m = None. Proof. intros (PRE & MEM & REG). - destruct i; simpl; unfold sist_set_local, ssem, ssem_local; simpl; try_simplify_someHyps. + 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 *) @@ -755,9 +755,9 @@ Qed. Lemma siexec_path_correct_false ge sp f rs0 m0 st' is: forall path, is.(icontinue)=false -> - forall st, ssem ge sp st rs0 m0 is -> + forall st, ssem_internal ge sp st rs0 m0 is -> siexec_path path f st = Some st' -> - ssem ge sp st' rs0 m0 is. + ssem_internal ge sp st' rs0 m0 is. Proof. induction path; simpl. - intros. congruence. @@ -765,9 +765,9 @@ Proof. destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate]. destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate]. eapply IHpath. 3: eapply STEQ'. eauto. - unfold ssem in SSEM. rewrite ICF in SSEM. + unfold ssem_internal in SSEM. rewrite ICF in SSEM. destruct SSEM as (ext & lx & SEXIT & ALLFU). - unfold ssem. rewrite ICF. exists ext, lx. + unfold ssem_internal. rewrite ICF. exists ext, lx. constructor; auto. eapply siexec_inst_preserves_allfu; eauto. Qed. @@ -805,19 +805,19 @@ Qed. Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is, is.(icontinue)=true -> - ssem ge sp st rs0 m0 is -> + ssem_internal ge sp st rs0 m0 is -> nth_default_succ (fn_code f) path st.(si_pc) <> None -> - ssem_opt2 ge sp (siexec_path path f st) rs0 m0 + 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, sist_set_local in * |- *; + unfold ssem_internal, sist_set_local in * |- *; try_simplify_someHyps. simpl. destruct is; simpl in * |- *; subst; intuition auto. - + intros st is CONT; unfold ssem at 1; rewrite CONT. + + 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. @@ -981,11 +981,11 @@ Record sstate := { internal:> sistate; final: sfval }. Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := | smatch_early is: is.(icontinue) = false -> - ssem ge sp st rs0 m0 is -> + ssem_internal ge sp st rs0 m0 is -> smatch pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) | smatch_normal is t s: is.(icontinue) = true -> - ssem ge sp st rs0 m0 is -> + ssem_internal ge sp st rs0 m0 is -> sfmatch pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> smatch pge ge sp st stack f rs0 m0 t s . @@ -1149,9 +1149,9 @@ Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}. -Lemma init_ssem ge sp pc rs m: ssem ge sp (init_sistate pc) rs m (mk_istate true pc rs m). +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, ssem_local, all_fallthrough; simpl. intuition. + unfold ssem_internal, ssem_local, all_fallthrough; simpl. intuition. Qed. Definition sexec (f: function) (pc:node): option sstate := @@ -1192,7 +1192,7 @@ Theorem sexec_correct f pc pge ge sp path stack rs m t s: path_step ge pge path.(psize) stack f sp rs m pc t s -> exists st, sexec f pc = Some st /\ smatch pge ge sp st stack f rs m t s. Proof. - Local Hint Resolve init_ssem: core. + 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. @@ -1200,9 +1200,9 @@ Proof. (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_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence); + (rewrite STEPS; unfold ssem_internal_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence); (* intro SEM *) - (simpl; unfold ssem; simpl; rewrite CONT; 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; @@ -1213,7 +1213,7 @@ Proof. (eexists; constructor; eauto). - (* early *) eapply smatch_early; eauto. - unfold ssem; simpl; rewrite CONT. + 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. @@ -1229,7 +1229,7 @@ Proof. intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT. { (* icontinue mk_istate = true *) eapply smatch_normal; simpl; eauto. - unfold ssem in SEM. + unfold ssem_internal in SEM. rewrite CONT in SEM. destruct SEM as (SEM & PC & HNYE). rewrite <- PC. @@ -1237,7 +1237,7 @@ Proof. { eapply smatch_early; eauto. } + (* normal non-Snone instruction *) eapply smatch_normal; eauto. - * unfold ssem; simpl; rewrite CONT; intuition. + * unfold ssem_internal; simpl; rewrite CONT; intuition. * simpl. eapply sexec_final_correct; eauto. rewrite PC; auto. Qed. @@ -1350,7 +1350,7 @@ Theorem sexec_exact f pc pge ge sp path stack st 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: core. + 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. @@ -1367,7 +1367,7 @@ Proof. 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 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))). + 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. @@ -1376,13 +1376,13 @@ Proof. 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 in ISTEPS. destruct (icontinue is0) eqn: ICONT0. + * 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_opt_determ; eauto. + intros; exploit ssem_internal_opt_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). eexists. econstructor 1. *** eapply exec_normal_exit; eauto. @@ -1401,16 +1401,16 @@ Proof. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. - (* early exit *) intros. - exploit ssem_opt_determ; eauto. + 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_opt_determ; eauto. + exploit ssem_internal_opt_determ; eauto. destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4). - unfold ssem in SEM1. + unfold ssem_internal in SEM1. rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0). exploit sfmatch_equiv; eauto. clear SEM2; destruct 1 as (s' & Ms' & SEM2). @@ -1541,8 +1541,8 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> - forall rs m is1, ssem ge1 sp st1 rs m is1 -> - exists is2, ssem ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. + forall rs m is1, ssem_internal ge1 sp st1 rs m is1 -> + exists is2, ssem_internal ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := OK tt. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index bec13681..8bb601ed 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -157,11 +157,11 @@ Proof. Qed. Lemma sistate_simu f dupmap sp st st' rs m is: - ssem ge sp st rs m is -> + ssem_internal ge sp st rs m is -> liveness_ok_function f -> sistate_simu f dupmap st st' -> exists is', - ssem tge sp st' rs m is' /\ istate_simu f dupmap is is'. + ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. intros SEM LIVE X; eapply X; eauto. Qed. -- cgit From f90a270150e882c1b83a5ea778202d14b37ddf00 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 18 Jun 2020 11:48:05 +0200 Subject: sfmatch -> sstep_final ; smatch -> ssem --- kvx/lib/RTLpathSE_theory.v | 50 ++++++++++++++++++++--------------------- kvx/lib/RTLpathSchedulerproof.v | 22 +++++++++--------- 2 files changed, 36 insertions(+), 36 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index ff36bd65..45e91853 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -941,14 +941,14 @@ Qed. End SEVAL_BUILTIN_ARG. -Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive ssem_final (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: - sfmatch pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(si_pc) rs m) + ssem_final pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(si_pc) rs m) | exec_Scall rs m sig svos lsv args res pc fd: sfind_function pge ge sp svos rs0 m0 = Some fd -> funsig fd = sig -> seval_list_sval ge sp lsv rs0 m0 = Some args -> - sfmatch pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + ssem_final pge ge sp st 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 pge ge sp svos rs0 m0 = Some fd -> @@ -956,38 +956,38 @@ Inductive sfmatch (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stac sp = Vptr stk Ptrofs.zero -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> seval_list_sval ge sp lsv rs0 m0 = Some args -> - sfmatch pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m + ssem_final pge ge sp st 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' -> - sfmatch pge ge sp st stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m + ssem_final pge ge sp st 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: seval_sval ge sp sv rs0 m0 = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - sfmatch pge ge sp st stack f rs0 m0 (Sjumptable sv tbl) rs m + ssem_final pge ge sp st stack f rs0 m0 (Sjumptable sv tbl) 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 => seval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> - sfmatch pge ge sp st stack f rs0 m0 (Sreturn osv) rs m + ssem_final pge ge sp st stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . Record sstate := { internal:> sistate; final: sfval }. -Inductive smatch pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := - | smatch_early is: +Inductive ssem pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop := + | ssem_early is: is.(icontinue) = false -> ssem_internal ge sp st rs0 m0 is -> - smatch pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) - | smatch_normal is t s: + ssem pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem)) + | ssem_normal is t s: is.(icontinue) = true -> ssem_internal ge sp st rs0 m0 is -> - sfmatch pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> - smatch pge ge sp st stack f rs0 m0 t s + ssem_final pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> + ssem pge ge sp st stack f rs0 m0 t s . Fixpoint builtin_arg_inj (f: reg -> sval) (arg: builtin_arg reg) : builtin_arg sval := @@ -1084,7 +1084,7 @@ Lemma sexec_final_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: ssem_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> siexec_inst i st = None -> - sfmatch pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s. + ssem_final pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + (* Snone *) destruct i; simpl in Hi |- *; unfold sist_set_local in Hi; try congruence. @@ -1107,7 +1107,7 @@ Lemma sexec_final_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> - sfmatch pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s -> + ssem_final pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s -> siexec_inst i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. @@ -1190,7 +1190,7 @@ Qed. 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 /\ smatch pge ge sp st stack f rs m 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. @@ -1212,7 +1212,7 @@ Proof. (* eexists *) (eexists; constructor; eauto). - (* early *) - eapply smatch_early; eauto. + 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. @@ -1228,15 +1228,15 @@ Proof. 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 smatch_normal; simpl; eauto. + eapply ssem_normal; simpl; eauto. unfold ssem_internal in SEM. rewrite CONT in SEM. destruct SEM as (SEM & PC & HNYE). rewrite <- PC. eapply exec_Snone. } - { eapply smatch_early; eauto. } + { eapply ssem_early; eauto. } + (* normal non-Snone instruction *) - eapply smatch_normal; eauto. + eapply ssem_normal; eauto. * unfold ssem_internal; simpl; rewrite CONT; intuition. * simpl. eapply sexec_final_correct; eauto. rewrite PC; auto. @@ -1311,10 +1311,10 @@ Proof. - repeat (rewrite Regmap.gso); auto. Qed. -Lemma sfmatch_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: - sfmatch pge ge sp st stack f rs0 m0 sv rs1 m t s -> +Lemma ssem_final_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s: + ssem_final 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' /\ sfmatch pge ge sp st stack f rs0 m0 sv rs2 m t s'. + exists s', equiv_state s s' /\ ssem_final pge ge sp st stack f rs0 m0 sv rs2 m t s'. Proof. Local Hint Resolve equiv_stack_refl: core. destruct 1. @@ -1346,7 +1346,7 @@ Qed. 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 -> - smatch pge ge sp st stack f rs m t s1 -> + 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. @@ -1412,7 +1412,7 @@ Proof. 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 sfmatch_equiv; eauto. + exploit ssem_final_equiv; eauto. clear SEM2; destruct 1 as (s' & Ms' & SEM2). rewrite ! EQ4 in * |-; clear EQ4. rewrite ! EQ2 in * |-; clear EQ2. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 8bb601ed..7599a439 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -166,14 +166,14 @@ Proof. intros SEM LIVE X; eapply X; eauto. Qed. -Lemma sfmatch_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: +Lemma ssem_internal_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dm f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) sfval_simu dm f st.(si_pc) st'.(si_pc) sv sv' -> - sfmatch pge ge sp st stk f rs0 m0 sv rs m t s -> - exists s', sfmatch tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + ssem_internal pge ge sp st stk f rs0 m0 sv rs m t s -> + exists s', ssem_internal tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct: core. intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. @@ -221,13 +221,13 @@ Proof. Qed. (* The main theorem on simulation of symbolic states ! *) -Theorem smatch_sstate_simu dm f f' stk stk' sp st st' rs m t s: +Theorem ssem_sstate_simu dm f f' stk stk' sp st st' rs m t s: match_function dm f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> - smatch pge ge sp st stk f rs m t s -> + ssem pge ge sp st stk f rs m t s -> sstate_simu dm f st st' -> - exists s', smatch tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. + exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) @@ -236,19 +236,19 @@ Proof. intros (is' & SEM' & (path & PATH & (CONT' & RS' & M') & PC')). exists (State stk' f' sp (ipc is') (irs is') (imem is')). split. - + eapply smatch_early; auto. congruence. + + eapply ssem_early; auto. congruence. + rewrite M'. econstructor; eauto. - (* sem_normal *) exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. intros (is' & SEM' & (CONT' & RS' & M')). rewrite <- eqlive_reg_triv in RS'. - exploit sfmatch_simu; eauto. + exploit ssem_internal_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). - exploit sfmatch_equiv; eauto. + exploit ssem_internal_equiv; eauto. clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s1 & EQ & SEM2). exists s1; split. - + eapply smatch_normal; eauto. + + eapply ssem_normal; eauto. + eapply match_states_equiv; eauto. Qed. @@ -269,7 +269,7 @@ Proof. intros (st & SYMB & SEM); clear STEP. exploit dupmap_correct; eauto. clear SYMB; intros (st' & SYMB & SIMU). - exploit smatch_sstate_simu; eauto. + exploit ssem_sstate_simu; eauto. intros (s0 & SEM0 & MATCH). exploit sexec_exact; eauto. intros (s' & STEP' & EQ). -- cgit From 4954c3841398bbd9bd42aae2f0d565f0b3813ed4 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 18 Jun 2020 11:52:45 +0200 Subject: Fix typo --- kvx/lib/RTLpathSchedulerproof.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 7599a439..38d44ff0 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -166,14 +166,14 @@ Proof. intros SEM LIVE X; eapply X; eauto. Qed. -Lemma ssem_internal_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: +Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: match_function dm f f' -> liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) sfval_simu dm f st.(si_pc) st'.(si_pc) sv sv' -> - ssem_internal pge ge sp st stk f rs0 m0 sv rs m t s -> - exists s', ssem_internal tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + ssem_final pge ge sp st stk f rs0 m0 sv rs m t s -> + exists s', ssem_final tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct: core. intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. @@ -243,9 +243,9 @@ Proof. unfold istate_simu; rewrite CONT. intros (is' & SEM' & (CONT' & RS' & M')). rewrite <- eqlive_reg_triv in RS'. - exploit ssem_internal_simu; eauto. + exploit ssem_final_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). - exploit ssem_internal_equiv; eauto. + exploit ssem_final_equiv; eauto. clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s1 & EQ & SEM2). exists s1; split. + eapply ssem_normal; eauto. -- cgit From 131851fe54b6b805ff169b31bf5c5e67597f1a5a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 19 Jun 2020 15:44:43 +0200 Subject: Proof of 2nd admitted of sexec_exact --- kvx/lib/RTLpathSE_theory.v | 102 +++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 99 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 45e91853..ef5603fb 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -741,6 +741,15 @@ Proof. intros. inv H. simpl. reflexivity. Qed. +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; try discriminate. + all: try (intros; left; eapply sist_set_local_preserves_si_exits; eauto; fail). + intros. inv H. right. simpl. 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' -> @@ -1340,6 +1349,78 @@ Proof. 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. + +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). *) @@ -1367,7 +1448,8 @@ Proof. 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))). + 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. @@ -1388,14 +1470,27 @@ Proof. *** eapply exec_normal_exit; eauto. eapply exec_istate; eauto. *** rewrite EQ1. - cutrewrite ((ipc st) = (if icontinue st then si_pc st' else ipc is)). + 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 *) - admit. + 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 *) admit. + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *. @@ -1423,6 +1518,7 @@ Proof. * unfold ssem_local in * |- *. intuition try congruence. Admitted. + (** * Simulation of RTLpath code w.r.t symbolic execution *) Section SymbValPreserved. -- cgit From f847df7db2271cbeb0224373490b5f0c5b415c71 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 19 Jun 2020 15:56:45 +0200 Subject: Proof of sexec_exact --- kvx/lib/RTLpathSE_theory.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index ef5603fb..a23a75ba 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1492,7 +1492,8 @@ Proof. 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 *) - admit. + 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. @@ -1516,8 +1517,7 @@ Proof. eapply sexec_final_complete; eauto. * congruence. * unfold ssem_local in * |- *. intuition try congruence. -Admitted. - +Qed. (** * Simulation of RTLpath code w.r.t symbolic execution *) -- cgit From e608bbe2676926fc8e81b04b8e261d84cea87b9a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 22 Jun 2020 13:42:39 +0200 Subject: Splitting sistate_simub --- kvx/lib/RTLpathSE_theory.v | 74 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 72 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index a23a75ba..40b410d4 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1640,12 +1640,82 @@ Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sis forall rs m is1, ssem_internal ge1 sp st1 rs m is1 -> exists is2, ssem_internal ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. -Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := OK tt. +Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local): Prop := + forall sp ge1 ge2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + liveness_ok_function f -> + forall rs m is1, ssem_local ge1 sp sl1 rs m (irs is1) (imem is1) -> + exists is2, ssem_local ge2 sp sl2 rs m (irs is2) (imem is2) /\ istate_simu f dm is1 is2. + +(* TODO *) +Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) := OK tt. + +Lemma silocal_simub_correct dm f sl1 sl2: + silocal_simub dm f sl1 sl2 = OK tt -> + silocal_simu dm f sl1 sl2. +Proof. +Admitted. + +Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := + forall sp ge1 ge2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + liveness_ok_function f -> + forall rs m is1, ssem_exit ge1 sp se1 rs m (irs is1) (imem is1) (ipc is1) -> + exists is2, + ssem_exit ge2 sp se2 rs m (irs is2) (imem is2) (ipc is2) + /\ istate_simu f dm is1 is2. + +(* TODO *) +Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := OK tt. + +Lemma siexit_simub_correct dm f se1 se2: + siexit_simub dm f se1 se2 = OK tt -> + siexit_simu dm f se1 se2. +Proof. +Admitted. + +Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := + forall sp ge1 ge2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + liveness_ok_function f -> + forall rs m is1, + (exists ext1 lx1, ssem_exit ge1 sp ext1 rs m (irs is1) (imem is1) (ipc is1) /\ + all_fallthrough_upto_exit ge1 sp ext1 lx1 lse1 rs m) -> + exists is2, + (exists ext2 lx2, ssem_exit ge2 sp ext2 rs m (irs is2) (imem is2) (ipc is2) /\ + all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m) + /\ istate_simu f dm is1 is2. + +Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := + match lse1 with + | nil => OK tt + | se1 :: lse1 => + match lse2 with + | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") + | se2 :: lse2 => + do _ <- siexit_simub dm f se1 se2; + do _ <- siexits_simub dm f lse1 lse2; + OK tt + end + end. + +Lemma siexits_simub_correct dm f lse1 lse2: + siexits_simub dm f lse1 lse2 = OK tt -> + siexits_simu dm f lse1 lse2. +Proof. +Admitted. + +Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := + do _ <- silocal_simub dm f (si_local st1) (si_local st2); + do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); + OK tt. Lemma sistate_simub_correct dm f st1 st2: sistate_simub dm f st1 st2 = OK tt -> sistate_simu f dm st1 st2. Proof. +(* unfold sistate_simub. intro. explore. + exploit silocal_simub_correct; eauto. *) Admitted. Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := @@ -1731,7 +1801,7 @@ Lemma sstate_simub_correct dm f st1 st2: sstate_simub dm f st1 st2 = OK tt -> sstate_simu dm f st1 st2. Proof. - unfold sstate_simub. intro. explore. constructor. + unfold sstate_simub. intros. explore. constructor. apply sistate_simub_correct; assumption. apply sfval_simub_correct; assumption. Qed. -- cgit From cdd61d05d567c9e74e4ad49fcb61cbf8ee32d8d0 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 22 Jun 2020 15:52:45 +0200 Subject: Start of sistate_simub_correct --- kvx/lib/RTLpathSE_theory.v | 59 ++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 54 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 40b410d4..3cf7dc4a 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -136,6 +136,17 @@ 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 @@ -1656,7 +1667,7 @@ Lemma silocal_simub_correct dm f sl1 sl2: Proof. Admitted. -Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := +Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> @@ -1665,6 +1676,16 @@ Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistat ssem_exit ge2 sp se2 rs m (irs is2) (imem is2) (ipc is2) /\ istate_simu f dm is1 is2. +Definition siexit_simu_cond (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := + forall sp ge1 ge2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + forall rs m, + seval_condition ge1 sp (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) rs m + = seval_condition ge2 sp (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) rs m. + +Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := + siexit_simu_ssem dm f se1 se2 /\ siexit_simu_cond dm f se1 se2. + (* TODO *) Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := OK tt. @@ -1675,7 +1696,23 @@ Proof. Admitted. Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := - forall sp ge1 ge2, + list_forall2 (siexit_simu dm f) lse1 lse2. + +Lemma siexits_simu_all_fallthrough dm f ge1 ge2 sp rs m: forall lse1 lse2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + all_fallthrough ge1 sp lse1 rs m -> + siexits_simu dm f lse1 lse2 -> + all_fallthrough ge2 sp lse2 rs m. +Proof. + induction 3; unfold all_fallthrough; [contradiction|]. + intros ext INEXT. apply all_fallthrough_revcons in H0. destruct H0 as (SEVAL & ALLFU). + apply IHlist_forall2 in ALLFU. + destruct H1 as (_ & CONDSIMU). + inv INEXT; [|eauto]. + erewrite <- CONDSIMU; eauto. +Qed. + +(* forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> forall rs m is1, @@ -1684,7 +1721,7 @@ Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: lis exists is2, (exists ext2 lx2, ssem_exit ge2 sp ext2 rs m (irs is2) (imem is2) (ipc is2) /\ all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m) - /\ istate_simu f dm is1 is2. + /\ istate_simu f dm is1 is2. *) Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := match lse1 with @@ -1714,8 +1751,20 @@ Lemma sistate_simub_correct dm f st1 st2: sistate_simub dm f st1 st2 = OK tt -> sistate_simu f dm st1 st2. Proof. -(* unfold sistate_simub. intro. explore. - exploit silocal_simub_correct; eauto. *) + unfold sistate_simub. intro. explore. unfold sistate_simu. + intros sp ge1 ge2 GFS LOK rs m is1 SEMI. + exploit siexits_simub_correct; eauto. intro ESIMU. + unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. + - destruct SEMI as (SSEML & PCEQ & ALLFU). + exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). + exists is2. constructor; auto. unfold ssem_internal. + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. + destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). + rewrite <- CONTEQ. rewrite ICONT. + repeat (constructor; auto). + + admit. (* TODO - not sure what to do yet with that one *) + + eapply siexits_simu_all_fallthrough; eauto. + - admit. Admitted. Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := -- cgit From c85e89c22fd372f8e6880e365123e502417b2ba9 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 22 Jun 2020 17:33:25 +0200 Subject: Modification on istate_simu to go further in the proof of sistate_simub_correct --- kvx/lib/RTLpathSE_theory.v | 28 +++++++++++++++++++++++++--- kvx/lib/RTLpathScheduler.v | 19 +++++++++++++++---- kvx/lib/RTLpathSchedulerproof.v | 9 ++++++--- 3 files changed, 46 insertions(+), 10 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 3cf7dc4a..fc330dfd 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1638,16 +1638,21 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := Typiquement, pour connaître la frame des registres vivants, il faudra faire une propagation en arrière sur la dernière instruction du superblock. *) istate_simulive (fun _ => True) srce is1 is2 + /\ srce!(is2.(ipc)) = Some is1.(ipc) 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). +Definition is_injective (dm: PTree.t node): Prop := + forall x y, dm!x = dm!y -> x = y. + (* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sistate): Prop := forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> + is_injective srce -> forall rs m is1, ssem_internal ge1 sp st1 rs m is1 -> exists is2, ssem_internal ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. @@ -1742,7 +1747,22 @@ Lemma siexits_simub_correct dm f lse1 lse2: Proof. Admitted. +Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := + match dm ! tn with + | None => Error (msg "revmap_check_single: no mapping for tn") + | Some res => if (Pos.eq_dec n res) then OK tt + else Error (msg "revmap_check_single: n and res do not match") + end. + +Lemma revmap_check_single_correct dm n tn: + revmap_check_single dm n tn = OK tt -> + dm ! tn = Some n. +Proof. + unfold revmap_check_single. explore; try discriminate. congruence. +Qed. + Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := + do _ <- revmap_check_single dm (si_pc st1) (si_pc st2); do _ <- silocal_simub dm f (si_local st1) (si_local st2); do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); OK tt. @@ -1752,17 +1772,19 @@ Lemma sistate_simub_correct dm f st1 st2: sistate_simu f dm st1 st2. Proof. unfold sistate_simub. intro. explore. unfold sistate_simu. - intros sp ge1 ge2 GFS LOK rs m is1 SEMI. + intros sp ge1 ge2 GFS LOK INJ rs m is1 SEMI. exploit siexits_simub_correct; eauto. intro ESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). exists is2. constructor; auto. unfold ssem_internal. unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. - destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). + destruct ISIMU as ((CONTEQ & RSEQ & MEMEQ) & REVBIND). rewrite <- CONTEQ. rewrite ICONT. repeat (constructor; auto). - + admit. (* TODO - not sure what to do yet with that one *) + + eapply revmap_check_single_correct in EQ. + assert (DMEQ: dm ! (si_pc st2) = dm ! (ipc is2)) by congruence. + apply INJ. assumption. + eapply siexits_simu_all_fallthrough; eauto. - admit. Admitted. diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 8b96db74..25013696 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -204,6 +204,14 @@ Proof. eapply simu_check_correct; eauto. Qed. +(* TODO *) +Definition injective_checker (dm: PTree.t node): res unit := OK tt. + +Lemma injective_checker_correct dm: + injective_checker dm = OK tt -> is_injective dm. +Proof. +Admitted. + Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (PTree.t node)) := let (tctetpm, dm) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in @@ -211,6 +219,7 @@ Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (P let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in do tf <- proj1_sig (function_builder tfr tpm); do tt <- function_equiv_checker dm f tf; + do _ <- injective_checker dm; OK (tf, dm). Theorem verified_scheduler_correct f tf dm: @@ -221,12 +230,13 @@ Theorem verified_scheduler_correct f tf dm: /\ dm ! (fn_entrypoint tf) = Some (fn_entrypoint f) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2) - /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2). + /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2) + /\ is_injective dm. Proof. intros VERIF. unfold verified_scheduler in VERIF. explore. Local Hint Resolve function_equiv_checker_entrypoint function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 - function_equiv_checker_correct: core. + function_equiv_checker_correct injective_checker_correct: core. destruct (function_builder _ _) as [res H]; simpl in * |- *; auto. apply H in EQ2. rewrite EQ2. simpl. repeat (constructor; eauto). @@ -240,7 +250,8 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint); dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; - dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sexec_simu dupmap f1 f2 pc1 pc2 + dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sexec_simu dupmap f1 f2 pc1 pc2; + dupmap_injective: is_injective dupmap }. Program Definition transf_function (f: RTLpath.function): @@ -251,7 +262,7 @@ Program Definition transf_function (f: RTLpath.function): end. Next Obligation. exploit verified_scheduler_correct; eauto. - intros (A & B & C & D & E & F & G). + intros (A & B & C & D & E & F & G & H). exists dm. econstructor; eauto. Defined. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 38d44ff0..b16010ed 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -159,11 +159,12 @@ Qed. Lemma sistate_simu f dupmap sp st st' rs m is: ssem_internal ge sp st rs m is -> liveness_ok_function f -> + is_injective dupmap -> sistate_simu f dupmap st st' -> exists is', ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. - intros SEM LIVE X; eapply X; eauto. + intros SEM LIVE INJ X; eapply X; eauto. Qed. Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: @@ -229,7 +230,9 @@ Theorem ssem_sstate_simu dm f f' stk stk' sp st st' rs m t s: sstate_simu dm f st st' -> exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. - intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. + intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). + assert (INJ: is_injective dm) by (inv MFUNC; assumption). + destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. @@ -241,7 +244,7 @@ Proof. - (* sem_normal *) exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. - intros (is' & SEM' & (CONT' & RS' & M')). + intros (is' & SEM' & (CONT' & RS' & M') & DMEQ). rewrite <- eqlive_reg_triv in RS'. exploit ssem_final_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). -- cgit From 3b995cfb54ef47b3b1d04acdf8afcce02fd31d79 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 23 Jun 2020 14:31:38 +0200 Subject: Reverting injectivity ? --- kvx/lib/RTLpathSE_theory.v | 17 +++++++++-------- 1 file changed, 9 insertions(+), 8 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index fc330dfd..467b9d3e 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1638,21 +1638,20 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := Typiquement, pour connaître la frame des registres vivants, il faudra faire une propagation en arrière sur la dernière instruction du superblock. *) istate_simulive (fun _ => True) srce is1 is2 - /\ srce!(is2.(ipc)) = Some is1.(ipc) 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). -Definition is_injective (dm: PTree.t node): Prop := - forall x y, dm!x = dm!y -> x = y. +(* Definition is_injective (dm: PTree.t node): Prop := + forall x y, dm!x = dm!y -> x = y. *) (* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sistate): Prop := forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> - is_injective srce -> +(* is_injective srce -> *) forall rs m is1, ssem_internal ge1 sp st1 rs m is1 -> exists is2, ssem_internal ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. @@ -1772,19 +1771,21 @@ Lemma sistate_simub_correct dm f st1 st2: sistate_simu f dm st1 st2. Proof. unfold sistate_simub. intro. explore. unfold sistate_simu. - intros sp ge1 ge2 GFS LOK INJ rs m is1 SEMI. + intros sp ge1 ge2 GFS LOK (* INJ *) rs m is1 SEMI. exploit siexits_simub_correct; eauto. intro ESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). exists is2. constructor; auto. unfold ssem_internal. unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. - destruct ISIMU as ((CONTEQ & RSEQ & MEMEQ) & REVBIND). +(* destruct ISIMU as ((CONTEQ & RSEQ & MEMEQ) & REVBIND). *) + destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). rewrite <- CONTEQ. rewrite ICONT. repeat (constructor; auto). - + eapply revmap_check_single_correct in EQ. + + admit. + (* eapply revmap_check_single_correct in EQ. assert (DMEQ: dm ! (si_pc st2) = dm ! (ipc is2)) by congruence. - apply INJ. assumption. + apply INJ. assumption. *) + eapply siexits_simu_all_fallthrough; eauto. - admit. Admitted. -- cgit From 1ee25a402c286648f8f8af91a096b75bcc186d5f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 23 Jun 2020 17:52:00 +0200 Subject: Fixing sistate_simub_correct ; no need for injection --- kvx/lib/RTLpathSE_theory.v | 22 ++++++++++------------ kvx/lib/RTLpathScheduler.v | 19 ++++++------------- kvx/lib/RTLpathSchedulerproof.v | 8 ++++---- 3 files changed, 20 insertions(+), 29 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 467b9d3e..618b4712 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1761,7 +1761,6 @@ Proof. Qed. Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := - do _ <- revmap_check_single dm (si_pc st1) (si_pc st2); do _ <- silocal_simub dm f (si_local st1) (si_local st2); do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); OK tt. @@ -1776,17 +1775,16 @@ Proof. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). - exists is2. constructor; auto. unfold ssem_internal. - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. -(* destruct ISIMU as ((CONTEQ & RSEQ & MEMEQ) & REVBIND). *) - destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). - rewrite <- CONTEQ. rewrite ICONT. - repeat (constructor; auto). - + admit. - (* eapply revmap_check_single_correct in EQ. - assert (DMEQ: dm ! (si_pc st2) = dm ! (ipc is2)) by congruence. - apply INJ. assumption. *) - + eapply siexits_simu_all_fallthrough; eauto. + destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. + exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. + + unfold ssem_internal. simpl. + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. + destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. + rewrite <- CONTEQ. rewrite ICONT. + repeat (constructor; auto). + eapply siexits_simu_all_fallthrough; eauto. + + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. + destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - admit. Admitted. diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 25013696..e4640bc0 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -204,14 +204,6 @@ Proof. eapply simu_check_correct; eauto. Qed. -(* TODO *) -Definition injective_checker (dm: PTree.t node): res unit := OK tt. - -Lemma injective_checker_correct dm: - injective_checker dm = OK tt -> is_injective dm. -Proof. -Admitted. - Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (PTree.t node)) := let (tctetpm, dm) := untrusted_scheduler f in let (tcte, tpm) := tctetpm in @@ -219,7 +211,7 @@ Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (P let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in do tf <- proj1_sig (function_builder tfr tpm); do tt <- function_equiv_checker dm f tf; - do _ <- injective_checker dm; +(* do _ <- injective_checker dm; *) OK (tf, dm). Theorem verified_scheduler_correct f tf dm: @@ -231,12 +223,13 @@ Theorem verified_scheduler_correct f tf dm: /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2) /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2) - /\ is_injective dm. +(* /\ is_injective dm *) +. Proof. intros VERIF. unfold verified_scheduler in VERIF. explore. Local Hint Resolve function_equiv_checker_entrypoint function_equiv_checker_pathentry1 function_equiv_checker_pathentry2 - function_equiv_checker_correct injective_checker_correct: core. + function_equiv_checker_correct (* injective_checker_correct *): core. destruct (function_builder _ _) as [res H]; simpl in * |- *; auto. apply H in EQ2. rewrite EQ2. simpl. repeat (constructor; eauto). @@ -251,7 +244,7 @@ Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1; dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2; dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sexec_simu dupmap f1 f2 pc1 pc2; - dupmap_injective: is_injective dupmap +(* dupmap_injective: is_injective dupmap *) }. Program Definition transf_function (f: RTLpath.function): @@ -262,7 +255,7 @@ Program Definition transf_function (f: RTLpath.function): end. Next Obligation. exploit verified_scheduler_correct; eauto. - intros (A & B & C & D & E & F & G & H). + intros (A & B & C & D & E & F & G (* & H *)). exists dm. econstructor; eauto. Defined. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index b16010ed..0e55dc6b 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -159,12 +159,12 @@ Qed. Lemma sistate_simu f dupmap sp st st' rs m is: ssem_internal ge sp st rs m is -> liveness_ok_function f -> - is_injective dupmap -> +(* is_injective dupmap -> *) sistate_simu f dupmap st st' -> exists is', ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. - intros SEM LIVE INJ X; eapply X; eauto. + intros SEM LIVE (* INJ *) X; eapply X; eauto. Qed. Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: @@ -231,7 +231,7 @@ Theorem ssem_sstate_simu dm f f' stk stk' sp st st' rs m t s: exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). - assert (INJ: is_injective dm) by (inv MFUNC; assumption). +(* assert (INJ: is_injective dm) by (inv MFUNC; assumption). *) destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) exploit sistate_simu; eauto. @@ -244,7 +244,7 @@ Proof. - (* sem_normal *) exploit sistate_simu; eauto. unfold istate_simu; rewrite CONT. - intros (is' & SEM' & (CONT' & RS' & M') & DMEQ). + intros (is' & SEM' & (CONT' & RS' & M')(* & DMEQ *)). rewrite <- eqlive_reg_triv in RS'. exploit ssem_final_simu; eauto. clear SEM2; intros (s0 & SEM2 & MATCH0). -- cgit From 65f3ee1ca172618a671068f41faa4280e470a061 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 24 Jun 2020 11:56:56 +0200 Subject: Some progress --- kvx/lib/RTLpathSE_theory.v | 54 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 52 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 618b4712..b7112289 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1716,6 +1716,26 @@ Proof. erewrite <- CONDSIMU; eauto. Qed. +Lemma siexits_simu_all_fallthrough_upto dm f ge1 ge2 sp rs m ext1 lx1: forall lse1 lse2, + (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + all_fallthrough_upto_exit ge1 sp ext1 lx1 lse1 rs m -> + siexits_simu dm f lse1 lse2 -> + exists ext2 lx2, + all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m. +Proof. + induction 3. + - destruct H0 as (ITAIL & ALLFU). apply is_tail_false in ITAIL. contradiction. + - destruct H0 as (ITAIL & ALLFU). +(* + + intros GFS (ITAIL & ALLFU) SIMU. + assert (exists lx2, siexits_simu dm f lx1 lx2). + eapply siexits_simu_all_fallthrough in ALLFU. 3: eassumption. eauto. + eexists; eexists; constructor. + constructor. *) +Admitted. + + (* forall sp ge1 ge2, (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> liveness_ok_function f -> @@ -1785,8 +1805,38 @@ Proof. eapply siexits_simu_all_fallthrough; eauto. + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - - admit. -Admitted. + - destruct SEMI as (ext & lx & SSEME & ALLFU). + exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2). + assert (INE: In ext (si_exits st1)). { inversion ALLFU. eapply is_tail_in; eauto. } + assert (INE2: In ext2 (si_exits st2)). { inversion ALLFU2. eapply is_tail_in; eauto. } + Search (list_nth_z _ _ -> list_forall2 _ _ _ -> _). admit. + (* IDEA : + 1) extend siexits_simu_all_fallthrough_upto to include a conclusion list_nth_z (the position of the exit *is* determined) + 2) prove a lemma with list_nth_z and list_forall2 to ensure the (exists ext) is actually the same one from the above lemma + 3) continue + *) +(* + + Search (list_forall2 _ _ _ -> In _ _ -> _). + eapply list_forall2_in_left in INE; eauto. destruct INE as (ext2 & INEXT2 & EXSIMU2). + destruct EXSIMU2 as (SEMSIMU & CONDSIMU). + eapply SEMSIMU in SSEME; eauto. destruct SSEME as (is2 & SSEME2 & ISIMU). + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). + destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). + (* TRY : is2 is the right one? *) + exists is2. constructor; auto. + + unfold ssem_internal. rewrite <- CONTEQ. rewrite ICONT. + exploit siexits_simu_all_fallthrough_upto; eauto. + exists ext2. eexists. constructor; eauto. + + Search siexits_simu. + + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). + + + exists (mk_istate (icont + eexists. constructor; auto. + + unfold ssem_internal. + *)Admitted. Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := match lp with -- cgit From f0f58dd1e4c88f3ca6eff539cb0c5bd69adbb17b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 24 Jun 2020 17:47:24 +0200 Subject: Proof of sistate_simub_correct with smaller admitted lemmas --- kvx/lib/RTLpathSE_theory.v | 100 ++++++++++++++++++++------------------------- 1 file changed, 45 insertions(+), 55 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index b7112289..9d2374c1 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1716,36 +1716,26 @@ Proof. erewrite <- CONDSIMU; eauto. Qed. -Lemma siexits_simu_all_fallthrough_upto dm f ge1 ge2 sp rs m ext1 lx1: forall lse1 lse2, +Lemma siexits_simu_all_fallthrough_upto dm f ge1 ge2 sp rs m lse1 lse2: (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + siexits_simu dm f lse1 lse2 -> forall ext1 lx1, all_fallthrough_upto_exit ge1 sp ext1 lx1 lse1 rs m -> - siexits_simu dm f lse1 lse2 -> exists ext2 lx2, - all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m. -Proof. - induction 3. - - destruct H0 as (ITAIL & ALLFU). apply is_tail_false in ITAIL. contradiction. - - destruct H0 as (ITAIL & ALLFU). -(* - - intros GFS (ITAIL & ALLFU) SIMU. - assert (exists lx2, siexits_simu dm f lx1 lx2). - eapply siexits_simu_all_fallthrough in ALLFU. 3: eassumption. eauto. - eexists; eexists; constructor. - constructor. *) -Admitted. - - -(* forall sp ge1 ge2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - liveness_ok_function f -> - forall rs m is1, - (exists ext1 lx1, ssem_exit ge1 sp ext1 rs m (irs is1) (imem is1) (ipc is1) /\ - all_fallthrough_upto_exit ge1 sp ext1 lx1 lse1 rs m) -> - exists is2, - (exists ext2 lx2, ssem_exit ge2 sp ext2 rs m (irs is2) (imem is2) (ipc is2) /\ - all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m) - /\ istate_simu f dm is1 is2. *) + all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m + /\ length lx1 = length lx2. +Proof. + induction 2. + - intros. destruct H0 as (ITAIL & ALLFU). apply is_tail_false in ITAIL. contradiction. + - intros ext1 lx1 ALLFUE. + destruct ALLFUE as (ITAIL & ALLFU). inv ITAIL. + + eexists; eexists. + constructor; [| eapply list_forall2_length; eauto]. + constructor; auto. eapply siexits_simu_all_fallthrough; eauto. + + exploit IHlist_forall2; [constructor; eauto|]. + intros (ext2 & lx2 & ALLFUE2 & LENEQ). + eexists; eexists. constructor; eauto. + eapply all_fallthrough_upto_exit_cons; eauto. +Qed. Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := match lse1 with @@ -1785,6 +1775,21 @@ Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sist do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); OK tt. +Lemma list_forall2_nth_error {A} (l1 l2: list A) P x1 x2 n: + list_forall2 P l1 l2 -> + nth_error l1 n = Some x1 -> + nth_error l2 n = Some x2 -> + P x1 x2. +Proof. +Admitted. + +Lemma is_tail_nth_error {A} (l1 l2: list A) x: + is_tail (x::l1) l2 -> + (* (length l1) < length l2 *) + nth_error l2 ((length l2) - length l1) = Some x. +Proof. +Admitted. + Lemma sistate_simub_correct dm f st1 st2: sistate_simub dm f st1 st2 = OK tt -> sistate_simu f dm st1 st2. @@ -1806,37 +1811,22 @@ Proof. + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - destruct SEMI as (ext & lx & SSEME & ALLFU). - exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2). - assert (INE: In ext (si_exits st1)). { inversion ALLFU. eapply is_tail_in; eauto. } - assert (INE2: In ext2 (si_exits st2)). { inversion ALLFU2. eapply is_tail_in; eauto. } - Search (list_nth_z _ _ -> list_forall2 _ _ _ -> _). admit. - (* IDEA : - 1) extend siexits_simu_all_fallthrough_upto to include a conclusion list_nth_z (the position of the exit *is* determined) - 2) prove a lemma with list_nth_z and list_forall2 to ensure the (exists ext) is actually the same one from the above lemma - 3) continue - *) -(* - - Search (list_forall2 _ _ _ -> In _ _ -> _). - eapply list_forall2_in_left in INE; eauto. destruct INE as (ext2 & INEXT2 & EXSIMU2). - destruct EXSIMU2 as (SEMSIMU & CONDSIMU). - eapply SEMSIMU in SSEME; eauto. destruct SSEME as (is2 & SSEME2 & ISIMU). + exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). + assert (EXTSIMU: siexit_simu dm f ext ext2). { + eapply list_forall2_nth_error; eauto. + - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. + - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. + assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). + congruence. } + destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. + destruct SSEME as (is2 & SSEME2 & ISIMU). unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). - (* TRY : is2 is the right one? *) - exists is2. constructor; auto. - + unfold ssem_internal. rewrite <- CONTEQ. rewrite ICONT. - exploit siexits_simu_all_fallthrough_upto; eauto. - exists ext2. eexists. constructor; eauto. - - Search siexits_simu. + unfold ssem_internal. exists is2. + rewrite <- CONTEQ. rewrite ICONT. constructor; auto. + + eexists; eexists; constructor; eauto. + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). - - - exists (mk_istate (icont - eexists. constructor; auto. - + unfold ssem_internal. - *)Admitted. +Qed. Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) := match lp with -- cgit From e3efc59d1d36ff1658559af8dc5ec35d206e1405 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 25 Jun 2020 16:59:03 +0200 Subject: Proving small lemmas --- kvx/lib/RTLpathSE_theory.v | 39 +++++++++++++++++++++++++++++++++------ 1 file changed, 33 insertions(+), 6 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 9d2374c1..1eda5505 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1775,20 +1775,47 @@ Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sist do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); OK tt. -Lemma list_forall2_nth_error {A} (l1 l2: list A) P x1 x2 n: +Lemma list_forall2_nth_error {A} (l1 l2: list A) P: list_forall2 P l1 l2 -> + forall x1 x2 n, nth_error l1 n = Some x1 -> nth_error l2 n = Some x2 -> P x1 x2. Proof. -Admitted. + induction 1. + - intros. rewrite nth_error_nil in H. discriminate. + - intros x1 x2 n. destruct n as [|n]; simpl. + + intros. inv H1. inv H2. assumption. + + apply IHlist_forall2. +Qed. + +Lemma is_tail_length {A} (l1 l2: list A): + is_tail l1 l2 -> + (length l1 <= length l2)%nat. +Proof. + induction l2. + - intro. destruct l1; auto. apply is_tail_false in H. contradiction. + - intros ITAIL. inv ITAIL; auto. + apply IHl2 in H1. clear IHl2. simpl. omega. +Qed. Lemma is_tail_nth_error {A} (l1 l2: list A) x: is_tail (x::l1) l2 -> - (* (length l1) < length l2 *) - nth_error l2 ((length l2) - length l1) = Some x. -Proof. -Admitted. + nth_error l2 ((length l2) - length l1 - 1) = Some x. +Proof. + induction l2. + - intro ITAIL. apply is_tail_false in ITAIL. contradiction. + - intros ITAIL. assert (length (a::l2) = S (length l2)) by auto. rewrite H. clear H. + assert (forall n n', ((S n) - n' - 1)%nat = (n - n')%nat) by (intros; omega). rewrite H. clear H. + inv ITAIL. + + assert (forall n, (n - n)%nat = 0%nat) by (intro; omega). rewrite H. + simpl. reflexivity. + + exploit IHl2; eauto. intros. clear IHl2. + assert (forall n n', (n > n')%nat -> (n - n')%nat = S (n - n' - 1)%nat) by (intros; omega). + exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2. + assert ((length l2 > length l1)%nat) by omega. clear H2. + rewrite H0; auto. +Qed. Lemma sistate_simub_correct dm f st1 st2: sistate_simub dm f st1 st2 = OK tt -> -- cgit From c2a80479ae69b2c241bafe6a03d29fbc715ddff0 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 25 Jun 2020 18:12:56 +0200 Subject: Start of sval_simub --- kvx/lib/RTLpathSE_theory.v | 96 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 93 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 1eda5505..3702c86d 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -8,7 +8,7 @@ Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. -Require Import Errors. +Require Import Errors Duplicate. Local Open Scope error_monad_scope. @@ -1915,13 +1915,103 @@ Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): | Sreturn_simu os: sfval_simu dm f opc1 opc2 (Sreturn os) (Sreturn os). -(* TODO *) -Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := OK tt. +Fixpoint sval_simub (sv1 sv2: sval) := + match sv1 with + | Sinput r => + match sv2 with + | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") + | _ => Error (msg "sval_simub: Sinput expected") + end + | Sop op lsv sm => + match sv2 with + | Sop op' lsv' sm' => + if (eq_operation op op') then + do _ <- list_sval_simub lsv lsv'; + smem_simub sm sm' + else Error (msg "sval_simub: different operations in Sop") + | _ => Error (msg "sval_simub: Sop expected") + end + | _ => OK tt (* TODO *) + end +with list_sval_simub (lsv1 lsv2: list_sval) := + match lsv1 with + | Snil => + match lsv2 with + | Snil => OK tt + | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") + end + | Scons sv1 lsv1 => + match lsv2 with + | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") + | Scons sv2 lsv2 => + do _ <- sval_simub sv1 sv2; + list_sval_simub lsv1 lsv2 + end + end +with smem_simub (sm1 sm2: smem) := + match sm1 with + | Sinit => + match sm2 with + | Sinit => OK tt + | _ => Error (msg "smem_simub: Sinit expected") + end + | Sstore sm chk addr lsv sv => + match sm2 with + | Sstore sm' chk' addr' lsv' sv' => sval_simub sv sv' (* TODO *) + | _ => Error (msg "smem_simub: Sstore expected") + end + end. + +Lemma sval_simub_correct sv1 sv2: + sval_simub sv1 sv2 = OK tt -> sv1 = sv2. +Proof. + unfold sval_simub. destruct sv1. + (* Sinput *) + - destruct sv2; try discriminate. + destruct (Pos.eq_dec r r0); try discriminate. + congruence. + (* Sop *) + - admit. + (* Sload *) + - admit. +Admitted. + +Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := + match fv1 with + | Snone => + match fv2 with + | Snone => revmap_check_single dm pc1 pc2 + | _ => Error (msg "sfval_simub: Snone expected") + end + (* TODO - to fill once we have the sval_simub !! *) +(* | Scall sig svos lsv res pc1 => + match fv2 with + | Scall sig2 svos2 lsv2 res2 pc2 => + do _ <- revmap_check_single dm pc1 pc2; + if (signature_eq sig sig2) then + if (product_eq + else Error (msg "sfval_simub: Scall different signatures") *) + | _ => OK tt + end. Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2: sfval_simub dm f pc1 pc2 fv1 fv2 = OK tt -> sfval_simu dm f pc1 pc2 fv1 fv2. Proof. + unfold sfval_simub. destruct fv1. + (* Snone *) + - destruct fv2; try discriminate. intro. + apply revmap_check_single_correct in H. constructor; auto. + (* Scall *) + - admit. + (* Stailcall *) + - admit. + (* Sbuiltin *) + - admit. + (* Sjumptable *) + - admit. + (* Sreturn *) + - admit. Admitted. Definition sstate_simu dm f (s1 s2: sstate): Prop := -- cgit From 2265471720f6934b5b09a0e942e49453f1c30bd7 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 29 Jun 2020 11:42:43 +0200 Subject: sval_simub + sval_simub_correct --- kvx/lib/RTLpathSE_theory.v | 60 +++++++++++++++++++++++++++++++++++++++------- 1 file changed, 51 insertions(+), 9 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 3702c86d..1c1224fd 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1931,7 +1931,20 @@ Fixpoint sval_simub (sv1 sv2: sval) := else Error (msg "sval_simub: different operations in Sop") | _ => Error (msg "sval_simub: Sop expected") end - | _ => OK tt (* TODO *) + | Sload sm trap chk addr lsv => + match sv2 with + | Sload sm' trap' chk' addr' lsv' => + if (trapping_mode_eq trap trap') then + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + list_sval_simub lsv lsv' + else Error (msg "sval_simub: addressings do not match") + else Error (msg "sval_simub: chunks do not match") + else Error (msg "sval_simub: trapping modes do not match") + (* FIXME - should be refined once we introduce non trapping loads *) + | _ => Error (msg "sval_simub: Sload expected") + end end with list_sval_simub (lsv1 lsv2: list_sval) := match lsv1 with @@ -1957,24 +1970,53 @@ with smem_simub (sm1 sm2: smem) := end | Sstore sm chk addr lsv sv => match sm2 with - | Sstore sm' chk' addr' lsv' sv' => sval_simub sv sv' (* TODO *) + | Sstore sm' chk' addr' lsv' sv' => + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + do _ <- list_sval_simub lsv lsv'; + sval_simub sv sv' + else Error (msg "smem_simub: addressings do not match") + else Error (msg "smem_simub: chunks not match") | _ => Error (msg "smem_simub: Sstore expected") end end. -Lemma sval_simub_correct sv1 sv2: +Lemma sval_simub_correct sv1: forall sv2, sval_simub sv1 sv2 = OK tt -> sv1 = sv2. Proof. - unfold sval_simub. destruct sv1. + induction sv1 using sval_mut with + (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) + (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. (* Sinput *) - destruct sv2; try discriminate. - destruct (Pos.eq_dec r r0); try discriminate. - congruence. + destruct (Pos.eq_dec r r0); [congruence|discriminate]. (* Sop *) - - admit. + - destruct sv2; try discriminate. + destruct (eq_operation _ _); [|discriminate]. subst. + intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. (* Sload *) - - admit. -Admitted. + - destruct sv2; try discriminate. + destruct (trapping_mode_eq _ _); [|discriminate]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. + congruence. + (* Snil *) + - destruct lsv2; [|discriminate]. congruence. + (* Scons *) + - destruct lsv2; [discriminate|]. intro. explore. + apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sinit *) + - destruct sm2; [|discriminate]. congruence. + (* Sstore *) + - destruct sm2; [discriminate|]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. + assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. + congruence. +Qed. Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := match fv1 with -- cgit From 476125efe5fb2bee34261d2388126b60dbe7453f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 29 Jun 2020 12:09:10 +0200 Subject: Going further in sfval_simub --- kvx/lib/RTLpathSE_theory.v | 66 +++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 60 insertions(+), 6 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 1c1224fd..5e56e8aa 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -2018,6 +2018,43 @@ Proof. congruence. Qed. +Lemma list_sval_simub_correct lsv1: forall lsv2, + list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. +Proof. + induction lsv1; simpl; auto. + - destruct lsv2; try discriminate. reflexivity. + - destruct lsv2; try discriminate. intro. explore. + apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. + congruence. +Qed. + +Definition svos_simub (svos1 svos2: sval + ident) := + match svos1 with + | inl sv1 => + match svos2 with + | inl sv2 => sval_simub sv1 sv2 + | _ => Error (msg "svos_simub: expected sval") + end + | inr id1 => + match svos2 with + | inr id2 => + if (ident_eq id1 id2) then OK tt + else Error (msg "svos_simub: ids do not match") + | _ => Error (msg "svos_simub: expected id") + end + end. + +Lemma svos_simub_correct svos1 svos2: + svos_simub svos1 svos2 = OK tt -> + svos1 = svos2. +Proof. + destruct svos1. + - simpl. destruct svos2; [|discriminate]. + intro. exploit sval_simub_correct; eauto. congruence. + - simpl. destruct svos2; [discriminate|]. + intro. explore. reflexivity. +Qed. + Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := match fv1 with | Snone => @@ -2025,14 +2062,27 @@ Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) | Snone => revmap_check_single dm pc1 pc2 | _ => Error (msg "sfval_simub: Snone expected") end - (* TODO - to fill once we have the sval_simub !! *) -(* | Scall sig svos lsv res pc1 => + | Scall sig svos lsv res pc1 => match fv2 with | Scall sig2 svos2 lsv2 res2 pc2 => do _ <- revmap_check_single dm pc1 pc2; if (signature_eq sig sig2) then - if (product_eq - else Error (msg "sfval_simub: Scall different signatures") *) + if (Pos.eq_dec res res2) then + do _ <- svos_simub svos svos2; + list_sval_simub lsv lsv2 + else Error (msg "sfval_simub: Scall res do not match") + else Error (msg "sfval_simub: Scall different signatures") + | _ => Error (msg "sfval_simub: Scall expected") + end + | Stailcall sig svos lsv => + match fv2 with + | Stailcall sig' svos' lsv' => + if (signature_eq sig sig') then + do _ <- svos_simub svos svos'; + list_sval_simub lsv lsv' + else Error (msg "sfval_simub: signatures do not match") + | _ => Error (msg "sfval_simub: Stailcall expected") + end | _ => OK tt end. @@ -2045,9 +2095,13 @@ Proof. - destruct fv2; try discriminate. intro. apply revmap_check_single_correct in H. constructor; auto. (* Scall *) - - admit. + - destruct fv2; try discriminate. intro. explore. + apply svos_simub_correct in EQ3. apply list_sval_simub_correct in EQ4. + subst. apply revmap_check_single_correct in EQ. constructor; auto. (* Stailcall *) - - admit. + - destruct fv2; try discriminate. intro. explore. + apply svos_simub_correct in EQ0. apply list_sval_simub_correct in EQ1. + subst. constructor; auto. (* Sbuiltin *) - admit. (* Sjumptable *) -- cgit From 10f8aa5bb4765e0132bfc73ba1f73a897371e628 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 29 Jun 2020 14:28:18 +0200 Subject: sfval_simub_correct --- kvx/lib/RTLpathSE_theory.v | 181 +++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 175 insertions(+), 6 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 5e56e8aa..bac66e21 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -4,7 +4,7 @@ NB: an efficient implementation with hash-consing will be defined in ... *) -Require Import Coqlib Maps. +Require Import Coqlib Maps Floats. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import Op Registers. Require Import RTL RTLpath. @@ -1895,6 +1895,30 @@ Proof. constructor; eauto. congruence. Qed. +Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := + match ln with + | nil => + match ln' with + | nil => OK tt + | _ => Error (msg "revmap_check_list: lists have different lengths") + end + | n::ln => + match ln' with + | nil => Error (msg "revmap_check_list: lists have different lengths") + | n'::ln' => do _ <- revmap_check_single dm n n'; revmap_check_list dm ln ln' + end + end. + +Lemma revmap_check_list_correct dm ln: forall ln', + revmap_check_list dm ln ln' = OK tt -> + ptree_get_list dm ln' = Some ln. +Proof. + induction ln. + - simpl. intros. destruct ln'; try discriminate. constructor; auto. + - simpl. intros; destruct ln'; try discriminate. explore. + apply IHln in EQ0. apply revmap_check_single_correct in EQ. + simpl. rewrite EQ. rewrite EQ0. reflexivity. +Qed. (* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): sfval -> sfval -> Prop := @@ -2055,6 +2079,112 @@ Proof. intro. explore. reflexivity. Qed. +Fixpoint builtin_arg_simub (bs bs': builtin_arg sval) := + match bs with + | BA sv => + match bs' with + | BA sv' => sval_simub sv sv' + | _ => Error (msg "builtin_arg_simub: BA expected") + end + | BA_int n => + match bs' with + | BA_int n' => if (Integers.int_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_int expected") + end + | BA_long n => + match bs' with + | BA_long n' => if (int64_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_long expected") + end + | BA_float f => + match bs' with + | BA_float f' => if (float_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_float expected") + end + | BA_single f => + match bs' with + | BA_single f' => if (float32_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_single expected") + end + | BA_loadstack chk ptr => + match bs' with + | BA_loadstack chk' ptr' => + if (chunk_eq chk chk') then + if (ptrofs_eq ptr ptr') then OK tt + else Error (msg "builtin_arg_simub: ptr do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadstack expected") + end + | BA_addrstack ptr => + match bs' with + | BA_addrstack ptr' => if (ptrofs_eq ptr ptr') then OK tt else Error (msg "builtin_arg_simub: ptr do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrstack expected") + end + | BA_loadglobal chk id ofs => + match bs' with + | BA_loadglobal chk' id' ofs' => + if (chunk_eq chk chk') then + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadglobal expected") + end + | BA_addrglobal id ofs => + match bs' with + | BA_addrglobal id' ofs' => + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrglobal expected") + end + | BA_splitlong lo hi => + match bs' with + | BA_splitlong lo' hi' => do _ <- builtin_arg_simub lo lo'; builtin_arg_simub hi hi' + | _ => Error (msg "builtin_arg_simub: BA_splitlong expected") + end + | BA_addptr b1 b2 => + match bs' with + | BA_addptr b1' b2' => do _ <- builtin_arg_simub b1 b1'; builtin_arg_simub b2 b2' + | _ => Error (msg "builtin_arg_simub: BA_addptr expected") + end + end. + +Lemma builtin_arg_simub_correct b1: forall b2, + builtin_arg_simub b1 b2 = OK tt -> b1 = b2. +Proof. + induction b1; simpl; destruct b2; try discriminate; auto; intros; try (explore; congruence). + - apply sval_simub_correct in H. congruence. + - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. + - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. +Qed. + +Fixpoint list_builtin_arg_simub lbs1 lbs2 := + match lbs1 with + | nil => + match lbs2 with + | nil => OK tt + | _ => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs2 is bigger)") + end + | bs1::lbs1 => + match lbs2 with + | nil => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs1 is bigger)") + | bs2::lbs2 => + do _ <- builtin_arg_simub bs1 bs2; + list_builtin_arg_simub lbs1 lbs2 + end + end. + +Lemma list_builtin_arg_simub_correct lsb1: forall lsb2, + list_builtin_arg_simub lsb1 lsb2 = OK tt -> lsb1 = lsb2. +Proof. + induction lsb1; intros; simpl; destruct lsb2; try discriminate; auto. + simpl in H. explore. apply builtin_arg_simub_correct in EQ. + assert (lsb1 = lsb2) by auto. congruence. +Qed. + Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := match fv1 with | Snone => @@ -2083,7 +2213,40 @@ Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) else Error (msg "sfval_simub: signatures do not match") | _ => Error (msg "sfval_simub: Stailcall expected") end - | _ => OK tt + | Sbuiltin ef lbs br pc => + match fv2 with + | Sbuiltin ef' lbs' br' pc' => + if (external_function_eq ef ef') then + if (builtin_res_eq_pos br br') then + do _ <- revmap_check_single dm pc pc'; + list_builtin_arg_simub lbs lbs' + else Error (msg "sfval_simub: builtin res do not match") + else Error (msg "sfval_simub: external functions do not match") + | _ => Error (msg "sfval_simub: Sbuiltin expected") + end + | Sjumptable sv ln => + match fv2 with + | Sjumptable sv' ln' => + do _ <- revmap_check_list dm ln ln'; sval_simub sv sv' + | _ => Error (msg "sfval_simub: Sjumptable expected") + end + | Sreturn osv => + match fv2 with + | Sreturn osv' => + match osv with + | Some sv => + match osv' with + | Some sv' => sval_simub sv sv' + | None => Error (msg "sfval_simub sv' expected") + end + | None => + match osv' with + | Some sv' => Error (msg "None expected") + | None => OK tt + end + end + | _ => Error (msg "sfval_simub: Sreturn expected") + end end. Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2: @@ -2103,12 +2266,18 @@ Proof. apply svos_simub_correct in EQ0. apply list_sval_simub_correct in EQ1. subst. constructor; auto. (* Sbuiltin *) - - admit. + - destruct fv2; try discriminate. intro. explore. + apply revmap_check_single_correct in EQ1. apply list_builtin_arg_simub_correct in EQ2. + subst. constructor; auto. (* Sjumptable *) - - admit. + - destruct fv2; try discriminate. intro. explore. + apply revmap_check_list_correct in EQ. apply sval_simub_correct in EQ0. subst. + constructor; auto. (* Sreturn *) - - admit. -Admitted. + - destruct fv2; try discriminate. destruct o; destruct o0; try discriminate. + + intro. apply sval_simub_correct in H. subst. constructor; auto. + + constructor; auto. +Qed. Definition sstate_simu dm f (s1 s2: sstate): Prop := sistate_simu f dm s1.(internal) s2.(internal) -- cgit From f3eba03ac9ea9072d946a0b5f1f0ec4d5712057b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 29 Jun 2020 16:13:27 +0200 Subject: Only silocal_simub left --- kvx/lib/RTLpathSE_theory.v | 303 +++++++++++++++++++++++++-------------------- 1 file changed, 171 insertions(+), 132 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index bac66e21..cf20b888 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1646,6 +1646,133 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := (* Definition is_injective (dm: PTree.t node): Prop := forall x y, dm!x = dm!y -> x = y. *) +Fixpoint sval_simub (sv1 sv2: sval) := + match sv1 with + | Sinput r => + match sv2 with + | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") + | _ => Error (msg "sval_simub: Sinput expected") + end + | Sop op lsv sm => + match sv2 with + | Sop op' lsv' sm' => + if (eq_operation op op') then + do _ <- list_sval_simub lsv lsv'; + smem_simub sm sm' + else Error (msg "sval_simub: different operations in Sop") + | _ => Error (msg "sval_simub: Sop expected") + end + | Sload sm trap chk addr lsv => + match sv2 with + | Sload sm' trap' chk' addr' lsv' => + if (trapping_mode_eq trap trap') then + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + list_sval_simub lsv lsv' + else Error (msg "sval_simub: addressings do not match") + else Error (msg "sval_simub: chunks do not match") + else Error (msg "sval_simub: trapping modes do not match") + (* FIXME - should be refined once we introduce non trapping loads *) + | _ => Error (msg "sval_simub: Sload expected") + end + end +with list_sval_simub (lsv1 lsv2: list_sval) := + match lsv1 with + | Snil => + match lsv2 with + | Snil => OK tt + | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") + end + | Scons sv1 lsv1 => + match lsv2 with + | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") + | Scons sv2 lsv2 => + do _ <- sval_simub sv1 sv2; + list_sval_simub lsv1 lsv2 + end + end +with smem_simub (sm1 sm2: smem) := + match sm1 with + | Sinit => + match sm2 with + | Sinit => OK tt + | _ => Error (msg "smem_simub: Sinit expected") + end + | Sstore sm chk addr lsv sv => + match sm2 with + | Sstore sm' chk' addr' lsv' sv' => + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + do _ <- list_sval_simub lsv lsv'; + sval_simub sv sv' + else Error (msg "smem_simub: addressings do not match") + else Error (msg "smem_simub: chunks not match") + | _ => Error (msg "smem_simub: Sstore expected") + end + end. + +Lemma sval_simub_correct sv1: forall sv2, + sval_simub sv1 sv2 = OK tt -> sv1 = sv2. +Proof. + induction sv1 using sval_mut with + (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) + (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. + (* Sinput *) + - destruct sv2; try discriminate. + destruct (Pos.eq_dec r r0); [congruence|discriminate]. + (* Sop *) + - destruct sv2; try discriminate. + destruct (eq_operation _ _); [|discriminate]. subst. + intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sload *) + - destruct sv2; try discriminate. + destruct (trapping_mode_eq _ _); [|discriminate]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. + congruence. + (* Snil *) + - destruct lsv2; [|discriminate]. congruence. + (* Scons *) + - destruct lsv2; [discriminate|]. intro. explore. + apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sinit *) + - destruct sm2; [|discriminate]. congruence. + (* Sstore *) + - destruct sm2; [discriminate|]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. + assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. + congruence. +Qed. + +Lemma list_sval_simub_correct lsv1: forall lsv2, + list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. +Proof. + induction lsv1; simpl; auto. + - destruct lsv2; try discriminate. reflexivity. + - destruct lsv2; try discriminate. intro. explore. + apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. + congruence. +Qed. + +Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := + match dm ! tn with + | None => Error (msg "revmap_check_single: no mapping for tn") + | Some res => if (Pos.eq_dec n res) then OK tt + else Error (msg "revmap_check_single: n and res do not match") + end. + +Lemma revmap_check_single_correct dm n tn: + revmap_check_single dm n tn = OK tt -> + dm ! tn = Some n. +Proof. + unfold revmap_check_single. explore; try discriminate. congruence. +Qed. + (* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *) Definition sistate_simu (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sistate): Prop := forall sp ge1 ge2, @@ -1665,6 +1792,12 @@ Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sista (* TODO *) Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) := OK tt. +Lemma silocal_simub_eq dm f sl1 sl2: + silocal_simub dm f sl1 sl2 = OK tt -> + si_sreg sl1 = si_sreg sl2 /\ si_smem sl1 = si_smem sl2. +Proof. +Admitted. + Lemma silocal_simub_correct dm f sl1 sl2: silocal_simub dm f sl1 sl2 = OK tt -> silocal_simu dm f sl1 sl2. @@ -1690,14 +1823,39 @@ Definition siexit_simu_cond (dm: PTree.t node) (f: RTLpath.function) (se1 se2: s Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := siexit_simu_ssem dm f se1 se2 /\ siexit_simu_cond dm f se1 se2. -(* TODO *) -Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := OK tt. +Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := + if (eq_condition (si_cond se1) (si_cond se2)) then + do _ <- list_sval_simub (si_scondargs se1) (si_scondargs se2); + do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); + revmap_check_single dm (si_ifso se1) (si_ifso se2) + else Error (msg "siexit_simub: conditions do not match") +. Lemma siexit_simub_correct dm f se1 se2: siexit_simub dm f se1 se2 = OK tt -> siexit_simu dm f se1 se2. Proof. -Admitted. + unfold siexit_simub. intro. explore. + apply list_sval_simub_correct in EQ0. exploit silocal_simub_eq; eauto. + intros (SREGEQ & SMEMEQ). + apply silocal_simub_correct in EQ2. + apply revmap_check_single_correct in EQ3. + constructor. + - unfold siexit_simu_ssem. intros sp ge ge' GFS LOK rs m is SEMEXIT. + destruct SEMEXIT as (SCOND & SLOCAL & SIFSO). + let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). + eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). + destruct is2 as [icont ipc irs imem]. simpl in *. + exists (mk_istate icont (si_ifso se2) irs imem). simpl. + constructor; auto. + + repeat (constructor; auto). + rewrite <- EQ0. rewrite <- e. rewrite <- SMEMEQ. erewrite seval_condition_preserved; eauto. + + unfold istate_simu in *. destruct (icontinue is) eqn:ICONT. + * destruct ISIMU as (A & B & C). constructor; auto. + * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. + repeat (constructor; auto). simpl in *. congruence. + - unfold siexit_simu_cond. intros sp ge ge' GFS rs m. erewrite seval_condition_preserved; eauto. congruence. +Qed. Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := list_forall2 (siexit_simu dm f) lse1 lse2. @@ -1739,7 +1897,11 @@ Qed. Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := match lse1 with - | nil => OK tt + | nil => + match lse2 with + | nil => OK tt + | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") + end | se1 :: lse1 => match lse2 with | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") @@ -1750,24 +1912,14 @@ Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list end end. -Lemma siexits_simub_correct dm f lse1 lse2: +Lemma siexits_simub_correct dm f lse1: forall lse2, siexits_simub dm f lse1 lse2 = OK tt -> siexits_simu dm f lse1 lse2. Proof. -Admitted. - -Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := - match dm ! tn with - | None => Error (msg "revmap_check_single: no mapping for tn") - | Some res => if (Pos.eq_dec n res) then OK tt - else Error (msg "revmap_check_single: n and res do not match") - end. - -Lemma revmap_check_single_correct dm n tn: - revmap_check_single dm n tn = OK tt -> - dm ! tn = Some n. -Proof. - unfold revmap_check_single. explore; try discriminate. congruence. + induction lse1. + - simpl. intros. destruct lse2; try discriminate. constructor. + - simpl. intros. destruct lse2; try discriminate. explore. + apply siexit_simub_correct in EQ. apply IHlse1 in EQ1. constructor; auto. Qed. Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := @@ -1939,119 +2091,6 @@ Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): | Sreturn_simu os: sfval_simu dm f opc1 opc2 (Sreturn os) (Sreturn os). -Fixpoint sval_simub (sv1 sv2: sval) := - match sv1 with - | Sinput r => - match sv2 with - | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") - | _ => Error (msg "sval_simub: Sinput expected") - end - | Sop op lsv sm => - match sv2 with - | Sop op' lsv' sm' => - if (eq_operation op op') then - do _ <- list_sval_simub lsv lsv'; - smem_simub sm sm' - else Error (msg "sval_simub: different operations in Sop") - | _ => Error (msg "sval_simub: Sop expected") - end - | Sload sm trap chk addr lsv => - match sv2 with - | Sload sm' trap' chk' addr' lsv' => - if (trapping_mode_eq trap trap') then - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - list_sval_simub lsv lsv' - else Error (msg "sval_simub: addressings do not match") - else Error (msg "sval_simub: chunks do not match") - else Error (msg "sval_simub: trapping modes do not match") - (* FIXME - should be refined once we introduce non trapping loads *) - | _ => Error (msg "sval_simub: Sload expected") - end - end -with list_sval_simub (lsv1 lsv2: list_sval) := - match lsv1 with - | Snil => - match lsv2 with - | Snil => OK tt - | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") - end - | Scons sv1 lsv1 => - match lsv2 with - | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") - | Scons sv2 lsv2 => - do _ <- sval_simub sv1 sv2; - list_sval_simub lsv1 lsv2 - end - end -with smem_simub (sm1 sm2: smem) := - match sm1 with - | Sinit => - match sm2 with - | Sinit => OK tt - | _ => Error (msg "smem_simub: Sinit expected") - end - | Sstore sm chk addr lsv sv => - match sm2 with - | Sstore sm' chk' addr' lsv' sv' => - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - do _ <- list_sval_simub lsv lsv'; - sval_simub sv sv' - else Error (msg "smem_simub: addressings do not match") - else Error (msg "smem_simub: chunks not match") - | _ => Error (msg "smem_simub: Sstore expected") - end - end. - -Lemma sval_simub_correct sv1: forall sv2, - sval_simub sv1 sv2 = OK tt -> sv1 = sv2. -Proof. - induction sv1 using sval_mut with - (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) - (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. - (* Sinput *) - - destruct sv2; try discriminate. - destruct (Pos.eq_dec r r0); [congruence|discriminate]. - (* Sop *) - - destruct sv2; try discriminate. - destruct (eq_operation _ _); [|discriminate]. subst. - intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sload *) - - destruct sv2; try discriminate. - destruct (trapping_mode_eq _ _); [|discriminate]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. - congruence. - (* Snil *) - - destruct lsv2; [|discriminate]. congruence. - (* Scons *) - - destruct lsv2; [discriminate|]. intro. explore. - apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sinit *) - - destruct sm2; [|discriminate]. congruence. - (* Sstore *) - - destruct sm2; [discriminate|]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. - assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. - congruence. -Qed. - -Lemma list_sval_simub_correct lsv1: forall lsv2, - list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. -Proof. - induction lsv1; simpl; auto. - - destruct lsv2; try discriminate. reflexivity. - - destruct lsv2; try discriminate. intro. explore. - apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. - congruence. -Qed. - Definition svos_simub (svos1 svos2: sval + ident) := match svos1 with | inl sv1 => -- cgit From 852db29f0c75fbc9c285201a0c01135673e9a5a3 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 29 Jun 2020 16:27:16 +0200 Subject: Must generalize silocal_simub with a list of regs --- kvx/lib/RTLpathSE_theory.v | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index cf20b888..ff359b97 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1789,8 +1789,10 @@ Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sista forall rs m is1, ssem_local ge1 sp sl1 rs m (irs is1) (imem is1) -> exists is2, ssem_local ge2 sp sl2 rs m (irs is2) (imem is2) /\ istate_simu f dm is1 is2. -(* TODO *) -Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) := OK tt. +(* TODO - generalize it with a list of registers to test *) +Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (* (regs: list reg) *) (sl1 sl2: sistate_local) := + OK tt +. Lemma silocal_simub_eq dm f sl1 sl2: silocal_simub dm f sl1 sl2 = OK tt -> @@ -1826,7 +1828,7 @@ Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistat Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := if (eq_condition (si_cond se1) (si_cond se2)) then do _ <- list_sval_simub (si_scondargs se1) (si_scondargs se2); - do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); + do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); (* TODO - generalize it with a list of registers to test *) revmap_check_single dm (si_ifso se1) (si_ifso se2) else Error (msg "siexit_simub: conditions do not match") . -- cgit From 8a2346a8f0eb9280550887c64f15e59b822b5b16 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 30 Jun 2020 12:17:27 +0200 Subject: starting refinement of (abstract) symbolic executions --- kvx/lib/RTLpathSE_impl.v | 160 +++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 160 insertions(+) create mode 100644 kvx/lib/RTLpathSE_impl.v diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v new file mode 100644 index 00000000..7e62f1c4 --- /dev/null +++ b/kvx/lib/RTLpathSE_impl.v @@ -0,0 +1,160 @@ +(* Implementation of the symbolic execution +*) + +Require Import Coqlib Maps Floats. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL RTLpath. +Require Import Errors Duplicate. +Require Import RTLpathSE_theory. + +Local Open Scope error_monad_scope. + +(* +(* TODO: for the implementation of the symbolic execution. + It is important to remove dependency of Op on memory. + Here is an way to formalize this *) + +Parameter mem_irrel: operation -> bool. (* TODO: TO PROVIDE FOR EACH BACKEND ? *) +Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_irrel o = true -> eval_operation ge sp o args m1 = eval_operation ge sp o args m2. + +*) + +(** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations *) + +(* Implementation of local symbolic internal states +TODO: use refined symbolic values instead of those from RTLpathSE_theory. +*) +(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) +Record hsistate_local := + { si_hsmems: list smem; + si_ok_hsval: list sval; + hsi_ptree: PTree.t sval + }. + +Definition hsi_sval_get (hst: hsistate_local) r: sval := + match PTree.get r (hsi_ptree hst) with + | None => Sinput r + | Some sv => sv + end. + +(* NB: short cut for (seval_sval ge sp (hsi_sval_get hst r) rs0 m0) *) +Definition hsi_sval_eval ge sp (hst: hsistate_local) r rs0 m0 : option val := + match PTree.get r (hsi_ptree hst) with + | None => Some (Regmap.get r rs0) + | Some sv => seval_sval ge sp sv rs0 m0 + end. + +Definition hsi_smem_get (hst: hsistate_local): smem := + match si_hsmems hst with + | nil => Sinit + | sm::_ => sm + end. + +(* NB: short cut for (seval_smem ge sp (hsi_smem_get hst) rs0 m0) *) +Definition hsi_smem_eval ge sp (hst: hsistate_local) rs0 m0 : option mem := + match si_hsmems hst with + | nil => Some m0 + | sm::_ => seval_smem ge sp sm rs0 m0 + end. + +(* negation of sabort_local *) +Definition sok_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem): Prop := + (st.(si_pre) ge sp rs0 m0) + /\ seval_smem ge sp st.(si_smem) rs0 m0 <> None + /\ forall (r: reg), seval_sval ge sp (si_sreg st r) rs0 m0 <> None. + +Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := + forall sv, List.In sv (si_ok_hsval hst) -> seval_sval ge sp sv rs0 m0 <> None + /\ forall sm, List.In sm (si_hsmems hst) -> seval_smem ge sp sm rs0 m0 <> None. + +(* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) +Definition si_local_model ge sp (st: sistate_local) (hst: hsistate_local) := + (forall rs0 m0, sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) + /\ (forall rs0 m0, hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) + /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sval_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). + +(* Syntax and semantics of symbolic exit states *) +(* TODO: add hash-consing *) +Record hsistate_exit := mk_hsistate_exit + { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. + +Definition si_exit_model (ge: RTL.genv) (sp: val) (ext: sistate_exit) (hext: hsistate_exit) : Prop := True. (* TODO *) + +(** * Syntax and Semantics of symbolic internal state *) +Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. + +Definition sistate_model (ge: RTL.genv) (sp: val) (st:sistate) (hst: hsistate) : Prop := True. (* TODO *) + +(** * Symbolic execution of one internal step + TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations +*) + +Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := + {| si_hsmems := sm::si_hsmems hst; + si_ok_hsval := si_ok_hsval hst; + hsi_ptree:= hsi_ptree hst + |}. + + +Definition is_trapping (sv: sval): bool := + match sv with + (* | Sinput _ => false *) (* TODO: is it useful here ? *) + | Sload _ NOTRAP _ _ _ => false + | _ => true (* TODO: to refine for the refined operations *) + end. + +(* untrap version of a trapping operation *) +Definition untrap (sv: sval): sval := + match sv with + | Sload sm TRAP chunk addr lsv => Sload sm NOTRAP chunk addr lsv + | _ => sv + end. + +(* naive version +Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := + {| si_hsmems := si_hsmems hst; + si_ok_hsval := sv::(si_ok_hsval hst); + hsi_ptree:= PTree.set r sv (hsi_ptree hst) |}. +*) + +Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := + {| si_hsmems := si_hsmems hst; + si_ok_hsval := if is_trapping sv then sv::(si_ok_hsval hst) else si_ok_hsval hst; + hsi_ptree := PTree.set r (untrap sv) (hsi_ptree hst) |}. + +Definition hsist_set_local (hst: hsistate) (pc: node) (nxt: hsistate_local): option hsistate := + Some {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= nxt |}. + +Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := + match i with + | Inop pc' => + hsist_set_local hst pc' hst.(hsi_local) + | Iop op args dst pc' => (* TODO: deal with non-trapping operations *) + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let next := hslocal_set_sreg prev dst (Sop op vargs (hsi_smem_get prev)) in + hsist_set_local hst pc' next + | Istore chunk addr args src pc' => + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sval_get prev src)) in + hsist_set_local hst pc' next + | Iload trap chunk addr args dst pc' => + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let next := hslocal_set_sreg prev dst (Sload (hsi_smem_get prev) trap chunk addr vargs) in + hsist_set_local hst pc' next + | Icond cond args ifso ifnot _ => + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in + Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |} + | _ => None (* TODO jumptable ? *) + end. + +Lemma hsiexec_inst_correct i hst hst': + hsiexec_inst i hst = Some hst' -> + forall ge sp st, sistate_model ge sp st hst -> exists st', siexec_inst i st = Some st' /\ sistate_model ge sp st' hst'. +Admitted. + -- cgit From 2757c07cc270d3f23798a0bb59658358b5190638 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 30 Jun 2020 15:31:44 +0200 Subject: a minor comment --- kvx/lib/RTLpathSE_impl.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 7e62f1c4..c63b3dd1 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -39,7 +39,7 @@ Definition hsi_sval_get (hst: hsistate_local) r: sval := end. (* NB: short cut for (seval_sval ge sp (hsi_sval_get hst r) rs0 m0) *) -Definition hsi_sval_eval ge sp (hst: hsistate_local) r rs0 m0 : option val := +Definition hsi_sval_eval ge sp (hst: hsistate_local) r rs0 m0: option val := match PTree.get r (hsi_ptree hst) with | None => Some (Regmap.get r rs0) | Some sv => seval_sval ge sp sv rs0 m0 @@ -96,7 +96,7 @@ Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := hsi_ptree:= hsi_ptree hst |}. - +(* NB: return [false] if the sv cannot fail *) Definition is_trapping (sv: sval): bool := match sv with (* | Sinput _ => false *) (* TODO: is it useful here ? *) -- cgit From 4b1f6f2c00347ffec7a9357dc4ef99bf056e6c1c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Jul 2020 12:04:23 +0200 Subject: Some renaming --- kvx/lib/RTLpathSE_impl.v | 87 +++++++++++++++++++++++++++--------------------- 1 file changed, 49 insertions(+), 38 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index c63b3dd1..c9cc4acf 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -1,5 +1,5 @@ -(* Implementation of the symbolic execution -*) +(** Implementation and refinement of the symbolic execution + *) Require Import Coqlib Maps Floats. Require Import AST Integers Values Events Memory Globalenvs Smallstep. @@ -25,35 +25,43 @@ Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_ (* Implementation of local symbolic internal states TODO: use refined symbolic values instead of those from RTLpathSE_theory. *) -(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *) + +(** name : Hash-consed Symbolic Internal state local. Later on we will use the + refinement to introduce hash consing *) Record hsistate_local := - { si_hsmems: list smem; - si_ok_hsval: list sval; - hsi_ptree: PTree.t sval + { + (** [hsi_lsmem] represents the list of smem symbolic evaluations. + The first one of the list is the current smem *) + hsi_lsmem: list smem; + (** For the values in registers: + 1) we store a list of sval evaluations + 2) we encode the symbolic regset by a PTree *) + hsi_ok_lsval: list sval; + hsi_sreg: PTree.t sval }. -Definition hsi_sval_get (hst: hsistate_local) r: sval := - match PTree.get r (hsi_ptree hst) with +Definition hsi_sreg_get (hst: hsistate_local) r: sval := + match PTree.get r (hsi_sreg hst) with | None => Sinput r | Some sv => sv end. -(* NB: short cut for (seval_sval ge sp (hsi_sval_get hst r) rs0 m0) *) -Definition hsi_sval_eval ge sp (hst: hsistate_local) r rs0 m0: option val := - match PTree.get r (hsi_ptree hst) with +(* NB: short cut for (seval_sval ge sp (hsi_sreg_get hst r) rs0 m0) *) +Definition hsi_sreg_eval ge sp (hst: hsistate_local) r rs0 m0: option val := + match PTree.get r (hsi_sreg hst) with | None => Some (Regmap.get r rs0) | Some sv => seval_sval ge sp sv rs0 m0 end. Definition hsi_smem_get (hst: hsistate_local): smem := - match si_hsmems hst with + match hsi_lsmem hst with | nil => Sinit | sm::_ => sm end. (* NB: short cut for (seval_smem ge sp (hsi_smem_get hst) rs0 m0) *) Definition hsi_smem_eval ge sp (hst: hsistate_local) rs0 m0 : option mem := - match si_hsmems hst with + match hsi_lsmem hst with | nil => Some m0 | sm::_ => seval_smem ge sp sm rs0 m0 end. @@ -65,47 +73,47 @@ Definition sok_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) ( /\ forall (r: reg), seval_sval ge sp (si_sreg st r) rs0 m0 <> None. Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := - forall sv, List.In sv (si_ok_hsval hst) -> seval_sval ge sp sv rs0 m0 <> None - /\ forall sm, List.In sm (si_hsmems hst) -> seval_smem ge sp sm rs0 m0 <> None. + forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None + /\ forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None. (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition si_local_model ge sp (st: sistate_local) (hst: hsistate_local) := +Definition si_local_refines ge sp (st: sistate_local) (hst: hsistate_local) := (forall rs0 m0, sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) /\ (forall rs0 m0, hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) - /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sval_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). + /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. -Definition si_exit_model (ge: RTL.genv) (sp: val) (ext: sistate_exit) (hext: hsistate_exit) : Prop := True. (* TODO *) +Definition si_exit_refines (ge: RTL.genv) (sp: val) (ext: sistate_exit) (hext: hsistate_exit) : Prop := True. (* TODO *) (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. -Definition sistate_model (ge: RTL.genv) (sp: val) (st:sistate) (hst: hsistate) : Prop := True. (* TODO *) +Definition sistate_refines (ge: RTL.genv) (sp: val) (st:sistate) (hst: hsistate) : Prop := True. (* TODO *) (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations *) Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := - {| si_hsmems := sm::si_hsmems hst; - si_ok_hsval := si_ok_hsval hst; - hsi_ptree:= hsi_ptree hst + {| hsi_lsmem := sm::hsi_lsmem hst; + hsi_ok_lsval := hsi_ok_lsval hst; + hsi_sreg:= hsi_sreg hst |}. (* NB: return [false] if the sv cannot fail *) -Definition is_trapping (sv: sval): bool := +Definition may_trap (sv: sval): bool := match sv with (* | Sinput _ => false *) (* TODO: is it useful here ? *) | Sload _ NOTRAP _ _ _ => false | _ => true (* TODO: to refine for the refined operations *) end. -(* untrap version of a trapping operation *) -Definition untrap (sv: sval): sval := +(* version without trapping of a may_trap operation. Identity if there is no such version *) +Definition try_untrap (sv: sval): sval := match sv with | Sload sm TRAP chunk addr lsv => Sload sm NOTRAP chunk addr lsv | _ => sv @@ -113,15 +121,15 @@ Definition untrap (sv: sval): sval := (* naive version Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := - {| si_hsmems := si_hsmems hst; - si_ok_hsval := sv::(si_ok_hsval hst); - hsi_ptree:= PTree.set r sv (hsi_ptree hst) |}. + {| hsi_lsmem := hsi_lsmem hst; + hsi_ok_lsval := sv::(hsi_ok_lsval hst); + hsi_sreg:= PTree.set r sv (hsi_sreg hst) |}. *) Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := - {| si_hsmems := si_hsmems hst; - si_ok_hsval := if is_trapping sv then sv::(si_ok_hsval hst) else si_ok_hsval hst; - hsi_ptree := PTree.set r (untrap sv) (hsi_ptree hst) |}. + {| hsi_lsmem := hsi_lsmem hst; + hsi_ok_lsval := if may_trap sv then sv::(hsi_ok_lsval hst) else hsi_ok_lsval hst; + hsi_sreg := PTree.set r (try_untrap sv) (hsi_sreg hst) |}. Definition hsist_set_local (hst: hsistate) (pc: node) (nxt: hsistate_local): option hsistate := Some {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= nxt |}. @@ -132,29 +140,32 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := hsist_set_local hst pc' hst.(hsi_local) | Iop op args dst pc' => (* TODO: deal with non-trapping operations *) let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in let next := hslocal_set_sreg prev dst (Sop op vargs (hsi_smem_get prev)) in hsist_set_local hst pc' next | Istore chunk addr args src pc' => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in - let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sval_get prev src)) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in hsist_set_local hst pc' next | Iload trap chunk addr args dst pc' => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in let next := hslocal_set_sreg prev dst (Sload (hsi_smem_get prev) trap chunk addr vargs) in hsist_set_local hst pc' next | Icond cond args ifso ifnot _ => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sval_get prev) args) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |} | _ => None (* TODO jumptable ? *) end. +(** Our refinement is correct ; ie it simulates the (non refined) symbolic states *) Lemma hsiexec_inst_correct i hst hst': - hsiexec_inst i hst = Some hst' -> - forall ge sp st, sistate_model ge sp st hst -> exists st', siexec_inst i st = Some st' /\ sistate_model ge sp st' hst'. + hsiexec_inst i hst = Some hst' -> + forall ge sp st, sistate_refines ge sp st hst -> + exists st', siexec_inst i st = Some st' /\ sistate_refines ge sp st' hst'. +Proof. Admitted. -- cgit From 824faa3396204b695407032af8db74fee1af0bd8 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Jul 2020 15:02:25 +0200 Subject: hsexec_correct --- kvx/lib/RTLpathSE_impl.v | 117 ++++++++++++++++++++++++++++++++++++++++++--- kvx/lib/RTLpathSE_theory.v | 48 +++++++++++-------- 2 files changed, 138 insertions(+), 27 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index c9cc4acf..9e955cb4 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -9,6 +9,7 @@ Require Import Errors Duplicate. Require Import RTLpathSE_theory. Local Open Scope error_monad_scope. +Local Open Scope option_monad_scope. (* (* TODO: for the implementation of the symbolic execution. @@ -77,7 +78,7 @@ Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := /\ forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None. (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition si_local_refines ge sp (st: sistate_local) (hst: hsistate_local) := +Definition hsistate_local_refines ge sp (hst: hsistate_local) (st: sistate_local) := (forall rs0 m0, sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) /\ (forall rs0 m0, hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). @@ -87,12 +88,19 @@ Definition si_local_refines ge sp (st: sistate_local) (hst: hsistate_local) := Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. -Definition si_exit_refines (ge: RTL.genv) (sp: val) (ext: sistate_exit) (hext: hsistate_exit) : Prop := True. (* TODO *) +Definition hsistate_exit_refines (ge: RTL.genv) (sp: val) (hext: hsistate_exit) (ext: sistate_exit): Prop := + hsi_cond hext = si_cond ext + /\ hsi_scondargs hext = si_scondargs ext + /\ hsistate_local_refines ge sp (hsi_elocal hext) (si_elocal ext) + /\ hsi_ifso hext = si_ifso ext. (** * Syntax and Semantics of symbolic internal state *) -Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. +Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. -Definition sistate_refines (ge: RTL.genv) (sp: val) (st:sistate) (hst: hsistate) : Prop := True. (* TODO *) +Definition hsistate_refines (ge: RTL.genv) (sp: val) (hst: hsistate) (st:sistate): Prop := + hsi_pc hst = si_pc st + /\ list_forall2 (hsistate_exit_refines ge sp) (hsi_exits hst) (si_exits st) + /\ hsistate_local_refines ge sp (hsi_local hst) (si_local st). (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations @@ -164,8 +172,105 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := (** Our refinement is correct ; ie it simulates the (non refined) symbolic states *) Lemma hsiexec_inst_correct i hst hst': hsiexec_inst i hst = Some hst' -> - forall ge sp st, sistate_refines ge sp st hst -> - exists st', siexec_inst i st = Some st' /\ sistate_refines ge sp st' hst'. + forall ge sp st, hsistate_refines ge sp hst st -> + exists st', siexec_inst i st = Some st' /\ hsistate_refines ge sp hst' st'. Proof. Admitted. +Lemma hsiexec_inst_none i hst ge sp st: + hsiexec_inst i hst = None -> + hsistate_refines ge sp hst st -> + siexec_inst i st = None. +Proof. +Admitted. + +Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := + match path with + | O => Some hst + | S p => + SOME i <- (fn_code f)!(hst.(hsi_pc)) IN + SOME hst1 <- hsiexec_inst i hst IN + hsiexec_path p f hst1 + end. + +Lemma hsiexec_path_correct path f hst hst': + hsiexec_path path f hst = Some hst' -> + forall ge sp st, hsistate_refines ge sp hst st -> + exists st', siexec_path path f st = Some st' /\ hsistate_refines ge sp hst' st'. +Proof. +Admitted. + +Record hsstate := { hinternal:> hsistate; hfinal: sfval }. + +Definition hsstate_refines (ge: RTL.genv) (sp: val) (hst: hsstate) (st:sstate): Prop := + hsistate_refines ge sp (hinternal hst) (internal st) + /\ hfinal hst = final st. (* TODO *) + +Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := + match i with + | Icall sig ros args res pc => + let svos := sum_left_map (hsi_sreg_get prev) ros in + let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + Scall sig svos sargs res pc + | Itailcall sig ros args => + let svos := sum_left_map (hsi_sreg_get prev) ros in + let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + Stailcall sig svos sargs + | Ibuiltin ef args res pc => + let sargs := List.map (builtin_arg_inj (hsi_sreg_get prev)) args in + Sbuiltin ef sargs res pc + | Ireturn or => + let sor := SOME r <- or IN Some (hsi_sreg_get prev r) in + Sreturn sor + | Ijumptable reg tbl => + let sv := hsi_sreg_get prev reg in + Sjumptable sv tbl + | _ => Snone + end. + +Lemma hsexec_final_correct ge sp hsl sl i: + hsistate_local_refines ge sp hsl sl -> + hsexec_final i hsl = sexec_final i sl. +Proof. +Admitted. + +Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. + +Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. + +Remark init_hsistate_correct ge sp pc: + hsistate_refines ge sp (init_hsistate pc) (init_sistate pc). +Proof. +Admitted. + +Definition hsexec (f: function) (pc:node): option hsstate := + SOME path <- (fn_path f)!pc IN + SOME hst <- hsiexec_path path.(psize) f (init_hsistate pc) IN + SOME i <- (fn_code f)!(hst.(hsi_pc)) IN + Some (match hsiexec_inst i hst with + | Some hst' => {| hinternal := hst'; hfinal := Snone |} + | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} + end). + +Lemma hsexec_correct f pc hst: + hsexec f pc = Some hst -> forall ge sp, + exists st', sexec f pc = Some st' /\ hsstate_refines ge sp hst st'. +Proof. + unfold hsexec. intros. explore_hyp. + eapply hsiexec_path_correct in EQ0; [|eapply init_hsistate_correct]. + destruct EQ0 as (st & EPATH & SISREF). + destruct (hsiexec_inst i h) eqn:HEINST. + - eapply hsiexec_inst_correct in HEINST; eauto. + destruct HEINST as (st' & EINST & SISREF'). + destruct SISREF as (PCEQ & EXREF & LOCREF). + destruct SISREF' as (PCEQ' & EXREF' & LOCREF'). + eexists. constructor. + + unfold sexec. rewrite EQ. rewrite EPATH. rewrite <- PCEQ. rewrite EQ1. reflexivity. + + rewrite EINST. constructor; auto. constructor; [|constructor]; simpl; eauto. + - eapply hsiexec_inst_none in HEINST; eauto. assert (SISREF' := SISREF). + destruct SISREF' as (PCEQ & EXREF & LOCREF). + eexists. constructor. + + unfold sexec. rewrite EQ. rewrite EPATH. rewrite <- PCEQ. rewrite EQ1. reflexivity. + + rewrite HEINST. erewrite hsexec_final_correct; eauto. constructor; simpl; auto. +Qed. + diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index ff359b97..ae702510 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -13,18 +13,36 @@ Require Import Errors Duplicate. Local Open Scope error_monad_scope. (* Enhanced from kvx/Asmblockgenproof.v *) -Ltac explore := +Ltac explore_hyp := repeat match goal with - | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in destruct var eqn:EQ1) + | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in (destruct var eqn:EQ1; try discriminate)) | [ H : OK _ = OK _ |- _ ] => monadInv H - | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "EQ" in destruct b eqn:EQ1) - | [ |- context[match ?m with | _ => _ end] ] => destruct m - | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m | [ H : bind _ _ = OK _ |- _ ] => monadInv H | [ H : Error _ = OK _ |- _ ] => inversion H + | [ H : Some _ = Some _ |- _ ] => inv H | [ x : unit |- _ ] => destruct x end. +Ltac explore := explore_hyp; + repeat match goal with + | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "IEQ" in destruct b eqn:EQ1) + | [ |- context[match ?m with | _ => _ end] ] => (let DEQ1 := fresh "DEQ" in destruct m eqn:DEQ1) + | [ |- context[match ?m as _ return _ with | _ => _ end]] => (let DREQ1 := fresh "DREQ" in destruct m eqn:DREQ1) + end. + +(* Ltac explore := + repeat match goal with + | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in (destruct var eqn:EQ1; try discriminate)) + | [ H : OK _ = OK _ |- _ ] => monadInv H + | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "IEQ" in destruct b eqn:EQ1) + | [ |- context[match ?m with | _ => _ end] ] => (let DEQ1 := fresh "DEQ" in destruct m eqn:DEQ1) + | [ |- context[match ?m as _ return _ with | _ => _ end]] => (let DREQ1 := fresh "DREQ" in destruct m eqn:DREQ1) + | [ H : bind _ _ = OK _ |- _ ] => monadInv H + | [ H : Error _ = OK _ |- _ ] => inversion H + | [ H : Some _ = Some _ |- _ ] => inv H + | [ x : unit |- _ ] => destruct x + end. *) + (* NOTE for the implementation of the symbolic execution. It is important to remove dependency of Op on memory. Here is an way to formalize this @@ -499,8 +517,6 @@ Proof. try_simplify_someHyps. Qed. -(* TODO - continue updating the rest *) - 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. @@ -518,16 +534,6 @@ Proof. unfold sabort_local. simpl; intuition. Qed. -Lemma sist_set_local_preserves_sabort ge sp rs0 m0 st st' loc' pc': - sist_set_local st pc' loc' = Some st' -> - sabort ge sp st rs0 m0 -> - sabort ge sp st' rs0 m0. -Proof. - intros SSL ABORT. inv SSL. unfold sabort. simpl. - destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)]. - - -Abort. - 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. @@ -2025,8 +2031,8 @@ Lemma ptree_get_list_nth dm p2: forall lp2 lp1, Proof. induction lp2. - simpl. intros. inv H. simpl in *. discriminate. - - intros lp1 PGL n LNZ. simpl in PGL. explore; try discriminate. - inv PGL. inv LNZ. destruct (zeq n 0) eqn:ZEQ. + - 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. @@ -2041,8 +2047,8 @@ Lemma ptree_get_list_nth_rev dm p1: forall lp2 lp1, Proof. induction lp2. - simpl. intros. inv H. simpl in *. discriminate. - - intros lp1 PGL n LNZ. simpl in PGL. explore; try discriminate. - inv PGL. inv LNZ. destruct (zeq n 0) eqn:ZEQ. + - 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. -- cgit From 61fca98669eac4f4a83de88a35ccee4833f8ec65 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Jul 2020 16:33:13 +0200 Subject: Avancement --- kvx/lib/RTLpathSE_impl.v | 140 ++++++++++++++++++++++++++++++++++++++------- kvx/lib/RTLpathScheduler.v | 54 +---------------- 2 files changed, 120 insertions(+), 74 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 9e955cb4..17cb3d1a 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -78,7 +78,7 @@ Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := /\ forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None. (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition hsistate_local_refines ge sp (hst: hsistate_local) (st: sistate_local) := +Definition hsistate_local_refines (hst: hsistate_local) (st: sistate_local) := forall ge sp, (forall rs0 m0, sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) /\ (forall rs0 m0, hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). @@ -88,19 +88,19 @@ Definition hsistate_local_refines ge sp (hst: hsistate_local) (st: sistate_local Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. -Definition hsistate_exit_refines (ge: RTL.genv) (sp: val) (hext: hsistate_exit) (ext: sistate_exit): Prop := +Definition hsistate_exit_refines (hext: hsistate_exit) (ext: sistate_exit): Prop := hsi_cond hext = si_cond ext /\ hsi_scondargs hext = si_scondargs ext - /\ hsistate_local_refines ge sp (hsi_elocal hext) (si_elocal ext) + /\ hsistate_local_refines (hsi_elocal hext) (si_elocal ext) /\ hsi_ifso hext = si_ifso ext. (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. -Definition hsistate_refines (ge: RTL.genv) (sp: val) (hst: hsistate) (st:sistate): Prop := +Definition hsistate_refines (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st - /\ list_forall2 (hsistate_exit_refines ge sp) (hsi_exits hst) (si_exits st) - /\ hsistate_local_refines ge sp (hsi_local hst) (si_local st). + /\ list_forall2 hsistate_exit_refines (hsi_exits hst) (si_exits st) + /\ hsistate_local_refines (hsi_local hst) (si_local st). (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations @@ -169,17 +169,17 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := | _ => None (* TODO jumptable ? *) end. -(** Our refinement is correct ; ie it simulates the (non refined) symbolic states *) +(* (** Our refinement is correct ; ie it simulates the (non refined) symbolic states *) Lemma hsiexec_inst_correct i hst hst': hsiexec_inst i hst = Some hst' -> - forall ge sp st, hsistate_refines ge sp hst st -> - exists st', siexec_inst i st = Some st' /\ hsistate_refines ge sp hst' st'. + forall st, hsistate_refines hst st -> + exists st', siexec_inst i st = Some st' /\ hsistate_refines hst' st'. Proof. -Admitted. +Admitted. *) -Lemma hsiexec_inst_none i hst ge sp st: +Lemma hsiexec_inst_none i hst st: hsiexec_inst i hst = None -> - hsistate_refines ge sp hst st -> + hsistate_refines hst st -> siexec_inst i st = None. Proof. Admitted. @@ -193,17 +193,17 @@ Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate hsiexec_path p f hst1 end. -Lemma hsiexec_path_correct path f hst hst': +(* Lemma hsiexec_path_correct path f hst hst': hsiexec_path path f hst = Some hst' -> forall ge sp st, hsistate_refines ge sp hst st -> exists st', siexec_path path f st = Some st' /\ hsistate_refines ge sp hst' st'. Proof. -Admitted. +Admitted. *) Record hsstate := { hinternal:> hsistate; hfinal: sfval }. -Definition hsstate_refines (ge: RTL.genv) (sp: val) (hst: hsstate) (st:sstate): Prop := - hsistate_refines ge sp (hinternal hst) (internal st) +Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := + hsistate_refines (hinternal hst) (internal st) /\ hfinal hst = final st. (* TODO *) Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := @@ -228,8 +228,8 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := | _ => Snone end. -Lemma hsexec_final_correct ge sp hsl sl i: - hsistate_local_refines ge sp hsl sl -> +Lemma hsexec_final_correct hsl sl i: + hsistate_local_refines hsl sl -> hsexec_final i hsl = sexec_final i sl. Proof. Admitted. @@ -238,8 +238,8 @@ Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := ni Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. -Remark init_hsistate_correct ge sp pc: - hsistate_refines ge sp (init_hsistate pc) (init_sistate pc). +Remark init_hsistate_correct pc: + hsistate_refines (init_hsistate pc) (init_sistate pc). Proof. Admitted. @@ -252,7 +252,7 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -Lemma hsexec_correct f pc hst: +(* Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> forall ge sp, exists st', sexec f pc = Some st' /\ hsstate_refines ge sp hst st'. Proof. @@ -272,5 +272,103 @@ Proof. eexists. constructor. + unfold sexec. rewrite EQ. rewrite EPATH. rewrite <- PCEQ. rewrite EQ1. reflexivity. + rewrite HEINST. erewrite hsexec_final_correct; eauto. constructor; simpl; auto. +Qed. *) + +Lemma hsexec_correct f pc st hst: + hsexec f pc = Some hst -> + hsstate_refines hst st -> + sexec f pc = Some st. +Proof. +Admitted. + +Lemma hsexec_complete f pc st: + sexec f pc = Some st -> + exists hst, hsexec f pc = Some hst /\ hsstate_refines hst st. +Proof. +Admitted. + +Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := + forall st1, + hsstate_refines hst1 st1 -> + exists st2, + hsstate_refines hst2 st2 + /\ sstate_simu dm f st1 st2. + +Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) + +Lemma hsstate_simub_correct dm f hst1 hst2: + hsstate_simub dm f hst1 hst2 = OK tt -> + hsstate_simu dm f hst1 hst2. +Proof. +Admitted. + +Definition hsexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := + forall hst1, hsexec f1 pc1 = Some hst1 -> + exists hst2, hsexec f2 pc2 = Some hst2 /\ hsstate_simu dm f1 hst1 hst2. + +Lemma hsexec_sexec_simu dm f1 f2 pc1 pc2: + hsexec_simu dm f1 f2 pc1 pc2 -> sexec_simu dm f1 f2 pc1 pc2. +Proof. + intro HESIMU. + intros st1 SEXEC. + eapply hsexec_complete in SEXEC. + destruct SEXEC as (hst & HSEXEC & HREF). + eapply HESIMU in HSEXEC. destruct HSEXEC as (hst' & HSEXEC' & HSSIMU). + eapply HSSIMU in HREF. destruct HREF as (st' & HREF' & SSIMU). + eapply hsexec_correct in HSEXEC'; eauto. +Qed. + +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := + let (pc2, pc1) := m in + match (hsexec f pc1) with + | None => Error (msg "simu_check_single: hsexec f pc1 failed") + | Some hst1 => + match (hsexec tf pc2) with + | None => Error (msg "simu_check_single: hsexec tf pc2 failed") + | Some hst2 => hsstate_simub dm f hst1 hst2 + end + end. + +Lemma simu_check_single_correct dm tf f pc1 pc2: + simu_check_single dm f tf (pc2, pc1) = OK tt -> + sexec_simu dm f tf pc1 pc2. +Proof. + unfold simu_check_single. intro. explore. + assert (hsexec_simu dm f tf pc1 pc2). { + unfold hsexec_simu. intros st1 STEQ. assert (h = st1) by congruence. subst. + exists h0. constructor; auto. + apply hsstate_simub_correct. assumption. + } + apply hsexec_sexec_simu. assumption. Qed. +Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := + match lm with + | nil => OK tt + | m :: lm => do u1 <- simu_check_single dm f tf m; + do u2 <- simu_check_rec dm f tf lm; + OK tt + end. + +Lemma simu_check_rec_correct dm f tf pc1 pc2: forall lm, + simu_check_rec dm f tf lm = OK tt -> + In (pc2, pc1) lm -> + sexec_simu dm f tf pc1 pc2. +Proof. + induction lm. + - simpl. intuition. + - simpl. intros. explore. destruct H0. + + subst. eapply simu_check_single_correct; eauto. + + eapply IHlm; assumption. +Qed. + +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). + +Lemma simu_check_correct dm f tf: + simu_check dm f tf = OK tt -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + sexec_simu dm f tf pc1 pc2. +Proof. + unfold simu_check. intros. eapply PTree.elements_correct in H0. + eapply simu_check_rec_correct; eassumption. +Qed. \ No newline at end of file diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index e4640bc0..5a1c7f1c 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -6,7 +6,7 @@ This module is inspired from [Duplicate] and [Duplicateproof] Require Import AST Linking Values Maps Globalenvs Smallstep Registers. Require Import Coqlib Maps Events Errors Op. -Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory. +Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory RTLpathSE_impl. Notation "'ASSERT' A 'WITH' MSG 'IN' B" := (if A then B else Error (msg MSG)) @@ -107,58 +107,6 @@ Proof. eapply path_entry_check_rec_correct; eassumption. Qed. -Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := - let (pc2, pc1) := m in - match (sexec f pc1) with - | None => Error (msg "simu_check_single: sexec f pc1 failed") - | Some st1 => - match (sexec tf pc2) with - | None => Error (msg "simu_check_single: sexec tf pc2 failed") - | Some st2 => sstate_simub dm f st1 st2 - end - end. - -Lemma simu_check_single_correct dm tf f pc1 pc2: - simu_check_single dm f tf (pc2, pc1) = OK tt -> - sexec_simu dm f tf pc1 pc2. -Proof. - unfold simu_check_single. intro. explore. - unfold sexec_simu. intros st1 STEQ. assert (s = st1) by congruence. subst. - exists s0. constructor; auto. - apply sstate_simub_correct. assumption. -Qed. - -Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := - match lm with - | nil => OK tt - | m :: lm => do u1 <- simu_check_single dm f tf m; - do u2 <- simu_check_rec dm f tf lm; - OK tt - end. - -Lemma simu_check_rec_correct dm f tf pc1 pc2: forall lm, - simu_check_rec dm f tf lm = OK tt -> - In (pc2, pc1) lm -> - sexec_simu dm f tf pc1 pc2. -Proof. - induction lm. - - simpl. intuition. - - simpl. intros. explore. destruct H0. - + subst. eapply simu_check_single_correct; eauto. - + eapply IHlm; assumption. -Qed. - -Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). - -Lemma simu_check_correct dm f tf: - simu_check dm f tf = OK tt -> - forall pc1 pc2, dm ! pc2 = Some pc1 -> - sexec_simu dm f tf pc1 pc2. -Proof. - unfold simu_check. intros. eapply PTree.elements_correct in H0. - eapply simu_check_rec_correct; eassumption. -Qed. - Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) : res unit := let pm := fn_path f in let fr := fn_RTL f in -- cgit From 5d1ef044ebf94a2ae92202f435ef3ed4904b7c52 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Jul 2020 18:00:32 +0200 Subject: hsexec_complete --- kvx/lib/RTLpathSE_impl.v | 73 ++++++++++++++++++++++++------------------------ 1 file changed, 36 insertions(+), 37 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 17cb3d1a..ef3dc090 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -102,6 +102,12 @@ Definition hsistate_refines (hst: hsistate) (st:sistate): Prop := /\ list_forall2 hsistate_exit_refines (hsi_exits hst) (si_exits st) /\ hsistate_local_refines (hsi_local hst) (si_local st). +Lemma hsistate_refines_pceq st hst: + hsistate_refines hst st -> + (hsi_pc hst) = (si_pc st). +Proof. +Admitted. + (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations *) @@ -169,18 +175,19 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := | _ => None (* TODO jumptable ? *) end. -(* (** Our refinement is correct ; ie it simulates the (non refined) symbolic states *) -Lemma hsiexec_inst_correct i hst hst': - hsiexec_inst i hst = Some hst' -> - forall st, hsistate_refines hst st -> - exists st', siexec_inst i st = Some st' /\ hsistate_refines hst' st'. +Lemma hsiexec_inst_none i st hst: + siexec_inst i st = None -> + hsistate_refines hst st -> + hsiexec_inst i hst = None. Proof. -Admitted. *) +Admitted. -Lemma hsiexec_inst_none i hst st: - hsiexec_inst i hst = None -> +Lemma hsiexec_inst_complete i st st' hst: + siexec_inst i st = Some st' -> hsistate_refines hst st -> - siexec_inst i st = None. + exists hst', + hsiexec_inst i hst = Some hst' + /\ hsistate_refines hst' st'. Proof. Admitted. @@ -193,12 +200,14 @@ Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate hsiexec_path p f hst1 end. -(* Lemma hsiexec_path_correct path f hst hst': - hsiexec_path path f hst = Some hst' -> - forall ge sp st, hsistate_refines ge sp hst st -> - exists st', siexec_path path f st = Some st' /\ hsistate_refines ge sp hst' st'. +Lemma hsiexec_path_complete ps f st st' hst: + siexec_path ps f st = Some st' -> + hsistate_refines hst st -> + exists hst', + hsiexec_path ps f hst = Some hst' + /\ hsistate_refines hst' st'. Proof. -Admitted. *) +Admitted. Record hsstate := { hinternal:> hsistate; hfinal: sfval }. @@ -252,28 +261,6 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -(* Lemma hsexec_correct f pc hst: - hsexec f pc = Some hst -> forall ge sp, - exists st', sexec f pc = Some st' /\ hsstate_refines ge sp hst st'. -Proof. - unfold hsexec. intros. explore_hyp. - eapply hsiexec_path_correct in EQ0; [|eapply init_hsistate_correct]. - destruct EQ0 as (st & EPATH & SISREF). - destruct (hsiexec_inst i h) eqn:HEINST. - - eapply hsiexec_inst_correct in HEINST; eauto. - destruct HEINST as (st' & EINST & SISREF'). - destruct SISREF as (PCEQ & EXREF & LOCREF). - destruct SISREF' as (PCEQ' & EXREF' & LOCREF'). - eexists. constructor. - + unfold sexec. rewrite EQ. rewrite EPATH. rewrite <- PCEQ. rewrite EQ1. reflexivity. - + rewrite EINST. constructor; auto. constructor; [|constructor]; simpl; eauto. - - eapply hsiexec_inst_none in HEINST; eauto. assert (SISREF' := SISREF). - destruct SISREF' as (PCEQ & EXREF & LOCREF). - eexists. constructor. - + unfold sexec. rewrite EQ. rewrite EPATH. rewrite <- PCEQ. rewrite EQ1. reflexivity. - + rewrite HEINST. erewrite hsexec_final_correct; eauto. constructor; simpl; auto. -Qed. *) - Lemma hsexec_correct f pc st hst: hsexec f pc = Some hst -> hsstate_refines hst st -> @@ -285,7 +272,19 @@ Lemma hsexec_complete f pc st: sexec f pc = Some st -> exists hst, hsexec f pc = Some hst /\ hsstate_refines hst st. Proof. -Admitted. + unfold sexec. intro. explore_hyp. + eapply hsiexec_path_complete in EQ0; [| apply init_hsistate_correct]. + destruct EQ0 as (hst & HPATH & HREF). + eexists. constructor. + + unfold hsexec. rewrite EQ. rewrite HPATH. + erewrite hsistate_refines_pceq; eauto. rewrite EQ1. reflexivity. + + destruct (siexec_inst i s) eqn:EINST. + - eapply hsiexec_inst_complete in EINST; eauto. destruct EINST as (hst' & HEINST & HREF'). + rewrite HEINST. constructor; auto. + - eapply hsiexec_inst_none in EINST; eauto. rewrite EINST. + constructor; auto. simpl. + destruct HREF as (_ & _ & HLREF). apply hsexec_final_correct. assumption. +Qed. Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := forall st1, -- cgit From 5e116ac1b0254e7904a543d357009c0125ecf423 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Jul 2020 18:22:09 +0200 Subject: Avancement, mais est-ce prouvable ? --- kvx/lib/RTLpathSE_impl.v | 37 +++++++++++++++++++++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index ef3dc090..1c62825a 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -191,6 +191,15 @@ Lemma hsiexec_inst_complete i st st' hst: Proof. Admitted. +Lemma hsiexec_inst_correct i hst hst' st: + hsiexec_inst i hst = Some hst' -> + hsistate_refines hst st -> + exists st', + siexec_inst i st = Some st' + /\ hsistate_refines hst' st'. +Proof. +Admitted. + Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := match path with | O => Some hst @@ -209,6 +218,15 @@ Lemma hsiexec_path_complete ps f st st' hst: Proof. Admitted. +Lemma hsiexec_path_correct ps f hst hst' st: + hsiexec_path ps f hst = Some hst' -> + hsistate_refines hst st -> + exists st', + siexec_path ps f st = Some st' + /\ hsistate_refines hst' st'. +Proof. +Admitted. + Record hsstate := { hinternal:> hsistate; hfinal: sfval }. Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := @@ -261,11 +279,30 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). +(* FIXME - this lemma might be unprovable ? *) +Lemma hsistate_refines_determ hst st st': + hsistate_refines hst st -> + hsistate_refines hst st' -> + st = st'. +Proof. +Admitted. + Lemma hsexec_correct f pc st hst: hsexec f pc = Some hst -> hsstate_refines hst st -> sexec f pc = Some st. Proof. + unfold hsexec. intros H HREF. explore. + unfold sexec. rewrite EQ. + destruct (hsiexec_inst i h) eqn:HINST. + - destruct HREF as (HIREF & HFREF). simpl in *. + eapply hsiexec_path_correct in EQ0; [| apply init_hsistate_correct]. + destruct EQ0 as (st2 & SIPATH & HIREF2). rewrite SIPATH. + erewrite <- hsistate_refines_pceq; eauto. rewrite EQ1. + eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST & HIREF'). + rewrite SINST. destruct st as (sti & stf). simpl in *. + eapply hsistate_refines_determ in HIREF; eauto. congruence. + - admit. Admitted. Lemma hsexec_complete f pc st: -- cgit From 9f440451e0694eddf66318bf2cde9bf25c32ab83 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 2 Jul 2020 07:48:50 +0200 Subject: remove useless (and unprovable) lemmas on completeness of the refinement --- configure | 2 +- kvx/lib/RTLpathSE_impl.v | 113 +++++++++++++++++++++++++++-------------------- 2 files changed, 67 insertions(+), 48 deletions(-) diff --git a/configure b/configure index 0db4b3a5..959d0590 100755 --- a/configure +++ b/configure @@ -845,7 +845,7 @@ EXECUTE=kvx-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __KVX_COS__ SIMU=kvx-cluster -- BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\ - RTLpath.v RTLpathproof.v RTLpathLivegen.v RTLpathLivegenproof.v RTLpathSE_theory.v RTLpathScheduler.v RTLpathSchedulerproof.v\\ + RTLpath.v RTLpathproof.v RTLpathLivegen.v RTLpathLivegenproof.v RTLpathSE_theory.v RTLpathSE_impl.v RTLpathScheduler.v RTLpathSchedulerproof.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 1c62825a..e4ae1af8 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -78,32 +78,32 @@ Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := /\ forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None. (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition hsistate_local_refines (hst: hsistate_local) (st: sistate_local) := forall ge sp, - (forall rs0 m0, sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) - /\ (forall rs0 m0, hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) - /\ (forall rs0 m0 r, hsok_local ge sp hst rs0 m0 -> hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). +Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistate_local) := + (sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) + /\ (hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) + /\ (hsok_local ge sp hst rs0 m0 -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. -Definition hsistate_exit_refines (hext: hsistate_exit) (ext: sistate_exit): Prop := +Definition hsistate_exit_refines ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := hsi_cond hext = si_cond ext /\ hsi_scondargs hext = si_scondargs ext - /\ hsistate_local_refines (hsi_elocal hext) (si_elocal ext) + /\ hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) /\ hsi_ifso hext = si_ifso ext. (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. -Definition hsistate_refines (hst: hsistate) (st:sistate): Prop := +Definition hsistate_refines ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st - /\ list_forall2 hsistate_exit_refines (hsi_exits hst) (si_exits st) - /\ hsistate_local_refines (hsi_local hst) (si_local st). + /\ list_forall2 (hsistate_exit_refines ge sp rs0 m0) (hsi_exits hst) (si_exits st) + /\ hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). -Lemma hsistate_refines_pceq st hst: - hsistate_refines hst st -> +Lemma hsistate_refines_pceq ge sp rs0 m0 st hst: + hsistate_refines ge sp rs0 m0 hst st -> (hsi_pc hst) = (si_pc st). Proof. Admitted. @@ -175,13 +175,16 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := | _ => None (* TODO jumptable ? *) end. -Lemma hsiexec_inst_none i st hst: +(* SB: est-ce utile ? +Lemma hsiexec_inst_none ge sp rs0 m0 i st hst: siexec_inst i st = None -> - hsistate_refines hst st -> + hsistate_refines ge sp rs0 m0 hst st -> hsiexec_inst i hst = None. Proof. Admitted. +*) +(* SB: inutile en principe Lemma hsiexec_inst_complete i st st' hst: siexec_inst i st = Some st' -> hsistate_refines hst st -> @@ -190,13 +193,14 @@ Lemma hsiexec_inst_complete i st st' hst: /\ hsistate_refines hst' st'. Proof. Admitted. +*) -Lemma hsiexec_inst_correct i hst hst' st: +Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: hsiexec_inst i hst = Some hst' -> - hsistate_refines hst st -> + hsistate_refines ge sp rs0 m0 hst st -> exists st', siexec_inst i st = Some st' - /\ hsistate_refines hst' st'. + /\ hsistate_refines ge sp rs0 m0 hst' st'. Proof. Admitted. @@ -209,6 +213,7 @@ Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate hsiexec_path p f hst1 end. +(* On ne devrait pas avoir besoin de ça ! Lemma hsiexec_path_complete ps f st st' hst: siexec_path ps f st = Some st' -> hsistate_refines hst st -> @@ -217,20 +222,21 @@ Lemma hsiexec_path_complete ps f st st' hst: /\ hsistate_refines hst' st'. Proof. Admitted. +*) Lemma hsiexec_path_correct ps f hst hst' st: hsiexec_path ps f hst = Some hst' -> - hsistate_refines hst st -> + forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 hst st -> exists st', siexec_path ps f st = Some st' - /\ hsistate_refines hst' st'. + /\ hsistate_refines ge sp rs0 m0 hst' st'. Proof. Admitted. Record hsstate := { hinternal:> hsistate; hfinal: sfval }. -Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := - hsistate_refines (hinternal hst) (internal st) +Definition hsstate_refines ge sp rs0 m0 (hst: hsstate) (st:sstate): Prop := + hsistate_refines ge sp rs0 m0 (hinternal hst) (internal st) /\ hfinal hst = final st. (* TODO *) Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := @@ -255,8 +261,8 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := | _ => Snone end. -Lemma hsexec_final_correct hsl sl i: - hsistate_local_refines hsl sl -> +Lemma hsexec_final_correct ge sp rs0 m0 hsl sl i: + hsistate_local_refines ge sp rs0 m0 hsl sl -> hsexec_final i hsl = sexec_final i sl. Proof. Admitted. @@ -265,8 +271,8 @@ Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := ni Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. -Remark init_hsistate_correct pc: - hsistate_refines (init_hsistate pc) (init_sistate pc). +Remark init_hsistate_correct ge sp rs0 m0 pc: + hsistate_refines ge sp rs0 m0 (init_hsistate pc) (init_sistate pc). Proof. Admitted. @@ -279,18 +285,21 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -(* FIXME - this lemma might be unprovable ? *) +(* FIXME - this lemma might be unprovable ? SB: inutile en principe Lemma hsistate_refines_determ hst st st': hsistate_refines hst st -> hsistate_refines hst st' -> st = st'. Proof. Admitted. +*) -Lemma hsexec_correct f pc st hst: +Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> - hsstate_refines hst st -> - sexec f pc = Some st. + exists st, sexec f pc = Some st /\ + forall ge sp rs0 m0, hsstate_refines ge sp rs0 m0 hst st. +Admitted. +(* Proof. unfold hsexec. intros H HREF. explore. unfold sexec. rewrite EQ. @@ -304,7 +313,10 @@ Proof. eapply hsistate_refines_determ in HIREF; eauto. congruence. - admit. Admitted. +*) + +(* NB: inutile en principe. Lemma hsexec_complete f pc st: sexec f pc = Some st -> exists hst, hsexec f pc = Some hst /\ hsstate_refines hst st. @@ -322,13 +334,12 @@ Proof. constructor; auto. simpl. destruct HREF as (_ & _ & HLREF). apply hsexec_final_correct. assumption. Qed. +*) Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := - forall st1, - hsstate_refines hst1 st1 -> - exists st2, - hsstate_refines hst2 st2 - /\ sstate_simu dm f st1 st2. + forall ge sp rs0 m0 st1 st2, + hsstate_refines ge sp rs0 m0 hst1 st1 -> + hsstate_refines ge sp rs0 m0 hst2 st2 -> sstate_simu dm f st1 st2. Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) @@ -338,21 +349,23 @@ Lemma hsstate_simub_correct dm f hst1 hst2: Proof. Admitted. +(* SB: Hum, le "exists st2" ne me plaît pas du tout ici... Definition hsexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := forall hst1, hsexec f1 pc1 = Some hst1 -> exists hst2, hsexec f2 pc2 = Some hst2 /\ hsstate_simu dm f1 hst1 hst2. -Lemma hsexec_sexec_simu dm f1 f2 pc1 pc2: - hsexec_simu dm f1 f2 pc1 pc2 -> sexec_simu dm f1 f2 pc1 pc2. + +Lemma hsexec_sexec_simu dm f1 f2 pc1 pc2 hst1 hst2: + hsexec f1 pc1 = Some hst1 -> + hsexec f2 pc2 = Some hst2 -> + hsstate_simu dm f1 hst1 hst2 -> sexec_simu dm f1 f2 pc1 pc2. Proof. - intro HESIMU. - intros st1 SEXEC. - eapply hsexec_complete in SEXEC. - destruct SEXEC as (hst & HSEXEC & HREF). + intros HSEXEC1 HSEXEC2 HESIMU. eapply HESIMU in HSEXEC. destruct HSEXEC as (hst' & HSEXEC' & HSSIMU). eapply HSSIMU in HREF. destruct HREF as (st' & HREF' & SSIMU). eapply hsexec_correct in HSEXEC'; eauto. Qed. +*) Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in @@ -369,14 +382,20 @@ Lemma simu_check_single_correct dm tf f pc1 pc2: simu_check_single dm f tf (pc2, pc1) = OK tt -> sexec_simu dm f tf pc1 pc2. Proof. - unfold simu_check_single. intro. explore. - assert (hsexec_simu dm f tf pc1 pc2). { - unfold hsexec_simu. intros st1 STEQ. assert (h = st1) by congruence. subst. - exists h0. constructor; auto. - apply hsstate_simub_correct. assumption. - } - apply hsexec_sexec_simu. assumption. -Qed. + unfold simu_check_single. intro. + unfold sexec_simu. + intros st1 SEXEC. + explore. + exploit hsexec_correct; eauto. + intros (st2 & SEXEC2 & REF2). + clear EQ0. (* now, useless in principle and harmful for the next [exploit] *) + exploit hsexec_correct; eauto. + intros (st0 & SEXEC1 & REF1). + rewrite SEXEC1 in SEXEC; inversion SEXEC; subst. + eexists; split; eauto. + eapply hsstate_simub_correct; eauto. + (* SB: issue: remaining goals are on the shelf... *) +Admitted. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with -- cgit From d6cc8135074a5c5650737b2f4d02982c53dc64b6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 2 Jul 2020 08:19:10 +0200 Subject: Fix hypothesis on environment in hsstate_simu --- kvx/lib/RTLpathSE_impl.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index e4ae1af8..cb645a92 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -337,9 +337,10 @@ Qed. *) Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := - forall ge sp rs0 m0 st1 st2, - hsstate_refines ge sp rs0 m0 hst1 st1 -> - hsstate_refines ge sp rs0 m0 hst2 st2 -> sstate_simu dm f st1 st2. + forall ge1 ge2 sp rs0 m0 st1 st2, + (forall s : ident, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> + hsstate_refines ge1 sp rs0 m0 hst1 st1 -> + hsstate_refines ge2 sp rs0 m0 hst2 st2 -> sstate_simu dm f st1 st2. Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) -- cgit From 54e9832f1c974921fbbee4b7a8d069d09eb55dc2 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 2 Jul 2020 18:27:24 +0200 Subject: Optimization of Iop symbolic execution --- kvx/lib/RTLpathSE_impl.v | 130 ++++++++++++++++++++++++++++++++++++--------- kvx/lib/RTLpathSE_theory.v | 12 +---- 2 files changed, 107 insertions(+), 35 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index cb645a92..fc95ab71 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -11,17 +11,7 @@ Require Import RTLpathSE_theory. Local Open Scope error_monad_scope. Local Open Scope option_monad_scope. -(* -(* TODO: for the implementation of the symbolic execution. - It is important to remove dependency of Op on memory. - Here is an way to formalize this *) - -Parameter mem_irrel: operation -> bool. (* TODO: TO PROVIDE FOR EACH BACKEND ? *) -Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_irrel o = true -> eval_operation ge sp o args m1 = eval_operation ge sp o args m2. - -*) - -(** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations *) +(** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) (* Implementation of local symbolic internal states TODO: use refined symbolic values instead of those from RTLpathSE_theory. @@ -38,18 +28,18 @@ Record hsistate_local := 1) we store a list of sval evaluations 2) we encode the symbolic regset by a PTree *) hsi_ok_lsval: list sval; - hsi_sreg: PTree.t sval + hsi_sreg:> PTree.t sval }. -Definition hsi_sreg_get (hst: hsistate_local) r: sval := - match PTree.get r (hsi_sreg hst) with +Definition hsi_sreg_get (hst: PTree.t sval) r: sval := + match PTree.get r hst with | None => Sinput r | Some sv => sv end. (* NB: short cut for (seval_sval ge sp (hsi_sreg_get hst r) rs0 m0) *) -Definition hsi_sreg_eval ge sp (hst: hsistate_local) r rs0 m0: option val := - match PTree.get r (hsi_sreg hst) with +Definition hsi_sreg_eval ge sp (hst: PTree.t sval) r rs0 m0: option val := + match PTree.get r hst with | None => Some (Regmap.get r rs0) | Some sv => seval_sval ge sp sv rs0 m0 end. @@ -118,22 +108,114 @@ Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := hsi_sreg:= hsi_sreg hst |}. +Local Open Scope lazy_bool_scope. + +(* naive function but sufficient on small lists *) +Fixpoint list_sval_expand (l: list_sval): list sval := + match l with + | Snil => nil + | Scons v l => v::(list_sval_expand l) + end. + (* NB: return [false] if the sv cannot fail *) Definition may_trap (sv: sval): bool := match sv with - (* | Sinput _ => false *) (* TODO: is it useful here ? *) - | Sload _ NOTRAP _ _ _ => false - | _ => true (* TODO: to refine for the refined operations *) + | Sload _ TRAP _ _ _ => true + | Sop op lsv _ => is_trapping_op op ||| negb (Nat.eqb (length (list_sval_expand lsv)) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) + | _ => false + end. + +Definition ok_on_args (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem) : Prop := + match sv with + | Sinput r => True + | Sop op l sm => seval_list_sval ge sp l rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None + | Sload sm trap chunk addr l => seval_list_sval ge sp l rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None end. -(* version without trapping of a may_trap operation. Identity if there is no such version *) -Definition try_untrap (sv: sval): sval := +Lemma lazy_orb_negb_false (b1 b2:bool): + (b1 ||| negb b2) = false <-> (b1 = false /\ b2 = true). +Proof. + unfold negb; explore; simpl; intuition (try congruence). +Qed. + +Lemma may_trap_correct (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): + may_trap sv = false -> + ok_on_args ge sp sv rs0 m0 -> + seval_sval ge sp sv rs0 m0 <> None. +Proof. + destruct sv; simpl; try congruence. + - rewrite lazy_orb_negb_false. intros (TRAP1 & TRAP2) (OK1 & OK2). + explore; try congruence. + eapply is_trapping_op_sound; eauto. + admit. (* TODO *) + - intros X (OK1 & OK2). + explore; try congruence. +Admitted. + +(* simplify a symbolic value before assignment to a register *) +Definition simplify (sv: sval): sval := match sv with | Sload sm TRAP chunk addr lsv => Sload sm NOTRAP chunk addr lsv + | Sop op lsv sm => + match is_move_operation op (list_sval_expand lsv) with + | Some arg => arg (* optimization of Omove *) + | None => + if op_depends_on_memory op then + sv + else + Sop op lsv Sinit (* magically remove the dependency on sm ! *) + end | _ => sv end. + +Lemma simplify_correct (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem) v: + seval_sval ge sp sv rs0 m0 = Some v -> seval_sval ge sp (simplify sv) rs0 m0 = Some v. +Proof. + destruct sv; simpl; auto. + - (* Sop *) + destruct (seval_list_sval _ _ _ _) as [args|] eqn: Hargs; try congruence. + destruct (seval_smem _ _ _ _) as [m|] eqn: Hm; try congruence. + intros Hv. + destruct (is_move_operation op (list_sval_expand lsv)) eqn: Hmove. + + exploit is_move_operation_correct; eauto. + intros (Hop & Hlsv); subst; simpl in *. + destruct lsv as [|v0 lsv]; simpl in *; try congruence. + destruct lsv as [|v1 lsv]; simpl in *; try congruence. + inversion Hmove; subst; clear Hmove Hlsv. + explore. auto. + + clear Hmove; destruct (op_depends_on_memory op) eqn: Hop; simpl; explore; try congruence. + inversion Hargs; subst. + erewrite op_depends_on_memory_correct; eauto. + - (* Sload *) + destruct trap; simpl; auto. + destruct (seval_list_sval _ _ _ _) as [args|] eqn: Hargs; try congruence. + destruct (eval_addressing _ _ _ _) as [a|] eqn: Ha; try congruence. + destruct (seval_smem _ _ _ _) as [m|] eqn: Hm; try congruence. + intros H; rewrite H; auto. +Qed. + +Definition red_PTree_set (r:reg) (sv: sval) (hst: PTree.t sval): PTree.t sval := + match sv with + | Sinput r' => + if Pos.eq_dec r r' + then PTree.remove r' hst + else PTree.set r sv hst + | _ => PTree.set r sv hst + end. + +Lemma red_PTree_set_correct (r:reg) (sv: sval) (hst: PTree.t sval) ge sp rs0 m0: + hsi_sreg_eval ge sp (red_PTree_set r sv hst) r rs0 m0 = hsi_sreg_eval ge sp (PTree.set r sv hst) r rs0 m0. +Proof. + destruct sv; simpl; auto. + destruct (Pos.eq_dec r r0); auto. + subst; unfold hsi_sreg_eval. + rewrite PTree.grs, PTree.gss; simpl; auto. +Qed. + +(* naive version: + +@Cyril: tu peux utiliser la version naive dans un premier temps pour simplifier les preuves... -(* naive version Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := {| hsi_lsmem := hsi_lsmem hst; hsi_ok_lsval := sv::(hsi_ok_lsval hst); @@ -143,7 +225,7 @@ Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := {| hsi_lsmem := hsi_lsmem hst; hsi_ok_lsval := if may_trap sv then sv::(hsi_ok_lsval hst) else hsi_ok_lsval hst; - hsi_sreg := PTree.set r (try_untrap sv) (hsi_sreg hst) |}. + hsi_sreg := red_PTree_set r (simplify sv) (hsi_sreg hst) |}. Definition hsist_set_local (hst: hsistate) (pc: node) (nxt: hsistate_local): option hsistate := Some {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= nxt |}. @@ -152,7 +234,7 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := match i with | Inop pc' => hsist_set_local hst pc' hst.(hsi_local) - | Iop op args dst pc' => (* TODO: deal with non-trapping operations *) + | Iop op args dst pc' => let prev := hst.(hsi_local) in let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in let next := hslocal_set_sreg prev dst (Sop op vargs (hsi_smem_get prev)) in diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index ae702510..57621029 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -43,22 +43,12 @@ Ltac explore := explore_hyp; | [ x : unit |- _ ] => destruct x end. *) -(* NOTE for the implementation of the symbolic execution. - It is important to remove dependency of Op on memory. - Here is an way to formalize this - -Parameter mem_irrel: operation -> bool. -Hypothesis eval_operation_mem_irrel: forall (ge: RTL.genv) sp o args m1 m2, mem_irrel o = true -> eval_operation ge sp o args m1 = eval_operation ge sp o args m2. - -*) - (** * Syntax and semantics of symbolic values *) -(* FIXME - might have to add non trapping loads *) (* symbolic value *) Inductive sval := | Sinput (r: reg) - | Sop (op:operation) (lsv: list_sval) (sm: smem) (* TODO for the implementation: add an additionnal case without dependency on sm*) + | 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 -- cgit From 5e2b7a9d07482c75660295f927234cdec84679f2 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 2 Jul 2020 22:52:00 +0200 Subject: slightly simpler code --- kvx/lib/RTLpathSE_impl.v | 152 +++++++++++++++++++++++++++-------------------- 1 file changed, 88 insertions(+), 64 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index fc95ab71..2f6adcfa 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -23,7 +23,7 @@ Record hsistate_local := { (** [hsi_lsmem] represents the list of smem symbolic evaluations. The first one of the list is the current smem *) - hsi_lsmem: list smem; + hsi_lsmem:> list smem; (** For the values in registers: 1) we store a list of sval evaluations 2) we encode the symbolic regset by a PTree *) @@ -44,15 +44,15 @@ Definition hsi_sreg_eval ge sp (hst: PTree.t sval) r rs0 m0: option val := | Some sv => seval_sval ge sp sv rs0 m0 end. -Definition hsi_smem_get (hst: hsistate_local): smem := - match hsi_lsmem hst with +Definition hsi_smem_get (hst: list smem): smem := + match hst with | nil => Sinit | sm::_ => sm end. (* NB: short cut for (seval_smem ge sp (hsi_smem_get hst) rs0 m0) *) -Definition hsi_smem_eval ge sp (hst: hsistate_local) rs0 m0 : option mem := - match hsi_lsmem hst with +Definition hsi_smem_eval ge sp (hst: list smem) rs0 m0 : option mem := + match hst with | nil => Some m0 | sm::_ => seval_smem ge sp sm rs0 m0 end. @@ -64,8 +64,8 @@ Definition sok_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) ( /\ forall (r: reg), seval_sval ge sp (si_sreg st r) rs0 m0 <> None. Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := - forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None - /\ forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None. + (forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None) + /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistate_local) := @@ -99,7 +99,7 @@ Proof. Admitted. (** * Symbolic execution of one internal step - TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values + non-trapping operations + non-memory dependent operations + TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values *) Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := @@ -108,28 +108,53 @@ Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := hsi_sreg:= hsi_sreg hst |}. -Local Open Scope lazy_bool_scope. +Lemma hsok_local_set_mem hst sm ge sp rs0 m0: + hsok_local ge sp (hslocal_set_smem hst sm) rs0 m0 + <-> (hsok_local ge sp hst rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). +Proof. + unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). +Qed. -(* naive function but sufficient on small lists *) -Fixpoint list_sval_expand (l: list_sval): list sval := - match l with - | Snil => nil - | Scons v l => v::(list_sval_expand l) - end. +Lemma sok_local_set_mem st sm ge sp rs0 m0: + sok_local ge sp (slocal_set_smem st sm) rs0 m0 + <-> (sok_local ge sp st rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). +Proof. + unfold slocal_set_smem, sok_local; simpl; intuition (subst; eauto). +Qed. -(* NB: return [false] if the sv cannot fail *) -Definition may_trap (sv: sval): bool := - match sv with - | Sload _ TRAP _ _ _ => true - | Sop op lsv _ => is_trapping_op op ||| negb (Nat.eqb (length (list_sval_expand lsv)) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) - | _ => false +Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st sm: + hsistate_local_refines ge sp rs0 m0 hst st -> + hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst sm) (slocal_set_smem st sm). +Proof. + unfold hsistate_local_refines. + intros (OK & MEM & VAL); simpl. + rewrite hsok_local_set_mem. + rewrite sok_local_set_mem. + rewrite! OK. + intuition. +Qed. + +(* 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) (lsv: list sval) (sm: smem): sval := + match rsv with + | Rop op => Sop op (list_sval_inj lsv) sm + | Rload trap chunk addr => Sload sm trap chunk addr (list_sval_inj lsv) end. +Coercion root_apply: root_sval >-> Funclass. -Definition ok_on_args (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem) : Prop := - match sv with - | Sinput r => True - | Sop op l sm => seval_list_sval ge sp l rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None - | Sload sm trap chunk addr l => seval_list_sval ge sp l rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None +Local Open Scope lazy_bool_scope. + +(* NB: return [false] if the rsv cannot fail *) +Definition may_trap (rsv: root_sval) (lsv: list sval) (sm: smem): bool := + match rsv with + | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lsv) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) + | Rload TRAP _ _ => true + | _ => false end. Lemma lazy_orb_negb_false (b1 b2:bool): @@ -138,55 +163,53 @@ Proof. unfold negb; explore; simpl; intuition (try congruence). Qed. -Lemma may_trap_correct (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): - may_trap sv = false -> - ok_on_args ge sp sv rs0 m0 -> - seval_sval ge sp sv rs0 m0 <> None. +Lemma may_trap_correct (ge: RTL.genv) (sp:val) (rsv: root_sval) (rs0: regset) (m0: mem) (lsv: list sval) (sm: smem): + may_trap rsv lsv sm = false -> + seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None -> + seval_smem ge sp sm rs0 m0 <> None -> + seval_sval ge sp (rsv lsv sm) rs0 m0 <> None. Proof. - destruct sv; simpl; try congruence. - - rewrite lazy_orb_negb_false. intros (TRAP1 & TRAP2) (OK1 & OK2). + destruct rsv; simpl; try congruence. + - rewrite lazy_orb_negb_false. intros (TRAP1 & TRAP2) OK1 OK2. explore; try congruence. eapply is_trapping_op_sound; eauto. admit. (* TODO *) - - intros X (OK1 & OK2). + - intros X OK1 OK2. explore; try congruence. Admitted. (* simplify a symbolic value before assignment to a register *) -Definition simplify (sv: sval): sval := - match sv with - | Sload sm TRAP chunk addr lsv => Sload sm NOTRAP chunk addr lsv - | Sop op lsv sm => - match is_move_operation op (list_sval_expand lsv) with +Definition simplify (rsv: root_sval) lsv sm: sval := + match rsv with + | Rload TRAP chunk addr => Sload sm NOTRAP chunk addr (list_sval_inj lsv) + | Rop op => + match is_move_operation op lsv with | Some arg => arg (* optimization of Omove *) | None => if op_depends_on_memory op then - sv + rsv lsv sm else - Sop op lsv Sinit (* magically remove the dependency on sm ! *) + Sop op (list_sval_inj lsv) Sinit (* magically remove the dependency on sm ! *) end - | _ => sv + | _ => rsv lsv sm end. -Lemma simplify_correct (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem) v: - seval_sval ge sp sv rs0 m0 = Some v -> seval_sval ge sp (simplify sv) rs0 m0 = Some v. +Lemma simplify_correct (rsv: root_sval) lsv sm (ge: RTL.genv) (sp:val) (rs0: regset) (m0: mem) v: + seval_sval ge sp (rsv lsv sm) rs0 m0 = Some v -> seval_sval ge sp (simplify rsv lsv sm) rs0 m0 = Some v. Proof. - destruct sv; simpl; auto. - - (* Sop *) + destruct rsv; simpl; auto. + - (* Rop *) destruct (seval_list_sval _ _ _ _) as [args|] eqn: Hargs; try congruence. destruct (seval_smem _ _ _ _) as [m|] eqn: Hm; try congruence. intros Hv. - destruct (is_move_operation op (list_sval_expand lsv)) eqn: Hmove. + destruct (is_move_operation _ _) eqn: Hmove. + exploit is_move_operation_correct; eauto. intros (Hop & Hlsv); subst; simpl in *. - destruct lsv as [|v0 lsv]; simpl in *; try congruence. - destruct lsv as [|v1 lsv]; simpl in *; try congruence. - inversion Hmove; subst; clear Hmove Hlsv. explore. auto. + clear Hmove; destruct (op_depends_on_memory op) eqn: Hop; simpl; explore; try congruence. inversion Hargs; subst. erewrite op_depends_on_memory_correct; eauto. - - (* Sload *) + - (* Rload *) destruct trap; simpl; auto. destruct (seval_list_sval _ _ _ _) as [args|] eqn: Hargs; try congruence. destruct (eval_addressing _ _ _ _) as [a|] eqn: Ha; try congruence. @@ -215,17 +238,18 @@ Qed. (* naive version: @Cyril: tu peux utiliser la version naive dans un premier temps pour simplifier les preuves... - -Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := +*) + +Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := + let sv := rsv lsv sm in {| hsi_lsmem := hsi_lsmem hst; hsi_ok_lsval := sv::(hsi_ok_lsval hst); - hsi_sreg:= PTree.set r sv (hsi_sreg hst) |}. -*) + hsi_sreg:= PTree.set r sv (hsi_sreg hst) |}. -Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (sv:sval) := +Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := {| hsi_lsmem := hsi_lsmem hst; - hsi_ok_lsval := if may_trap sv then sv::(hsi_ok_lsval hst) else hsi_ok_lsval hst; - hsi_sreg := red_PTree_set r (simplify sv) (hsi_sreg hst) |}. + hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; + hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. Definition hsist_set_local (hst: hsistate) (pc: node) (nxt: hsistate_local): option hsistate := Some {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= nxt |}. @@ -236,18 +260,18 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := hsist_set_local hst pc' hst.(hsi_local) | Iop op args dst pc' => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - let next := hslocal_set_sreg prev dst (Sop op vargs (hsi_smem_get prev)) in + let vargs := List.map (hsi_sreg_get prev) args in + let next := hslocal_set_sreg prev dst (Rop op) vargs (hsi_smem_get prev) in hsist_set_local hst pc' next - | Istore chunk addr args src pc' => + | Iload trap chunk addr args dst pc' => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in + let vargs := List.map (hsi_sreg_get prev) args in + let next := hslocal_set_sreg prev dst (Rload trap chunk addr) vargs (hsi_smem_get prev) in hsist_set_local hst pc' next - | Iload trap chunk addr args dst pc' => + | Istore chunk addr args src pc' => let prev := hst.(hsi_local) in let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - let next := hslocal_set_sreg prev dst (Sload (hsi_smem_get prev) trap chunk addr vargs) in + let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in hsist_set_local hst pc' next | Icond cond args ifso ifnot _ => let prev := hst.(hsi_local) in -- cgit From 1dc5ea9d02a46413940398a415519ae4b9eb4108 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 3 Jul 2020 08:21:32 +0200 Subject: sketch the proof of symbolic execution of one instruction --- kvx/lib/RTLpathSE_impl.v | 75 +++++++++++++++++++++++++++++++++++----------- kvx/lib/RTLpathSE_theory.v | 28 +++++++---------- 2 files changed, 68 insertions(+), 35 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 2f6adcfa..902d08ab 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -96,25 +96,27 @@ Lemma hsistate_refines_pceq ge sp rs0 m0 st hst: hsistate_refines ge sp rs0 m0 hst st -> (hsi_pc hst) = (si_pc st). Proof. -Admitted. + unfold hsistate_refines; intuition. +Qed. + +Lemma hsistate_refines_imp_local_refines ge sp rs0 m0 hst st: + hsistate_refines ge sp rs0 m0 hst st -> + hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). +Proof. + unfold hsistate_refines; intuition. +Qed. (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values *) +(** ** Assignment of memory *) Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := {| hsi_lsmem := sm::hsi_lsmem hst; hsi_ok_lsval := hsi_ok_lsval hst; hsi_sreg:= hsi_sreg hst |}. -Lemma hsok_local_set_mem hst sm ge sp rs0 m0: - hsok_local ge sp (hslocal_set_smem hst sm) rs0 m0 - <-> (hsok_local ge sp hst rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). -Proof. - unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). -Qed. - Lemma sok_local_set_mem st sm ge sp rs0 m0: sok_local ge sp (slocal_set_smem st sm) rs0 m0 <-> (sok_local ge sp st rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). @@ -122,6 +124,13 @@ Proof. unfold slocal_set_smem, sok_local; simpl; intuition (subst; eauto). Qed. +Lemma hsok_local_set_mem hst sm ge sp rs0 m0: + hsok_local ge sp (hslocal_set_smem hst sm) rs0 m0 + <-> (hsok_local ge sp hst rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). +Proof. + unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). +Qed. + Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st sm: hsistate_local_refines ge sp rs0 m0 hst st -> hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst sm) (slocal_set_smem st sm). @@ -134,6 +143,21 @@ Proof. intuition. Qed. +(** ** Assignment of local state *) + +Definition hsist_set_local (hst: hsistate) (pc: node) (hnxt: hsistate_local): hsistate := + {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= hnxt |}. + +Lemma hsist_set_local_correct ge sp rs0 m0 hst st pc hnxt nxt: + hsistate_refines ge sp rs0 m0 hst st -> + hsistate_local_refines ge sp rs0 m0 hnxt nxt -> + hsistate_refines ge sp rs0 m0 (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt). +Proof. + unfold hsistate_refines; simpl; intuition. +Qed. + +(** ** Assignment of registers *) + (* locally new symbolic values during symbolic execution *) Inductive root_sval: Type := | Rop (op:operation) @@ -236,43 +260,49 @@ Proof. Qed. (* naive version: - -@Cyril: tu peux utiliser la version naive dans un premier temps pour simplifier les preuves... -*) +@Cyril: éventuellement, tu peux utiliser la version naive dans un premier temps pour simplifier les preuves... Definition naive_hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := let sv := rsv lsv sm in {| hsi_lsmem := hsi_lsmem hst; hsi_ok_lsval := sv::(hsi_ok_lsval hst); hsi_sreg:= PTree.set r sv (hsi_sreg hst) |}. +*) Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := {| hsi_lsmem := hsi_lsmem hst; hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. -Definition hsist_set_local (hst: hsistate) (pc: node) (nxt: hsistate_local): option hsistate := - Some {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= nxt |}. +Lemma hslocal_set_sreg_correct ge sp rs0 m0 hst st r (rsv:root_sval) lsv sm sv': + hsistate_local_refines ge sp rs0 m0 hst st -> + seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None -> + seval_smem ge sp sm rs0 m0 <> None -> + sv' = rsv lsv sm -> + hsistate_local_refines ge sp rs0 m0 (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). +Admitted. + +(** ** Execution of one instruction *) Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := match i with | Inop pc' => - hsist_set_local hst pc' hst.(hsi_local) + Some (hsist_set_local hst pc' hst.(hsi_local)) | Iop op args dst pc' => let prev := hst.(hsi_local) in let vargs := List.map (hsi_sreg_get prev) args in let next := hslocal_set_sreg prev dst (Rop op) vargs (hsi_smem_get prev) in - hsist_set_local hst pc' next + Some (hsist_set_local hst pc' next) | Iload trap chunk addr args dst pc' => let prev := hst.(hsi_local) in let vargs := List.map (hsi_sreg_get prev) args in let next := hslocal_set_sreg prev dst (Rload trap chunk addr) vargs (hsi_smem_get prev) in - hsist_set_local hst pc' next + Some (hsist_set_local hst pc' next) | Istore chunk addr args src pc' => let prev := hst.(hsi_local) in let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in - hsist_set_local hst pc' next + Some (hsist_set_local hst pc' next) | Icond cond args ifso ifnot _ => let prev := hst.(hsi_local) in let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in @@ -301,6 +331,9 @@ Proof. Admitted. *) + +Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines: core. + Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: hsiexec_inst i hst = Some hst' -> hsistate_refines ge sp rs0 m0 hst st -> @@ -308,6 +341,14 @@ Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: siexec_inst i st = Some st' /\ hsistate_refines ge sp rs0 m0 hst' st'. Proof. + destruct i; simpl; intros STEPI; inversion_clear STEPI; try (discriminate). + - (* Inop *) intros REF; eexists; split; eauto. + - (* Iop *) intros REF; eexists; split; eauto. + eapply hsist_set_local_correct; eauto. + eapply hslocal_set_sreg_correct; eauto. + + admit. (* TODO *) + + admit. (* TODO *) + + simpl. (* TODO *) Admitted. Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 57621029..a8affbba 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -353,6 +353,7 @@ Proof. inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence. Qed. +(* TODO: remove this JUNK ? Lemma ssem_local_exclude_sabort_exit ge sp st ext lx rs m rs' m': ssem_local ge sp (si_local st) rs m rs' m' -> all_fallthrough ge sp (si_exits st) rs m -> @@ -367,6 +368,7 @@ Proof. but we NEVER actually evaluated the si_elocal from the sistate_exit ! So we cannot prove a lack of abort on the si_elocal.. We must change the definitions *) Abort. +*) Lemma ssem_local_exclude_sabort ge sp st rs m rs' m': ssem_local ge sp (si_local st) rs m rs' m' -> @@ -464,29 +466,29 @@ Definition slocal_set_smem (st:sistate_local) (sm:smem) := si_sreg:= st.(si_sreg); si_smem:= sm |}. -Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): option sistate := - Some {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. +Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): sistate := + {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}. Definition siexec_inst (i: instruction) (st: sistate): option sistate := match i with | Inop pc' => - sist_set_local st pc' st.(si_local) + 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 - sist_set_local st pc' next + 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 - sist_set_local st pc' next + Some (sist_set_local st pc' next) | Istore chunk addr args src pc' => let prev := st.(si_local) in let vargs := list_sval_inj (List.map prev.(si_sreg) args) in let next := slocal_set_smem prev (Sstore prev.(si_smem) chunk addr vargs (prev.(si_sreg) src)) in - sist_set_local st pc' next + 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 @@ -741,20 +743,11 @@ Fixpoint siexec_path (path:nat) (f: function) (st: sistate): option sistate := siexec_path p f st1 end. -Lemma sist_set_local_preserves_si_exits st pc loc st': - sist_set_local st pc loc = Some st' -> - si_exits st' = si_exits st. -Proof. - intros. inv H. simpl. reflexivity. -Qed. - 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; try discriminate. - all: try (intros; left; eapply sist_set_local_preserves_si_exits; eauto; fail). - intros. inv H. right. simpl. eauto. + 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: @@ -764,8 +757,7 @@ Lemma siexec_inst_preserves_allfu ge sp ext lx rs0 m0 st st' i: Proof. intros ALLFU SISTEP. destruct ALLFU as (ISTAIL & ALLF). constructor; eauto. - destruct i; try discriminate; simpl in SISTEP; try (erewrite sist_set_local_preserves_si_exits; eauto; fail). - inv SISTEP. simpl. constructor. assumption. + destruct i; simpl in SISTEP; inversion_clear SISTEP; simpl; (discriminate || eauto). Qed. Lemma siexec_path_correct_false ge sp f rs0 m0 st' is: -- cgit From c2f5fc97ef81c01a07b1bc9063e80e512f6945c2 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 3 Jul 2020 12:26:05 +0200 Subject: fix the proof sketch --- kvx/lib/RTLpathSE_impl.v | 52 ++++++++++++++++++++---------------------------- 1 file changed, 22 insertions(+), 30 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 902d08ab..b47b7288 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -131,16 +131,17 @@ Proof. unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). Qed. -Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st sm: +Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st hsm sm: hsistate_local_refines ge sp rs0 m0 hst st -> - hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst sm) (slocal_set_smem st sm). + (hsok_local ge sp hst rs0 m0 -> seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> + hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm). Proof. unfold hsistate_local_refines. - intros (OK & MEM & VAL); simpl. + intros (OK & MEM & VAL) HSM; simpl. rewrite hsok_local_set_mem. rewrite sok_local_set_mem. rewrite! OK. - intuition. + intuition congruence. Qed. (** ** Assignment of local state *) @@ -274,11 +275,13 @@ Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. +Definition ok_args hst ge sp rs0 m0 lsv sm := + hsok_local ge sp hst rs0 m0 -> (seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None). + Lemma hslocal_set_sreg_correct ge sp rs0 m0 hst st r (rsv:root_sval) lsv sm sv': hsistate_local_refines ge sp rs0 m0 hst st -> - seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None -> - seval_smem ge sp sm rs0 m0 <> None -> - sv' = rsv lsv sm -> + ok_args hst ge sp rs0 m0 lsv sm -> + (hsok_local ge sp hst rs0 m0 -> seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0)-> hsistate_local_refines ge sp rs0 m0 (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). Admitted. @@ -311,28 +314,7 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := | _ => None (* TODO jumptable ? *) end. -(* SB: est-ce utile ? -Lemma hsiexec_inst_none ge sp rs0 m0 i st hst: - siexec_inst i st = None -> - hsistate_refines ge sp rs0 m0 hst st -> - hsiexec_inst i hst = None. -Proof. -Admitted. -*) - -(* SB: inutile en principe -Lemma hsiexec_inst_complete i st st' hst: - siexec_inst i st = Some st' -> - hsistate_refines hst st -> - exists hst', - hsiexec_inst i hst = Some hst' - /\ hsistate_refines hst' st'. -Proof. -Admitted. -*) - - -Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines: core. +Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines hslocal_set_mem_correct: core. Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: hsiexec_inst i hst = Some hst' -> @@ -347,8 +329,18 @@ Proof. eapply hsist_set_local_correct; eauto. eapply hslocal_set_sreg_correct; eauto. + admit. (* TODO *) + + simpl. admit. (* TODO *) + - (* Iload *) intros REF; eexists; split; eauto. + eapply hsist_set_local_correct; eauto. + eapply hslocal_set_sreg_correct; eauto. + admit. (* TODO *) - + simpl. (* TODO *) + + simpl. admit. (* TODO *) + - (* Istore *) intros REF; eexists; split; eauto. + eapply hsist_set_local_correct; eauto. + eapply hslocal_set_mem_correct; eauto. + admit. (* TODO: the easiest case of all these TODO ? *) + - (* Icond *) intros REF; eexists; split; eauto. + admit. (* TODO *) Admitted. Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := -- cgit From c0a3ff04d79468e15af4d81c7732ae5fc75d599c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 3 Jul 2020 15:10:32 +0200 Subject: Fixing ge, sp, rs, m issue --- kvx/lib/RTLpathSE_impl.v | 17 +++++++++-------- 1 file changed, 9 insertions(+), 8 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index b47b7288..b54ff42c 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -374,7 +374,8 @@ Admitted. Record hsstate := { hinternal:> hsistate; hfinal: sfval }. -Definition hsstate_refines ge sp rs0 m0 (hst: hsstate) (st:sstate): Prop := +(* CS: made hsstate_refines independent of ge, sp, rs0 and m0 *) +Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 (hinternal hst) (internal st) /\ hfinal hst = final st. (* TODO *) @@ -436,7 +437,7 @@ Admitted. Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> exists st, sexec f pc = Some st /\ - forall ge sp rs0 m0, hsstate_refines ge sp rs0 m0 hst st. + hsstate_refines hst st. Admitted. (* Proof. @@ -476,10 +477,11 @@ Qed. *) Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := - forall ge1 ge2 sp rs0 m0 st1 st2, - (forall s : ident, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - hsstate_refines ge1 sp rs0 m0 hst1 st1 -> - hsstate_refines ge2 sp rs0 m0 hst2 st2 -> sstate_simu dm f st1 st2. + forall st1 st2, + (* CS: this hypothesis is not needed? *) +(* (forall s : ident, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> *) + hsstate_refines hst1 st1 -> + hsstate_refines hst2 st2 -> sstate_simu dm f st1 st2. Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) @@ -534,8 +536,7 @@ Proof. rewrite SEXEC1 in SEXEC; inversion SEXEC; subst. eexists; split; eauto. eapply hsstate_simub_correct; eauto. - (* SB: issue: remaining goals are on the shelf... *) -Admitted. +Qed. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with -- cgit From bfb75cc289647db3489de79d2abfaacaef66330a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 3 Jul 2020 16:01:42 +0200 Subject: Scope issue.. --- kvx/lib/RTLpathSE_impl.v | 23 ++++++++++++----------- 1 file changed, 12 insertions(+), 11 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index b54ff42c..ca1f7602 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -438,22 +438,23 @@ Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> exists st, sexec f pc = Some st /\ hsstate_refines hst st. -Admitted. -(* Proof. - unfold hsexec. intros H HREF. explore. + unfold hsexec. intro. explore_hyp. unfold sexec. rewrite EQ. + eapply hsiexec_path_correct in EQ0; [| apply init_hsistate_correct]. + destruct EQ0 as (st0 & SPATH & REF0). rewrite SPATH. + erewrite <- hsistate_refines_pceq; eauto. rewrite EQ1. destruct (hsiexec_inst i h) eqn:HINST. - - destruct HREF as (HIREF & HFREF). simpl in *. - eapply hsiexec_path_correct in EQ0; [| apply init_hsistate_correct]. - destruct EQ0 as (st2 & SIPATH & HIREF2). rewrite SIPATH. - erewrite <- hsistate_refines_pceq; eauto. rewrite EQ1. - eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST & HIREF'). - rewrite SINST. destruct st as (sti & stf). simpl in *. - eapply hsistate_refines_determ in HIREF; eauto. congruence. + - destruct (siexec_inst i st0) eqn:SINST. + + eexists. constructor; [reflexivity|]. constructor; auto. simpl. + eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & HREF'). + assert (s = st') by congruence. subst. admit. (* TODO - issue with scope ! *) + + (* absurd case *) + eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & _). + rewrite SINST in SINST'. discriminate. - admit. Admitted. -*) + (* NB: inutile en principe. -- cgit From f4ff9e187efc0d344a80d7133412a1972b36de46 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 3 Jul 2020 16:54:21 +0200 Subject: Version without dependency on rs0, m0, ge and sp ? --- kvx/lib/RTLpathSE_impl.v | 104 +++++++++++++++++++++++++++-------------------- 1 file changed, 60 insertions(+), 44 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index ca1f7602..4feb44ba 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -68,7 +68,7 @@ Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistate_local) := +Definition hsistate_local_refines (hst: hsistate_local) (st: sistate_local) := forall ge sp rs0 m0, (sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) /\ (hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) /\ (hsok_local ge sp hst rs0 m0 -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). @@ -78,30 +78,30 @@ Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistat Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. -Definition hsistate_exit_refines ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := +Definition hsistate_exit_refines (hext: hsistate_exit) (ext: sistate_exit): Prop := hsi_cond hext = si_cond ext /\ hsi_scondargs hext = si_scondargs ext - /\ hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) + /\ hsistate_local_refines (hsi_elocal hext) (si_elocal ext) /\ hsi_ifso hext = si_ifso ext. (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. -Definition hsistate_refines ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop := +Definition hsistate_refines (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st - /\ list_forall2 (hsistate_exit_refines ge sp rs0 m0) (hsi_exits hst) (si_exits st) - /\ hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). + /\ list_forall2 hsistate_exit_refines (hsi_exits hst) (si_exits st) + /\ hsistate_local_refines (hsi_local hst) (si_local st). -Lemma hsistate_refines_pceq ge sp rs0 m0 st hst: - hsistate_refines ge sp rs0 m0 hst st -> +Lemma hsistate_refines_pceq st hst: + hsistate_refines hst st -> (hsi_pc hst) = (si_pc st). Proof. unfold hsistate_refines; intuition. Qed. -Lemma hsistate_refines_imp_local_refines ge sp rs0 m0 hst st: - hsistate_refines ge sp rs0 m0 hst st -> - hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). +Lemma hsistate_refines_imp_local_refines hst st: + hsistate_refines hst st -> + hsistate_local_refines (hsi_local hst) (si_local st). Proof. unfold hsistate_refines; intuition. Qed. @@ -131,17 +131,32 @@ Proof. unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). Qed. -Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st hsm sm: - hsistate_local_refines ge sp rs0 m0 hst st -> - (hsok_local ge sp hst rs0 m0 -> seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> - hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm). +Lemma hslocal_set_mem_correct hst st hsm sm: + hsistate_local_refines hst st -> + (forall ge sp rs0 m0, hsok_local ge sp hst rs0 m0 -> seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> + hsistate_local_refines (hslocal_set_smem hst hsm) (slocal_set_smem st sm). Proof. - unfold hsistate_local_refines. - intros (OK & MEM & VAL) HSM; simpl. + intros LOCREF. intros SMEMEQ. + intros ge sp rs0 m0. edestruct LOCREF as (OKEQ & SMEMEQ' & REGEQ). constructor; [|constructor]. + - rewrite hsok_local_set_mem. + rewrite sok_local_set_mem. + constructor. + + intros (OKL & SMEMN). constructor. 2: rewrite SMEMEQ; auto. + all: rewrite <- OKEQ; assumption. + + intros (HOKL & HSMEM). rewrite OKEQ. constructor; auto. + rewrite <- SMEMEQ; auto. + - rewrite hsok_local_set_mem. intros (HOKL & HSMEM). + simpl. apply SMEMEQ; assumption. + - rewrite hsok_local_set_mem. intros (HOKL & HSMEM). + simpl. intuition congruence. + + +(* unfold hsistate_local_refines. + intros (OK & MEM & VAL) HSM; simpl. intros ge sp rs0 m0. rewrite hsok_local_set_mem. rewrite sok_local_set_mem. rewrite! OK. - intuition congruence. + intuition congruence. *) Qed. (** ** Assignment of local state *) @@ -149,10 +164,10 @@ Qed. Definition hsist_set_local (hst: hsistate) (pc: node) (hnxt: hsistate_local): hsistate := {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= hnxt |}. -Lemma hsist_set_local_correct ge sp rs0 m0 hst st pc hnxt nxt: - hsistate_refines ge sp rs0 m0 hst st -> - hsistate_local_refines ge sp rs0 m0 hnxt nxt -> - hsistate_refines ge sp rs0 m0 (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt). +Lemma hsist_set_local_correct hst st pc hnxt nxt: + hsistate_refines hst st -> + hsistate_local_refines hnxt nxt -> + hsistate_refines (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt). Proof. unfold hsistate_refines; simpl; intuition. Qed. @@ -278,11 +293,12 @@ Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm Definition ok_args hst ge sp rs0 m0 lsv sm := hsok_local ge sp hst rs0 m0 -> (seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None). -Lemma hslocal_set_sreg_correct ge sp rs0 m0 hst st r (rsv:root_sval) lsv sm sv': - hsistate_local_refines ge sp rs0 m0 hst st -> +Lemma hslocal_set_sreg_correct hst st r (rsv:root_sval) lsv sm sv': + hsistate_local_refines hst st -> + (forall ge sp rs0 m0, ok_args hst ge sp rs0 m0 lsv sm -> - (hsok_local ge sp hst rs0 m0 -> seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0)-> - hsistate_local_refines ge sp rs0 m0 (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). + (hsok_local ge sp hst rs0 m0 -> seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0) ) -> + hsistate_local_refines (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). Admitted. (** ** Execution of one instruction *) @@ -316,25 +332,25 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines hslocal_set_mem_correct: core. -Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: +Lemma hsiexec_inst_correct i hst hst' st: hsiexec_inst i hst = Some hst' -> - hsistate_refines ge sp rs0 m0 hst st -> + hsistate_refines hst st -> exists st', siexec_inst i st = Some st' - /\ hsistate_refines ge sp rs0 m0 hst' st'. + /\ hsistate_refines hst' st'. Proof. destruct i; simpl; intros STEPI; inversion_clear STEPI; try (discriminate). - (* Inop *) intros REF; eexists; split; eauto. - (* Iop *) intros REF; eexists; split; eauto. eapply hsist_set_local_correct; eauto. - eapply hslocal_set_sreg_correct; eauto. - + admit. (* TODO *) - + simpl. admit. (* TODO *) + eapply hslocal_set_sreg_correct; eauto. admit. +(* + admit. (* TODO *) + + simpl. admit. (* TODO *) *) - (* Iload *) intros REF; eexists; split; eauto. eapply hsist_set_local_correct; eauto. - eapply hslocal_set_sreg_correct; eauto. - + admit. (* TODO *) - + simpl. admit. (* TODO *) + eapply hslocal_set_sreg_correct; eauto. admit. +(* + admit. (* TODO *) + + simpl. admit. (* TODO *) *) - (* Istore *) intros REF; eexists; split; eauto. eapply hsist_set_local_correct; eauto. eapply hslocal_set_mem_correct; eauto. @@ -365,18 +381,18 @@ Admitted. Lemma hsiexec_path_correct ps f hst hst' st: hsiexec_path ps f hst = Some hst' -> - forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 hst st -> + hsistate_refines hst st -> exists st', siexec_path ps f st = Some st' - /\ hsistate_refines ge sp rs0 m0 hst' st'. + /\ hsistate_refines hst' st'. Proof. Admitted. Record hsstate := { hinternal:> hsistate; hfinal: sfval }. (* CS: made hsstate_refines independent of ge, sp, rs0 and m0 *) -Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := forall ge sp rs0 m0, - hsistate_refines ge sp rs0 m0 (hinternal hst) (internal st) +Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := + hsistate_refines (hinternal hst) (internal st) /\ hfinal hst = final st. (* TODO *) Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := @@ -401,8 +417,8 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := | _ => Snone end. -Lemma hsexec_final_correct ge sp rs0 m0 hsl sl i: - hsistate_local_refines ge sp rs0 m0 hsl sl -> +Lemma hsexec_final_correct hsl sl i: + hsistate_local_refines hsl sl -> hsexec_final i hsl = sexec_final i sl. Proof. Admitted. @@ -411,8 +427,8 @@ Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := ni Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. -Remark init_hsistate_correct ge sp rs0 m0 pc: - hsistate_refines ge sp rs0 m0 (init_hsistate pc) (init_sistate pc). +Remark init_hsistate_correct pc: + hsistate_refines (init_hsistate pc) (init_sistate pc). Proof. Admitted. @@ -448,7 +464,7 @@ Proof. - destruct (siexec_inst i st0) eqn:SINST. + eexists. constructor; [reflexivity|]. constructor; auto. simpl. eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & HREF'). - assert (s = st') by congruence. subst. admit. (* TODO - issue with scope ! *) + assert (s = st') by congruence. subst. assumption. (* admit. (* TODO - issue with scope ! *) *) + (* absurd case *) eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & _). rewrite SINST in SINST'. discriminate. -- cgit From 3ca730035fa3bd051abfc7a2c8c3e78f31a8447c Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 3 Jul 2020 19:22:01 +0200 Subject: hsexec_correct proved (under admitted lemmas). --- kvx/lib/RTLpathSE_impl.v | 116 ++++++++++++++++++++--------------------------- 1 file changed, 48 insertions(+), 68 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index ca1f7602..b5ead763 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -79,26 +79,19 @@ Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. Definition hsistate_exit_refines ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := - hsi_cond hext = si_cond ext - /\ hsi_scondargs hext = si_scondargs ext + hsi_cond hext = si_cond ext (* FIXME: should not depend on ge sp rs0 m0 ? *) + /\ hsi_scondargs hext = si_scondargs ext (* FIXME: should not depend on ge sp rs0 m0 ? *) /\ hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) - /\ hsi_ifso hext = si_ifso ext. + /\ hsi_ifso hext = si_ifso ext. (* FIXME: should not depend on ge sp rs0 m0 ? *) (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. Definition hsistate_refines ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop := - hsi_pc hst = si_pc st - /\ list_forall2 (hsistate_exit_refines ge sp rs0 m0) (hsi_exits hst) (si_exits st) + (* FOR THE PROOF: the condition "hsi_pc hst = si_pc st" should not depend on "ge sp rs0 m0" *) + list_forall2 (hsistate_exit_refines ge sp rs0 m0) (hsi_exits hst) (si_exits st) /\ hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). -Lemma hsistate_refines_pceq ge sp rs0 m0 st hst: - hsistate_refines ge sp rs0 m0 hst st -> - (hsi_pc hst) = (si_pc st). -Proof. - unfold hsistate_refines; intuition. -Qed. - Lemma hsistate_refines_imp_local_refines ge sp rs0 m0 hst st: hsistate_refines ge sp rs0 m0 hst st -> hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). @@ -106,6 +99,8 @@ Proof. unfold hsistate_refines; intuition. Qed. +Local Hint Resolve hsistate_refines_imp_local_refines: core. + (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values *) @@ -316,33 +311,37 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines hslocal_set_mem_correct: core. -Lemma hsiexec_inst_correct ge sp rs0 m0 i hst hst' st: +Lemma hsiexec_inst_correct_Some i hst hst' st: hsiexec_inst i hst = Some hst' -> - hsistate_refines ge sp rs0 m0 hst st -> - exists st', - siexec_inst i st = Some st' - /\ hsistate_refines ge sp rs0 m0 hst' st'. + exists st', siexec_inst i st = Some st' /\ hsi_pc hst' = si_pc st' /\ + forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 hst st -> hsistate_refines ge sp rs0 m0 hst' st'. Proof. - destruct i; simpl; intros STEPI; inversion_clear STEPI; try (discriminate). - - (* Inop *) intros REF; eexists; split; eauto. - - (* Iop *) intros REF; eexists; split; eauto. + destruct i; simpl; intros STEPI; inversion_clear STEPI; discriminate || eexists; split; eauto; + simpl; split; auto; intros ge sp rs0 m0 REF. + - (* Iop *) eapply hsist_set_local_correct; eauto. eapply hslocal_set_sreg_correct; eauto. + admit. (* TODO *) + simpl. admit. (* TODO *) - - (* Iload *) intros REF; eexists; split; eauto. + - (* Iload *) eapply hsist_set_local_correct; eauto. eapply hslocal_set_sreg_correct; eauto. + admit. (* TODO *) + simpl. admit. (* TODO *) - - (* Istore *) intros REF; eexists; split; eauto. + - (* Istore *) eapply hsist_set_local_correct; eauto. eapply hslocal_set_mem_correct; eauto. admit. (* TODO: the easiest case of all these TODO ? *) - - (* Icond *) intros REF; eexists; split; eauto. + - (* Icond *) admit. (* TODO *) Admitted. +Lemma hsiexec_inst_correct_None i hst st: + hsiexec_inst i hst = None -> siexec_inst i st = None. +Proof. + destruct i; simpl; congruence. +Qed. + Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := match path with | O => Some hst @@ -352,32 +351,22 @@ Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate hsiexec_path p f hst1 end. -(* On ne devrait pas avoir besoin de ça ! -Lemma hsiexec_path_complete ps f st st' hst: - siexec_path ps f st = Some st' -> - hsistate_refines hst st -> - exists hst', - hsiexec_path ps f hst = Some hst' - /\ hsistate_refines hst' st'. -Proof. -Admitted. -*) - Lemma hsiexec_path_correct ps f hst hst' st: - hsiexec_path ps f hst = Some hst' -> - forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 hst st -> + hsiexec_path ps f hst = Some hst' -> hsi_pc hst = si_pc st -> exists st', - siexec_path ps f st = Some st' - /\ hsistate_refines ge sp rs0 m0 hst' st'. + siexec_path ps f st = Some st' /\ hsi_pc hst' = si_pc st' /\ + (forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 hst st + -> hsistate_refines ge sp rs0 m0 hst' st'). Proof. Admitted. Record hsstate := { hinternal:> hsistate; hfinal: sfval }. (* CS: made hsstate_refines independent of ge, sp, rs0 and m0 *) -Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := forall ge sp rs0 m0, - hsistate_refines ge sp rs0 m0 (hinternal hst) (internal st) - /\ hfinal hst = final st. (* TODO *) +Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := + hsi_pc (hinternal hst) = si_pc (internal st) + /\ (forall ge sp rs0 m0, hsistate_refines ge sp rs0 m0 (hinternal hst) (internal st)) + /\ hfinal hst = final st. (* FIXME: too strong ! *) Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := match i with @@ -401,8 +390,8 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := | _ => Snone end. -Lemma hsexec_final_correct ge sp rs0 m0 hsl sl i: - hsistate_local_refines ge sp rs0 m0 hsl sl -> +Lemma hsexec_final_correct hsl sl i: + (forall ge sp rs0 m0, hsistate_local_refines ge sp rs0 m0 hsl sl) -> hsexec_final i hsl = sexec_final i sl. Proof. Admitted. @@ -425,37 +414,28 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -(* FIXME - this lemma might be unprovable ? SB: inutile en principe -Lemma hsistate_refines_determ hst st st': - hsistate_refines hst st -> - hsistate_refines hst st' -> - st = st'. -Proof. -Admitted. -*) +Local Hint Resolve init_hsistate_correct hsexec_final_correct: core. Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> - exists st, sexec f pc = Some st /\ - hsstate_refines hst st. + exists st, sexec f pc = Some st /\ hsstate_refines hst st. Proof. unfold hsexec. intro. explore_hyp. - unfold sexec. rewrite EQ. - eapply hsiexec_path_correct in EQ0; [| apply init_hsistate_correct]. - destruct EQ0 as (st0 & SPATH & REF0). rewrite SPATH. - erewrite <- hsistate_refines_pceq; eauto. rewrite EQ1. + unfold sexec. + rewrite EQ. + eapply hsiexec_path_correct in EQ0. + destruct EQ0 as (st0 & SPATH & PC0 & REF0). rewrite SPATH. + erewrite <- PC0. rewrite EQ1. destruct (hsiexec_inst i h) eqn:HINST. - - destruct (siexec_inst i st0) eqn:SINST. - + eexists. constructor; [reflexivity|]. constructor; auto. simpl. - eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & HREF'). - assert (s = st') by congruence. subst. admit. (* TODO - issue with scope ! *) - + (* absurd case *) - eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & _). - rewrite SINST in SINST'. discriminate. - - admit. -Admitted. - - + + eapply hsiexec_inst_correct_Some in HINST; eauto. + destruct HINST as (st1 & EQ2 & PC2 & REF2). + rewrite EQ2. + repeat (econstructor; eauto). + + erewrite hsiexec_inst_correct_None; eauto. + repeat (econstructor; eauto). + simpl; eauto. + + simpl; auto. +Qed. (* NB: inutile en principe. Lemma hsexec_complete f pc st: -- cgit From f6b525b000b971f9faaa9b1b2d176bdf70b01ed1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 6 Jul 2020 17:07:41 +0200 Subject: init_hsistate_correct --- kvx/lib/RTLpathSE_impl.v | 34 +++++++++++++++++++++++++++++----- 1 file changed, 29 insertions(+), 5 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 4feb44ba..c21af3e6 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -359,6 +359,13 @@ Proof. admit. (* TODO *) Admitted. +Lemma hsiexec_inst_error i hst st: + hsiexec_inst i hst = None -> + hsistate_refines hst st -> + siexec_inst i st = None. +Proof. +Admitted. + Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := match path with | O => Some hst @@ -425,12 +432,27 @@ Admitted. Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. +Remark init_hsistate_local_correct: + hsistate_local_refines init_hsistate_local init_sistate_local. +Proof. + constructor; constructor. + - intro. destruct H as (_ & SMEM & SVAL). constructor; simpl in *; [contradiction|]. + intros. destruct H; [|contradiction]. subst. discriminate. + - intro. destruct H as (SVAL & SMEM). constructor; [simpl; auto|]. + constructor; simpl; discriminate. + - intro. reflexivity. + - intros. simpl. unfold hsi_sreg_eval. rewrite PTree.gempty. reflexivity. +Qed. + Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. Remark init_hsistate_correct pc: hsistate_refines (init_hsistate pc) (init_sistate pc). Proof. -Admitted. + constructor; constructor; simpl; auto. + - apply list_forall2_nil. + - apply init_hsistate_local_correct. +Qed. Definition hsexec (f: function) (pc:node): option hsstate := SOME path <- (fn_path f)!pc IN @@ -468,10 +490,12 @@ Proof. + (* absurd case *) eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & _). rewrite SINST in SINST'. discriminate. - - admit. -Admitted. - - + - destruct (siexec_inst i st0) eqn:SINST. + + (* absurd *) + eapply hsiexec_inst_error in HINST; eauto. rewrite SINST in HINST. discriminate. + + eexists. constructor; [reflexivity|]. constructor; auto. simpl. + apply hsexec_final_correct; auto. +Qed. (* NB: inutile en principe. Lemma hsexec_complete f pc st: -- cgit From 84359fec32837617a52a42784fb7d307d60ce00b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 6 Jul 2020 17:08:18 +0200 Subject: ssem_final with less dependencies --- kvx/lib/RTLpathSE_theory.v | 25 ++++++++++++++----------- 1 file changed, 14 insertions(+), 11 deletions(-) diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index a8affbba..7e83f772 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -949,14 +949,14 @@ Qed. End SEVAL_BUILTIN_ARG. -Inductive ssem_final (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := +Inductive ssem_final (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (npc: node) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop := | exec_Snone rs m: - ssem_final pge ge sp st stack f rs0 m0 Snone rs m E0 (State stack f sp st.(si_pc) rs m) + ssem_final pge ge sp npc stack f rs0 m0 Snone rs m E0 (State stack f sp npc rs m) | exec_Scall rs m sig svos lsv args res pc fd: sfind_function pge ge sp svos rs0 m0 = Some fd -> funsig fd = sig -> seval_list_sval ge sp lsv rs0 m0 = Some args -> - ssem_final pge ge sp st stack f rs0 m0 (Scall sig svos lsv res pc) rs m + ssem_final pge ge sp npc 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 pge ge sp svos rs0 m0 = Some fd -> @@ -964,23 +964,23 @@ Inductive ssem_final (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (st: sistate) s sp = Vptr stk Ptrofs.zero -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> seval_list_sval ge sp lsv rs0 m0 = Some args -> - ssem_final pge ge sp st stack f rs0 m0 (Stailcall sig svos lsv) rs m + ssem_final pge ge sp npc 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' -> - ssem_final pge ge sp st stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m + ssem_final pge ge sp npc 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: seval_sval ge sp sv rs0 m0 = Some (Vint n) -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - ssem_final pge ge sp st stack f rs0 m0 (Sjumptable sv tbl) rs m + ssem_final pge ge sp npc stack f rs0 m0 (Sjumptable sv tbl) 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 => seval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v -> - ssem_final pge ge sp st stack f rs0 m0 (Sreturn osv) rs m + ssem_final pge ge sp npc stack f rs0 m0 (Sreturn osv) rs m E0 (Returnstate stack v m') . @@ -994,7 +994,7 @@ Inductive ssem pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m | ssem_normal is t s: is.(icontinue) = true -> ssem_internal ge sp st rs0 m0 is -> - ssem_final pge ge sp st stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> + ssem_final pge ge sp st.(si_pc) stack f rs0 m0 st.(final) is.(irs) is.(imem) t s -> ssem pge ge sp st stack f rs0 m0 t s . @@ -1092,7 +1092,7 @@ Lemma sexec_final_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s: ssem_local ge sp (si_local st) rs0 m0 rs m -> path_last_step ge pge stack f sp pc rs m t s -> siexec_inst i st = None -> - ssem_final pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s. + ssem_final pge ge sp pc stack f rs0 m0 (sexec_final i (si_local st)) rs m t s. Proof. intros PC1 PC2 (PRE&MEM®) LAST Hi. destruct LAST; subst; try_simplify_someHyps; simpl. + (* Snone *) destruct i; simpl in Hi |- *; unfold sist_set_local in Hi; try congruence. @@ -1115,7 +1115,7 @@ Lemma sexec_final_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s: (fn_code f) ! pc = Some i -> pc = st.(si_pc) -> ssem_local ge sp (si_local st) rs0 m0 rs m -> - ssem_final pge ge sp st stack f rs0 m0 (sexec_final i (si_local st)) rs m t s -> + ssem_final pge ge sp pc stack f rs0 m0 (sexec_final i (si_local st)) rs m t s -> siexec_inst i st = None -> path_last_step ge pge stack f sp pc rs m t s. Proof. @@ -1515,7 +1515,10 @@ Proof. eapply exec_normal_exit; eauto. eapply sexec_final_complete; eauto. * congruence. - * unfold ssem_local in * |- *. intuition try congruence. + * unfold ssem_local in * |- *. + destruct SEM1 as (A & B & C). constructor; [|constructor]; eauto. + intro r. congruence. + * congruence. Qed. (** * Simulation of RTLpath code w.r.t symbolic execution *) -- cgit From 26eeda87795cee0a387d58cf4ff8ffa9f63039ba Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 7 Jul 2020 18:24:36 +0200 Subject: Avancement --- kvx/lib/RTLpathSE_impl.v | 374 +++++++++++++++++++++++++++++++--------- kvx/lib/RTLpathSchedulerproof.v | 4 +- 2 files changed, 290 insertions(+), 88 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index c21af3e6..2c4ae088 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -7,6 +7,7 @@ Require Import Op Registers. Require Import RTL RTLpath. Require Import Errors Duplicate. Require Import RTLpathSE_theory. +Require Import Axioms. Local Open Scope error_monad_scope. Local Open Scope option_monad_scope. @@ -57,7 +58,7 @@ Definition hsi_smem_eval ge sp (hst: list smem) rs0 m0 : option mem := | sm::_ => seval_smem ge sp sm rs0 m0 end. -(* negation of sabort_local *) +(* (* negation of sabort_local *) Definition sok_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem): Prop := (st.(si_pre) ge sp rs0 m0) /\ seval_smem ge sp st.(si_smem) rs0 m0 <> None @@ -65,13 +66,13 @@ Definition sok_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) ( Definition hsok_local ge sp (hst: hsistate_local) rs0 m0: Prop := (forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None) - /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). + /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). *) (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) Definition hsistate_local_refines (hst: hsistate_local) (st: sistate_local) := forall ge sp rs0 m0, - (sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) - /\ (hsok_local ge sp hst rs0 m0 -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) - /\ (hsok_local ge sp hst rs0 m0 -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). +(* (sok_local ge sp st rs0 m0 <-> hsok_local ge sp hst rs0 m0) *) + ((* hsok_local ge sp hst rs0 m0 -> *) hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) + /\ ((* hsok_local ge sp hst rs0 m0 -> *) forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) @@ -117,7 +118,7 @@ Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := hsi_sreg:= hsi_sreg hst |}. -Lemma sok_local_set_mem st sm ge sp rs0 m0: +(* Lemma sok_local_set_mem st sm ge sp rs0 m0: sok_local ge sp (slocal_set_smem st sm) rs0 m0 <-> (sok_local ge sp st rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). Proof. @@ -129,34 +130,26 @@ Lemma hsok_local_set_mem hst sm ge sp rs0 m0: <-> (hsok_local ge sp hst rs0 m0 /\ seval_smem ge sp sm rs0 m0 <> None). Proof. unfold hslocal_set_smem, hsok_local; simpl; intuition (subst; eauto). -Qed. +Qed. *) Lemma hslocal_set_mem_correct hst st hsm sm: hsistate_local_refines hst st -> - (forall ge sp rs0 m0, hsok_local ge sp hst rs0 m0 -> seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> + (forall ge sp rs0 m0, (* hsok_local ge sp hst rs0 m0 -> *) seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> hsistate_local_refines (hslocal_set_smem hst hsm) (slocal_set_smem st sm). Proof. intros LOCREF. intros SMEMEQ. - intros ge sp rs0 m0. edestruct LOCREF as (OKEQ & SMEMEQ' & REGEQ). constructor; [|constructor]. - - rewrite hsok_local_set_mem. + intros ge sp rs0 m0. edestruct LOCREF as ((* OKEQ & *) SMEMEQ' & REGEQ). constructor. +(* - rewrite hsok_local_set_mem. rewrite sok_local_set_mem. constructor. + intros (OKL & SMEMN). constructor. 2: rewrite SMEMEQ; auto. all: rewrite <- OKEQ; assumption. + intros (HOKL & HSMEM). rewrite OKEQ. constructor; auto. - rewrite <- SMEMEQ; auto. - - rewrite hsok_local_set_mem. intros (HOKL & HSMEM). + rewrite <- SMEMEQ; auto. *) + - (* rewrite hsok_local_set_mem. intros (HOKL & HSMEM). *) simpl. apply SMEMEQ; assumption. - - rewrite hsok_local_set_mem. intros (HOKL & HSMEM). - simpl. intuition congruence. - - -(* unfold hsistate_local_refines. - intros (OK & MEM & VAL) HSM; simpl. intros ge sp rs0 m0. - rewrite hsok_local_set_mem. - rewrite sok_local_set_mem. - rewrite! OK. - intuition congruence. *) + - (* rewrite hsok_local_set_mem. intros (HOKL & HSMEM). *) + simpl. eassumption. (* intuition congruence. *) Qed. (** ** Assignment of local state *) @@ -203,6 +196,7 @@ Proof. unfold negb; explore; simpl; intuition (try congruence). Qed. +(* not used yet *) Lemma may_trap_correct (ge: RTL.genv) (sp:val) (rsv: root_sval) (rs0: regset) (m0: mem) (lsv: list sval) (sm: smem): may_trap rsv lsv sm = false -> seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None -> @@ -290,14 +284,15 @@ Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. -Definition ok_args hst ge sp rs0 m0 lsv sm := - hsok_local ge sp hst rs0 m0 -> (seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None). +Definition ok_args ge sp rs0 m0 lsv sm := + (* hsok_local ge sp hst rs0 m0 -> *) + (seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None). Lemma hslocal_set_sreg_correct hst st r (rsv:root_sval) lsv sm sv': hsistate_local_refines hst st -> (forall ge sp rs0 m0, - ok_args hst ge sp rs0 m0 lsv sm -> - (hsok_local ge sp hst rs0 m0 -> seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0) ) -> + ok_args (* hst *) ge sp rs0 m0 lsv sm -> + ((* hsok_local ge sp hst rs0 m0 -> *) seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0) ) -> hsistate_local_refines (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). Admitted. @@ -332,6 +327,34 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := Local Hint Resolve hsist_set_local_correct hsistate_refines_imp_local_refines hslocal_set_mem_correct: core. +Lemma seval_sval_refines hst st ge sp rs0 m0 p: +(* sok_local ge sp st rs0 m0 -> *) + hsistate_local_refines hst st -> + seval_sval ge sp (hsi_sreg_get hst p) rs0 m0 = seval_sval ge sp (si_sreg st p) rs0 m0. +Proof. + intros (* OKL *) HREF. destruct (HREF ge sp rs0 m0) as ((* OKEQ & *) _ & RSEQ). + (* apply OKEQ in OKL. rewrite <- (RSEQ OKL p). *) + rewrite <- (RSEQ p). + unfold hsi_sreg_get, hsi_sreg_eval. destruct (hst ! p); auto. +Qed. + +Lemma seval_list_sval_refines hsl sl ge sp rs0 m0 l: +(* sok_local ge sp sl rs0 m0 -> *) + hsistate_local_refines hsl sl -> + seval_list_sval ge sp (list_sval_inj (map (hsi_sreg_get hsl) l)) rs0 m0 = + seval_list_sval ge sp (list_sval_inj (map (si_sreg sl) l)) rs0 m0. +Proof. + intros (* OKL *) HLREF. destruct (HLREF ge sp rs0 m0) as ((* OKEQ & *) _ & RSEQ). + induction l; simpl; auto. + erewrite seval_sval_refines; eauto. rewrite IHl. destruct (seval_sval ge sp _ _ _); reflexivity. +Qed. + +Lemma seval_smem_refines hst st rs0 m0 ge sp: + hsistate_local_refines hst st -> + seval_smem ge sp (hsi_smem_get hst) rs0 m0 = seval_smem ge sp (si_smem st) rs0 m0. +Proof. +Admitted. + Lemma hsiexec_inst_correct i hst hst' st: hsiexec_inst i hst = Some hst' -> hsistate_refines hst st -> @@ -354,11 +377,15 @@ Proof. - (* Istore *) intros REF; eexists; split; eauto. eapply hsist_set_local_correct; eauto. eapply hslocal_set_mem_correct; eauto. - admit. (* TODO: the easiest case of all these TODO ? *) + intros. simpl. destruct REF as (_ & _ & LREF). + erewrite seval_list_sval_refines; eauto. + erewrite seval_smem_refines; eauto. + erewrite seval_sval_refines; eauto. - (* Icond *) intros REF; eexists; split; eauto. admit. (* TODO *) Admitted. + Lemma hsiexec_inst_error i hst st: hsiexec_inst i hst = None -> hsistate_refines hst st -> @@ -395,12 +422,21 @@ Lemma hsiexec_path_correct ps f hst hst' st: Proof. Admitted. +(* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', + ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) + +(* FIXME - too strong, let's change it later.. *) +Definition hfinal_refines (hfv fv: sfval) := hfv = fv. + +Remark hfinal_refines_snone: hfinal_refines Snone Snone. +Proof. +Admitted. + Record hsstate := { hinternal:> hsistate; hfinal: sfval }. -(* CS: made hsstate_refines independent of ge, sp, rs0 and m0 *) Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := hsistate_refines (hinternal hst) (internal st) - /\ hfinal hst = final st. (* TODO *) + /\ hfinal_refines (hfinal hst) (final st). Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := match i with @@ -424,23 +460,59 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := | _ => Snone end. +(* Lemma local_refines_sreg_get hsl sl ge sp rs0 m0: + hsistate_local_refines hsl sl -> + sok_local ge sp sl rs0 m0 -> + hsi_sreg_get hsl = si_sreg sl. +Proof. + intros HREF SOKL. apply functional_extensionality. intro r. + destruct (HREF ge sp rs0 m0) as (OKEQ & MEMEQ & RSEQ). + apply OKEQ in SOKL. pose (RSEQ SOKL r) as EQ. + unfold hsi_sreg_get. +Admitted. *) + +Lemma sfind_function_conserves hsl sl pge ge sp s rs0 m0: + hsistate_local_refines hsl sl -> + sfind_function pge ge sp (sum_left_map (hsi_sreg_get hsl) s) rs0 m0 = + sfind_function pge ge sp (sum_left_map (si_sreg sl) s) rs0 m0. +Admitted. + Lemma hsexec_final_correct hsl sl i: hsistate_local_refines hsl sl -> - hsexec_final i hsl = sexec_final i sl. + hfinal_refines (hsexec_final i hsl) (sexec_final i sl). Proof. +(* destruct i; simpl; intros HLREF; try (apply hfinal_refines_snone). + (* Scall *) + - constructor. + + intro. inv H. constructor; auto. + ++ erewrite <- sfind_function_conserves; eauto. + ++ erewrite <- seval_list_sval_refines; eauto. + + intro. inv H. constructor; auto. + ++ erewrite sfind_function_conserves; eauto. + ++ erewrite seval_list_sval_refines; eauto. + (* Stailcall *) + - admit. + (* Sbuiltin *) + - admit. + (* Sjumptable *) + - admit. + (* Sreturn *) + - admit. *) Admitted. + + Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. Remark init_hsistate_local_correct: hsistate_local_refines init_hsistate_local init_sistate_local. Proof. - constructor; constructor. - - intro. destruct H as (_ & SMEM & SVAL). constructor; simpl in *; [contradiction|]. - intros. destruct H; [|contradiction]. subst. discriminate. - - intro. destruct H as (SVAL & SMEM). constructor; [simpl; auto|]. - constructor; simpl; discriminate. - - intro. reflexivity. + constructor. +(* - intro. destruct H as (_ & SMEM & SVAL). constructor; simpl in *; [contradiction|]. + intros. destruct H; [|contradiction]. subst. discriminate. *) +(* - intro. destruct H as (SVAL & SMEM). constructor; [simpl; auto|]. + constructor; simpl; discriminate. *) + - reflexivity. - intros. simpl. unfold hsi_sreg_eval. rewrite PTree.gempty. reflexivity. Qed. @@ -463,15 +535,6 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -(* FIXME - this lemma might be unprovable ? SB: inutile en principe -Lemma hsistate_refines_determ hst st st': - hsistate_refines hst st -> - hsistate_refines hst st' -> - st = st'. -Proof. -Admitted. -*) - Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> exists st, sexec f pc = Some st /\ @@ -484,9 +547,9 @@ Proof. erewrite <- hsistate_refines_pceq; eauto. rewrite EQ1. destruct (hsiexec_inst i h) eqn:HINST. - destruct (siexec_inst i st0) eqn:SINST. - + eexists. constructor; [reflexivity|]. constructor; auto. simpl. + + eexists. constructor; [reflexivity|]. constructor; auto; simpl; [|apply hfinal_refines_snone]. eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & HREF'). - assert (s = st') by congruence. subst. assumption. (* admit. (* TODO - issue with scope ! *) *) + assert (s = st') by congruence. subst. assumption. + (* absurd case *) eapply hsiexec_inst_correct in HINST; eauto. destruct HINST as (st' & SINST' & _). rewrite SINST in SINST'. discriminate. @@ -497,58 +560,197 @@ Proof. apply hsexec_final_correct; auto. Qed. -(* NB: inutile en principe. -Lemma hsexec_complete f pc st: - sexec f pc = Some st -> - exists hst, hsexec f pc = Some hst /\ hsstate_refines hst st. -Proof. - unfold sexec. intro. explore_hyp. - eapply hsiexec_path_complete in EQ0; [| apply init_hsistate_correct]. - destruct EQ0 as (hst & HPATH & HREF). - eexists. constructor. - + unfold hsexec. rewrite EQ. rewrite HPATH. - erewrite hsistate_refines_pceq; eauto. rewrite EQ1. reflexivity. - + destruct (siexec_inst i s) eqn:EINST. - - eapply hsiexec_inst_complete in EINST; eauto. destruct EINST as (hst' & HEINST & HREF'). - rewrite HEINST. constructor; auto. - - eapply hsiexec_inst_none in EINST; eauto. rewrite EINST. - constructor; auto. simpl. - destruct HREF as (_ & _ & HLREF). apply hsexec_final_correct. assumption. +Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. + +Definition hsilocal_simu dm f (hst1 hst2: hsistate_local): Prop := forall st1 st2, + hsistate_local_refines hst1 st1 -> + hsistate_local_refines hst2 st2 -> + silocal_simu dm f st1 st2. + +Lemma hsilocal_simub_correct dm f hst1 hst2: + hsilocal_simub dm f hst1 hst2 = OK tt -> + hsilocal_simu dm f hst1 hst2. +Proof. +Admitted. + +Lemma hsilocal_simub_eq dm f hsl1 hsl2: + hsilocal_simub dm f hsl1 hsl2 = OK tt -> + hsi_sreg_get (hsi_sreg hsl1) = hsi_sreg_get (hsi_sreg hsl2) + /\ hsi_smem_get (hsi_lsmem hsl1) = hsi_smem_get (hsi_lsmem hsl2). +Proof. +Admitted. + +Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := + if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then + do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); + do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); + revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) + else Error (msg "siexit_simub: conditions do not match") +. + +Definition hsiexit_simu dm f hse1 hse2: Prop := forall se1 se2, + hsistate_exit_refines hse1 se1 -> + hsistate_exit_refines hse2 se2 -> + siexit_simu dm f se1 se2. + +(* Lemma seval_condition_refines hst hst' ge sp cond args rs m: + hsi_smem_get hst = hsi_smem_get hst' -> + seval_condition ge sp cond args (si_smem hst) rs m = seval_condition ge sp cond args (si_smem hst') rs m. *) + + +Lemma seval_condition_refines hst st ge sp cond args rs m: + hsistate_local_refines hst st -> + seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. +Proof. +Admitted. + +Lemma hsiexit_simub_correct dm f hse1 hse2: + hsiexit_simub dm f hse1 hse2 = OK tt -> + hsiexit_simu dm f hse1 hse2. +Proof. + unfold hsiexit_simub. intro. explore. + apply list_sval_simub_correct in EQ0. exploit hsilocal_simub_eq; eauto. + intros (SREGEQ & SMEMEQ). + apply hsilocal_simub_correct in EQ2. + apply revmap_check_single_correct in EQ3. + intros se1 se2 HEREF1 HEREF2. + destruct HEREF1 as (HCOND1 & HARGS1 & HLREF1 & HIFSO1). + destruct HEREF2 as (HCOND2 & HARGS2 & HLREF2 & HIFSO2). + constructor. + - unfold siexit_simu_ssem. intros sp ge ge' GFS LOK rs m is SEMEXIT. + destruct SEMEXIT as (SCOND & SLOCAL & SIFSO). + let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). + eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). + destruct is2 as [icont ipc irs imem]. simpl in *. + exists (mk_istate icont (si_ifso se2) irs imem). simpl. + constructor; auto. + + repeat (constructor; auto). + rewrite <- HARGS2. rewrite <- EQ0. + rewrite <- HCOND2. rewrite <- e. + erewrite <- seval_condition_refines; eauto. rewrite <- SMEMEQ. + rewrite HARGS1. rewrite HCOND1. erewrite seval_condition_refines; eauto. + erewrite seval_condition_preserved; eauto. + + unfold istate_simu in *. destruct (icontinue is) eqn:ICONT. + * destruct ISIMU as (A & B & C). constructor; auto. + * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. + repeat (constructor; auto). simpl in *. congruence. + - unfold siexit_simu_cond. intros sp ge ge' GFS rs m. erewrite seval_condition_preserved; eauto. + rewrite <- HARGS2. rewrite <- HCOND2. erewrite <- seval_condition_refines; eauto. + rewrite <- HARGS1. rewrite <- HCOND1. erewrite <- seval_condition_refines; eauto. + congruence. +Qed. + +Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := + match lhse1 with + | nil => + match lhse2 with + | nil => OK tt + | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") + end + | hse1 :: lhse1 => + match lhse2 with + | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") + | hse2 :: lhse2 => + do _ <- hsiexit_simub dm f hse1 hse2; + do _ <- hsiexits_simub dm f lhse1 lhse2; + OK tt + end + end. + +Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit): Prop := forall se1 se2, + list_forall2 hsistate_exit_refines hse1 se1 -> + list_forall2 hsistate_exit_refines hse2 se2 -> + siexits_simu dm f se1 se2. + +Lemma hsiexits_simub_correct dm f lhse1: forall lhse2, + hsiexits_simub dm f lhse1 lhse2 = OK tt -> + hsiexits_simu dm f lhse1 lhse2. +Proof. + induction lhse1. + - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREF1 HEREF2. + inv HEREF1. inv HEREF2. constructor. + - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. + fold hsiexits_simub in EQ1. + apply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. + intros se1 se2 HEREF1 HEREF2. inv HEREF1. inv HEREF2. constructor; auto. + apply EQ1; assumption. +Qed. + +Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := + do _ <- hsilocal_simub dm f (hsi_local hst1) (hsi_local hst2); + do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); + OK tt. + +Definition hsistate_simu dm f (hst1 hst2: hsistate): Prop := forall st1 st2, + hsistate_refines hst1 st1 -> + hsistate_refines hst2 st2 -> + sistate_simu dm f st1 st2. + +(* Very adapted from sistate_simub_correct ---> should delete sistate_simub_correct after *) +Lemma hsistate_simub_correct dm f hst1 hst2: + hsistate_simub dm f hst1 hst2 = OK tt -> + hsistate_simu f dm hst1 hst2. +Proof. + unfold hsistate_simub. intro. explore. unfold hsistate_simu. + intros st1 st2 HREF1 HREF2 sp ge1 ge2 GFS LOK rs m is1 SEMI. + destruct HREF1 as (PCEQ1 & HEREF1 & HLREF1). destruct HREF2 as (PCEQ2 & HEREF2 & HLREF2). + exploit hsiexits_simub_correct; eauto. intro ESIMU. + unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. + - destruct SEMI as (SSEML & PCEQ & ALLFU). + exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). + destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. + exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. + + unfold ssem_internal. simpl. + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. + destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. + rewrite <- CONTEQ. rewrite ICONT. + constructor; [|constructor]; [eassumption | auto | eapply siexits_simu_all_fallthrough; eauto]. + + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. + destruct ISIMU as (A & B & C). simpl in *. constructor; auto. + - destruct SEMI as (ext & lx & SSEME & ALLFU). + exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). + assert (EXTSIMU: siexit_simu dm f ext ext2). { + eapply list_forall2_nth_error; eauto. + - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. + - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. + assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). + congruence. } + destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. + destruct SSEME as (is2 & SSEME2 & ISIMU). + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). + destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). + unfold ssem_internal. exists is2. + rewrite <- CONTEQ. rewrite ICONT. constructor; auto. + + eexists; eexists; constructor; eauto. + + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). + Unshelve. apply dm. (* Very weird to have this *) Qed. -*) Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := forall st1 st2, (* CS: this hypothesis is not needed? *) (* (forall s : ident, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> *) hsstate_refines hst1 st1 -> - hsstate_refines hst2 st2 -> sstate_simu dm f st1 st2. + hsstate_refines hst2 st2 -> sstate_simu f dm st1 st2. -Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) +Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := + do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); + do u2 <- sfval_simub dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2); + OK tt. Lemma hsstate_simub_correct dm f hst1 hst2: hsstate_simub dm f hst1 hst2 = OK tt -> - hsstate_simu dm f hst1 hst2. + hsstate_simu f dm hst1 hst2. Proof. -Admitted. - -(* SB: Hum, le "exists st2" ne me plaît pas du tout ici... -Definition hsexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall hst1, hsexec f1 pc1 = Some hst1 -> - exists hst2, hsexec f2 pc2 = Some hst2 /\ hsstate_simu dm f1 hst1 hst2. - - -Lemma hsexec_sexec_simu dm f1 f2 pc1 pc2 hst1 hst2: - hsexec f1 pc1 = Some hst1 -> - hsexec f2 pc2 = Some hst2 -> - hsstate_simu dm f1 hst1 hst2 -> sexec_simu dm f1 f2 pc1 pc2. -Proof. - intros HSEXEC1 HSEXEC2 HESIMU. - eapply HESIMU in HSEXEC. destruct HSEXEC as (hst' & HSEXEC' & HSSIMU). - eapply HSSIMU in HREF. destruct HREF as (st' & HREF' & SSIMU). - eapply hsexec_correct in HSEXEC'; eauto. + unfold hsstate_simub. intro. explore. + apply sfval_simub_correct in EQ1. apply hsistate_simub_correct in EQ. + intros st1 st2 (LREF & FREF) (LREF' & FREF'). + constructor; [auto|]. + destruct LREF as (PCEQ & _ & _). destruct LREF' as (PCEQ' & _ & _). + rewrite <- PCEQ. rewrite <- PCEQ'. rewrite <- FREF. rewrite <- FREF'. + assumption. + (* FIXME - that proof will have to be changed once hfinal_refines is more complex *) Qed. -*) Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 0e55dc6b..1e8adc71 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -173,8 +173,8 @@ Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) sfval_simu dm f st.(si_pc) st'.(si_pc) sv sv' -> - ssem_final pge ge sp st stk f rs0 m0 sv rs m t s -> - exists s', ssem_final tpge tge sp st' stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. + ssem_final pge ge sp st.(si_pc) stk f rs0 m0 sv rs m t s -> + exists s', ssem_final tpge tge sp st'.(si_pc) stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct: core. intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. -- cgit From 962a35719c04c5325f6034fa68743c92613c3fd2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 7 Jul 2020 18:26:33 +0200 Subject: begin computing OpWeights --- kvx/OpWeights.ml | 73 ++++++++++++++++++++++++++++++++++++++++++ kvx/lib/RTLpathScheduleraux.ml | 18 +++++++++++ 2 files changed, 91 insertions(+) create mode 100644 kvx/OpWeights.ml diff --git a/kvx/OpWeights.ml b/kvx/OpWeights.ml new file mode 100644 index 00000000..8396bde1 --- /dev/null +++ b/kvx/OpWeights.ml @@ -0,0 +1,73 @@ +open Op;; +open PostpassSchedulingOracle;; +let resource_bounds = PostpassSchedulingOracle.resource_bounds;; + + +let rec nlist_rec x l = function + | 0 -> l + | n when n > 0 -> nlist_rec x (x :: l) (n-1) + | _ -> failwith "nlist_rec";; +let nlist x n = nlist_rec x [] n;; + +let bogus_register = Machregs.R0;; +let bogus_inputs n = nlist bogus_register n;; + +let insns_of_op (op : operation) (nargs : int) = + match Asmblockgen.transl_op op + (bogus_inputs nargs) bogus_register [] with + | Errors.Error msg -> failwith "OpWeights.insns_of_op" + | Errors.OK insns -> insns;; + +let insn_of_op op nargs = + match insns_of_op op nargs with + | [] -> failwith "OpWeights.insn_of_op" + | h::_ -> h;; + +let insns_of_load trap chunk addressing (nargs : int) = + match Asmblockgen.transl_load trap chunk addressing + (bogus_inputs nargs) bogus_register [] with + | Errors.Error msg -> failwith "OpWeights.insns_of_load" + | Errors.OK insns -> insns;; + +let insn_of_load trap chunk addressing nargs = + match insns_of_load trap chunk addressing nargs with + | [] -> failwith "OpWeights.insn_of_load" + | h::_ -> h;; + +let insns_of_store chunk addressing (nargs : int) = + match Asmblockgen.transl_store chunk addressing + (bogus_inputs nargs) bogus_register [] with + | Errors.Error msg -> failwith "OpWeights.insns_of_store" + | Errors.OK insns -> insns;; + +let insn_of_store chunk addressing nargs = + match insns_of_store chunk addressing nargs with + | [] -> failwith "OpWeights.insn_of_store" + | h::_ -> h;; + +let latency_of_op (op : operation) (nargs : int) = + let insn = insn_of_op op nargs in + let record = basic_rec insn in + let latency = real_inst_to_latency record.inst in + latency;; + +let resources_of_op (op : operation) (nargs : int) = + let insn = insn_of_op op nargs in + let record = basic_rec insn in + rec_to_usage record;; + +let latency_of_load trap chunk (addr : addressing) (nargs : int) = 3;; +let latency_of_call _ _ = 6;; + +let resources_of_load trap chunk addressing nargs = + let insn = insn_of_load trap chunk addressing nargs in + let record = basic_rec insn in + rec_to_usage record;; + +let resources_of_store chunk addressing nargs = + let insn = insn_of_store chunk addressing nargs in + let record = basic_rec insn in + rec_to_usage record;; + +let resources_of_call _ _ = resource_bounds;; +let resources_of_builtin _ = resource_bounds;; diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 1a13da67..87d03fed 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -94,7 +94,25 @@ let get_superblocks code entry pm = end (* TODO David *) +let compute_latency (insn : RTL.instruction) = + match insn with + | Inop(successor) -> 0 + | Iop(op, args, dst, successor) -> OpWeights.latency_of_op op (List.length args) + | Iload(trap, chunk, addr, args, dst, successor) -> OpWeights.latency_of_load trap chunk addr (List.length args) + | Istore(chunk, addr, args, src, successor) -> 0 + | Icond(cond, args, ifso, ifnot, suggestion) -> 0 + + | Ijumptable(arg, targets) -> 0 + | Itailcall(sign, ros, args) -> 0 + | Icall(sign, ros, args, dst, successor) -> 0 + | Ibuiltin(ef, args, res, successor) -> 0 + | Ireturn(arg) -> 0;; + let schedule_superblock sb code = + let old_flag = !debug_flag in + debug_flag := true; + print_superblock sb code; + debug_flag := old_flag; (* stub2: reverse function *) (* let reversed = Array.of_list @@ List.rev @@ Array.to_list (sb.instructions) in -- cgit From 2396cc336c7ffce0075073591c3243ace2320d18 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 7 Jul 2020 20:46:40 +0200 Subject: begin prepass scheduling oracle --- kvx/OpWeights.ml | 16 ++ kvx/lib/PrepassSchedulingOracle.ml | 406 +++++++++++++++++++++++++++++++++++++ kvx/lib/RTLpathScheduleraux.ml | 18 +- 3 files changed, 426 insertions(+), 14 deletions(-) create mode 100644 kvx/lib/PrepassSchedulingOracle.ml diff --git a/kvx/OpWeights.ml b/kvx/OpWeights.ml index 8396bde1..4c3c40d0 100644 --- a/kvx/OpWeights.ml +++ b/kvx/OpWeights.ml @@ -23,6 +23,17 @@ let insn_of_op op nargs = | [] -> failwith "OpWeights.insn_of_op" | h::_ -> h;; +let insns_of_cond (cond : condition) (nargs : int) = + match Asmblockgen.transl_cond_op cond + Asmvliw.GPR0 (bogus_inputs nargs) [] with + | Errors.Error msg -> failwith "OpWeights.insns_of_cond" + | Errors.OK insns -> insns;; + +let insn_of_cond cond nargs = + match insns_of_cond cond nargs with + | [] -> failwith "OpWeights.insn_of_cond" + | h::_ -> h;; + let insns_of_load trap chunk addressing (nargs : int) = match Asmblockgen.transl_load trap chunk addressing (bogus_inputs nargs) bogus_register [] with @@ -56,6 +67,11 @@ let resources_of_op (op : operation) (nargs : int) = let record = basic_rec insn in rec_to_usage record;; +let resources_of_cond (cond : condition) (nargs : int) = + let insn = insn_of_cond cond nargs in + let record = basic_rec insn in + rec_to_usage record;; + let latency_of_load trap chunk (addr : addressing) (nargs : int) = 3;; let latency_of_call _ _ = 6;; diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml new file mode 100644 index 00000000..14239ed2 --- /dev/null +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -0,0 +1,406 @@ +open AST +open RTL +open Maps +open InstructionScheduler +open OpWeights +open Registers + +let use_alias_analysis () = false + +let length_of_chunk = function +| Mint8signed +| Mint8unsigned -> 1 +| Mint16signed +| Mint16unsigned -> 2 +| Mint32 +| Mfloat32 +| Many32 -> 4 +| Mint64 +| Mfloat64 +| Many64 -> 8;; + +let get_simple_dependencies (seqa : instruction array) = + 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 latency_constraints : latency_constraint list ref = 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_from; + instr_to = instr_to; + latency = 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 + begin + begin + (* Read after write *) + match !last_mem_write with + | None -> () + | Some j -> add_constraint j i 1 + end; + last_mem_reads := i :: !last_mem_reads + end + and add_output_mem i = + if not (use_alias_analysis ()) + then + begin + begin + (* Write after write *) + match !last_mem_write with + | None -> () + | Some j -> add_constraint j i 1 + end; + (* Write after read *) + List.iter (fun j -> add_constraint j i 0) !last_mem_reads; + last_mem_write := Some i; + last_mem_reads := [] + end + and add_input_reg i reg = + begin + (* Read after write *) + match PTree.get reg !last_reg_write with + | None -> () + | Some (j, latency) -> add_constraint j i latency + end; + last_reg_reads := PTree.set reg + (i :: get_last_reads reg) + !last_reg_reads + and add_output_reg i latency reg = + begin + (* Write after write *) + match PTree.get reg !last_reg_write with + | None -> () + | Some (j, _) -> add_constraint j i 1 + end; + begin + (* Write after read *) + List.iter (fun j -> add_constraint j i 0) (get_last_reads reg) + end; + 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 in + let rec add_builtin_res i (res : reg builtin_res) = + match res with + | BR r -> add_output_reg i 10 r + | BR_none -> () + | BR_splitlong (hi, lo) -> add_builtin_res i hi; + add_builtin_res i lo in + let rec add_builtin_arg i (ba : reg builtin_arg) = + match ba with + | BA r -> add_input_reg i r + | BA_int _ | BA_long _ | BA_float _ | BA_single _ -> () + | BA_loadstack(_,_) -> add_input_mem i + | BA_addrstack _ -> () + | BA_loadglobal(_, _, _) -> add_input_mem i + | BA_addrglobal _ -> () + | BA_splitlong(hi, lo) -> add_builtin_arg i hi; + add_builtin_arg i lo + | BA_addptr(a1, a2) -> add_builtin_arg i a1; + add_builtin_arg i a2 + in + Array.iteri + begin + fun i insn -> + match insn with + | Inop _ -> () + | Iop(op, inputs, output, _) -> + add_input_regs i inputs; + add_output_reg i (latency_of_op op (List.length inputs)) output + | Iload(trap, chunk, addressing, addr_regs, output, _) -> + add_input_mem i; + add_input_regs i addr_regs; + add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output + | Istore(chunk, addressing, addr_regs, input, _) -> + add_input_regs i addr_regs; + add_input_reg i input; + add_output_mem i + | Icall(signature, ef, inputs, output, _) -> + (match ef with + | Datatypes.Coq_inl r -> add_input_reg i r + | Datatypes.Coq_inr symbol -> () + ); + add_input_mem i; + add_input_regs i inputs; + add_output_reg i (latency_of_call signature ef) output; + add_output_mem i + | Itailcall(signature, ef, inputs) -> + (match ef with + | Datatypes.Coq_inl r -> add_input_reg i r + | Datatypes.Coq_inr symbol -> () + ); + add_input_mem i; + add_input_regs i inputs + | Ibuiltin(ef, builtin_inputs, builtin_output, _) -> + add_input_mem i; + List.iter (add_builtin_arg i) builtin_inputs; + add_builtin_res i builtin_output; + add_output_mem i + | Icond(cond, inputs, _, _, _) -> + add_input_regs i inputs + | Ijumptable(input, _) -> + add_input_reg i input + | Ireturn(Some input) -> + add_input_reg i input + | Ireturn(None) -> () + end seqa; + !latency_constraints;; + +let resources_of_instruction = function + | Inop _ -> Array.map (fun _ -> 0) resource_bounds + | Iop(op, inputs, output, _) -> resources_of_op op (List.length inputs) + | Iload(trap, chunk, addressing, addr_regs, output, _) -> + resources_of_load trap chunk addressing (List.length addr_regs) + | Istore(chunk, addressing, addr_regs, input, _) -> + resources_of_store chunk addressing (List.length addr_regs) + | Icall(signature, ef, inputs, output, _) -> + resources_of_call signature ef + | Ibuiltin(ef, builtin_inputs, builtin_output, _) -> + resources_of_builtin ef + | Icond(cond, args, _, _ , _) -> + resources_of_cond cond (List.length args) + | Itailcall _ | Ijumptable _ | Ireturn _ -> resource_bounds + +let print_sequence pp (seqa : instruction array) = + Array.iteri ( + fun i (insn : instruction) -> + PrintRTL.print_instruction pp (i, insn)) seqa;; + +type unique_id = int + +type 'a symbolic_term_node = + | STop of Op.operation * 'a list + | STinitial_reg of int + | STother of int;; + +type symbolic_term = { + hash_id : unique_id; + hash_ct : symbolic_term symbolic_term_node + };; + +let rec print_term channel term = + match term.hash_ct with + | STop(op, args) -> + PrintOp.print_operation print_term channel (op, args) + | STinitial_reg n -> Printf.fprintf channel "x%d" n + | STother n -> Printf.fprintf channel "y%d" n;; + +type symbolic_term_table = { + st_table : (unique_id symbolic_term_node, symbolic_term) Hashtbl.t; + mutable st_next_id : unique_id };; + +let hash_init () = { + st_table = Hashtbl.create 20; + st_next_id = 0 + };; + +let ground_to_id = function + | STop(op, l) -> STop(op, List.map (fun t -> t.hash_id) l) + | STinitial_reg r -> STinitial_reg r + | STother i -> STother i;; + +let hash_node (table : symbolic_term_table) (term : symbolic_term symbolic_term_node) : symbolic_term = + let grounded = ground_to_id term in + match Hashtbl.find_opt table.st_table grounded with + | Some x -> x + | None -> + let term' = { hash_id = table.st_next_id; + hash_ct = term } in + (if table.st_next_id = max_int then failwith "hash: max_int"); + table.st_next_id <- table.st_next_id + 1; + Hashtbl.add table.st_table grounded term'; + term';; + +type access = { + base : symbolic_term; + offset : int64; + length : int + };; + +let term_equal a b = (a.hash_id = b.hash_id);; + +let access_of_addressing get_reg chunk addressing args = + match addressing, args with + | (Op.Aindexed ofs), [reg] -> Some + { base = get_reg reg; + offset = Camlcoq.camlint64_of_ptrofs ofs; + length = length_of_chunk chunk + } + | _, _ -> None ;; +(* TODO: global *) + +let symbolic_execution (seqa : instruction array) = + let regs = ref PTree.empty + and table = hash_init() in + let assign reg term = regs := PTree.set reg term !regs + and hash term = hash_node table term in + let get_reg reg = + match PTree.get reg !regs with + | None -> hash (STinitial_reg (Camlcoq.P.to_int reg)) + | Some x -> x in + let targets = Array.make (Array.length seqa) None in + Array.iteri + begin + fun i insn -> + match insn with + | Iop(Op.Omove, [input], output, _) -> + assign output (get_reg input) + | Iop(op, inputs, output, _) -> + assign output (hash (STop(op, List.map get_reg inputs))) + + | Iload(trap, chunk, addressing, args, output, _) -> + let access = access_of_addressing get_reg chunk addressing args in + targets.(i) <- access; + assign output (hash (STother(i))) + + | Icall(_, _, _, output, _) + | Ibuiltin(_, _, BR output, _) -> + assign output (hash (STother(i))) + + | Istore(chunk, addressing, args, va, _) -> + let access = access_of_addressing get_reg chunk addressing args in + targets.(i) <- access + + | Inop _ -> () + | Ibuiltin(_, _, BR_none, _) -> () + | Ibuiltin(_, _, BR_splitlong _, _) -> failwith "BR_splitlong" + + | Itailcall (_, _, _) + |Icond (_, _, _, _, _) + |Ijumptable (_, _) + |Ireturn _ -> () + end seqa; + targets;; + +let print_access channel = function + | None -> Printf.fprintf channel "any" + | Some x -> Printf.fprintf channel "%a + %Ld" print_term x.base x.offset;; + +let print_targets channel seqa = + let targets = symbolic_execution seqa in + Array.iteri + (fun i insn -> + match insn with + | Iload _ -> Printf.fprintf channel "%d: load %a\n" + i print_access targets.(i) + | Istore _ -> Printf.fprintf channel "%d: store %a\n" + i print_access targets.(i) + | _ -> () + ) seqa;; + +let may_overlap a0 b0 = + match a0, b0 with + | (None, _) | (_ , None) -> true + | (Some a), (Some b) -> + if term_equal a.base b.base + then (max a.offset b.offset) < + (min (Int64.add (Int64.of_int a.length) a.offset) + (Int64.add (Int64.of_int b.length) b.offset)) + else match a.base.hash_ct, b.base.hash_ct with + | STop(Op.Oaddrsymbol(ida, ofsa),[]), + STop(Op.Oaddrsymbol(idb, ofsb),[]) -> + (ida=idb) && + let ao = Int64.add a.offset (Camlcoq.camlint64_of_ptrofs ofsa) + and bo = Int64.add b.offset (Camlcoq.camlint64_of_ptrofs ofsb) in + (max ao bo) < + (min (Int64.add (Int64.of_int a.length) ao) + (Int64.add (Int64.of_int b.length) bo)) + | STop(Op.Oaddrstack _, []), + STop(Op.Oaddrsymbol _, []) + | STop(Op.Oaddrsymbol _, []), + STop(Op.Oaddrstack _, []) -> false + | STop(Op.Oaddrstack(ofsa),[]), + STop(Op.Oaddrstack(ofsb),[]) -> + let ao = Int64.add a.offset (Camlcoq.camlint64_of_ptrofs ofsa) + and bo = Int64.add b.offset (Camlcoq.camlint64_of_ptrofs ofsb) in + (max ao bo) < + (min (Int64.add (Int64.of_int a.length) ao) + (Int64.add (Int64.of_int b.length) bo)) + | _ -> true;; + +(* +(* TODO suboptimal quadratic algorithm *) +let get_alias_dependencies seqa = + let targets = symbolic_execution seqa + and deps = ref [] in + let add_constraint instr_from instr_to latency = + deps := { instr_from = instr_from; + instr_to = instr_to; + latency = latency + }:: !deps in + for i=0 to (Array.length seqa)-1 + do + for j=0 to i-1 + do + match seqa.(j), seqa.(i) with + | (Istore _), ((Iload _) | (Istore _)) -> + if may_overlap targets.(j) targets.(i) + then add_constraint j i 1 + | (Iload _), (Istore _) -> + if may_overlap targets.(j) targets.(i) + then add_constraint j i 0 + | (Istore _ | Iload _), (Icall _ | Ibuiltin _) + | (Icall _ | Ibuiltin _), (Icall _ | Ibuiltin _ | Iload _ | Istore _) -> + add_constraint j i 1 + | (Inop _ | Iop _), _ + | _, (Inop _ | Iop _) + | (Iload _), (Iload _) -> () + done + done; + !deps;; + *) + +let define_problem seqa = + let simple_deps = get_simple_dependencies seqa in + { max_latency = -1; + resource_bounds = OpWeights.resource_bounds; + instruction_usages = Array.map resources_of_instruction seqa; + latency_constraints = + (* if (use_alias_analysis ()) + then (get_alias_dependencies seqa) @ simple_deps + else *) simple_deps };; + +let schedule_sequence (seqa : instruction array) = + try + if (Array.length seqa) <= 1 + then seqa + else + begin + let nr_instructions = Array.length seqa in + Printf.printf "prepass scheduling length = %d\n" (Array.length seqa); + let problem = define_problem seqa in + print_sequence stdout seqa; + print_problem stdout problem; + match reverse_list_scheduler problem + (* scheduler_by_name !Clflags.option_fprepass_sched problem *) with + | None -> Printf.printf "no solution in prepass scheduling\n"; + seqa + | 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; + let reordered = Array.init nr_instructions + (fun i -> seqa.(positions.(i))) in + reordered + end + with (Failure s) -> + Printf.printf "failure in prepass scheduling: %s\n" s; + seqa;; + diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 87d03fed..c304c6d3 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -94,24 +94,14 @@ let get_superblocks code entry pm = end (* TODO David *) -let compute_latency (insn : RTL.instruction) = - match insn with - | Inop(successor) -> 0 - | Iop(op, args, dst, successor) -> OpWeights.latency_of_op op (List.length args) - | Iload(trap, chunk, addr, args, dst, successor) -> OpWeights.latency_of_load trap chunk addr (List.length args) - | Istore(chunk, addr, args, src, successor) -> 0 - | Icond(cond, args, ifso, ifnot, suggestion) -> 0 - - | Ijumptable(arg, targets) -> 0 - | Itailcall(sign, ros, args) -> 0 - | Icall(sign, ros, args, dst, successor) -> 0 - | Ibuiltin(ef, args, res, successor) -> 0 - | Ireturn(arg) -> 0;; - let schedule_superblock sb code = let old_flag = !debug_flag in debug_flag := true; print_superblock sb code; + let _ = PrepassSchedulingOracle.schedule_sequence + (Array.map (fun i -> + match PTree.get i code with Some ii -> ii | None -> failwith "RTLpathScheduleraux.schedule_superblock") + sb.instructions) in debug_flag := old_flag; (* stub2: reverse function *) (* -- cgit From 988070c7bf7109aa1342a99e08b0cc1ddeea9ebb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Tue, 7 Jul 2020 10:41:45 +0200 Subject: Characterizing Op dependency on memory --- common/Memory.v | 12 ++++++++++++ kvx/Op.v | 39 +++++++++++++++++++++++++++++++++++++++ 2 files changed, 51 insertions(+) diff --git a/common/Memory.v b/common/Memory.v index cd8a2001..65f36966 100644 --- a/common/Memory.v +++ b/common/Memory.v @@ -1322,6 +1322,18 @@ Proof. eapply load_store_same. eassumption. Qed. + +Theorem storev_preserv_valid (b : block) (ofs: Z): valid_pointer m1 b ofs = valid_pointer m2 b ofs. +Proof. + unfold storev in STORE. + cut (valid_pointer m1 b ofs = true <-> valid_pointer m2 b ofs = true). + { destruct (valid_pointer _ _ _), (valid_pointer _ _ _); intuition congruence. } + destruct addr; try congruence. + rewrite! valid_pointer_valid_access. split. + - intros; eapply store_valid_access_1; eauto. + - intros; eapply store_valid_access_2; eauto. +Qed. + End STOREV. Lemma load_store_overlap: diff --git a/kvx/Op.v b/kvx/Op.v index 544bb081..b78b7b97 100644 --- a/kvx/Op.v +++ b/kvx/Op.v @@ -1238,6 +1238,45 @@ Proof. unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity. Qed. +Lemma Val_cmpu_bool_valid_pointer_eq m1 m2 c v1 v2: + (forall (b : block) (z : Z), Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) -> + Val.cmpu_bool( Mem.valid_pointer m1) c v1 v2 = Val.cmpu_bool (Mem.valid_pointer m2) c v1 v2. +Proof. + intros MEM; unfold Val.cmpu_bool; destruct v1; try congruence; + destruct v2; try congruence; + rewrite !MEM; auto. +Qed. + +Lemma Val_cmplu_bool_valid_pointer_eq m1 m2 c v1 v2: + (forall (b : block) (z : Z), Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) -> + Val.cmplu_bool( Mem.valid_pointer m1) c v1 v2 = Val.cmplu_bool (Mem.valid_pointer m2) c v1 v2. +Proof. + intros MEM; unfold Val.cmplu_bool; destruct v1; try congruence; + destruct v2; try congruence; + rewrite !MEM; auto. +Qed. + +Lemma op_valid_pointer_eq: + forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2, + (forall b z, Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) -> + eval_operation ge sp op args m1 = eval_operation ge sp op args m2. +Proof. + intros until m2. destruct op; simpl; try congruence. + - intros MEM; destruct cond; simpl; try congruence; + repeat (destruct args; simpl; try congruence); + erewrite Val_cmpu_bool_valid_pointer_eq || erewrite Val_cmplu_bool_valid_pointer_eq; eauto. + - intros MEM; destruct c0; simpl; try congruence; + repeat (destruct args; simpl; try congruence); + erewrite Val_cmpu_bool_valid_pointer_eq || erewrite Val_cmplu_bool_valid_pointer_eq; eauto. + - intros MEM; destruct c0; simpl; try congruence; + repeat (destruct args; simpl; try congruence); + erewrite Val_cmpu_bool_valid_pointer_eq || erewrite Val_cmplu_bool_valid_pointer_eq; eauto. + - intros MEM; destruct c0; simpl; try congruence; + repeat (destruct args; simpl; try congruence); + erewrite Val_cmpu_bool_valid_pointer_eq || erewrite Val_cmplu_bool_valid_pointer_eq; eauto. +Qed. + + (** Global variables mentioned in an operation or addressing mode *) Definition globals_addressing (addr: addressing) : list ident := -- cgit From aae7ef702536618ca0e59912987421f4b021b782 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Jul 2020 09:09:25 +0200 Subject: prepass reordering activated --- kvx/lib/PrepassSchedulingOracle.ml | 30 ++++++++++++++++++++++-------- kvx/lib/RTLpathScheduleraux.ml | 18 ++++++++++++++---- 2 files changed, 36 insertions(+), 12 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 14239ed2..eec2d9b5 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -24,6 +24,7 @@ let get_simple_dependencies (seqa : instruction array) = 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 latency_constraints : latency_constraint list ref = ref [] in let add_constraint instr_from instr_to latency = assert (instr_from <= instr_to); @@ -111,7 +112,14 @@ let get_simple_dependencies (seqa : instruction array) = | BA_splitlong(hi, lo) -> add_builtin_arg i hi; add_builtin_arg i lo | BA_addptr(a1, a2) -> add_builtin_arg i a1; - add_builtin_arg i a2 + add_builtin_arg i a2 in + let 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 in Array.iteri begin @@ -126,10 +134,12 @@ let get_simple_dependencies (seqa : instruction array) = add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output | Istore(chunk, addressing, addr_regs, input, _) -> + irreversible_action i; add_input_regs i addr_regs; add_input_reg i input; add_output_mem i | Icall(signature, ef, inputs, output, _) -> + irreversible_action i; (match ef with | Datatypes.Coq_inl r -> add_input_reg i r | Datatypes.Coq_inr symbol -> () @@ -139,6 +149,7 @@ let get_simple_dependencies (seqa : instruction array) = add_output_reg i (latency_of_call signature ef) output; add_output_mem i | Itailcall(signature, ef, inputs) -> + set_branch i; (match ef with | Datatypes.Coq_inl r -> add_input_reg i r | Datatypes.Coq_inr symbol -> () @@ -146,17 +157,22 @@ let get_simple_dependencies (seqa : instruction array) = add_input_mem i; add_input_regs i inputs | Ibuiltin(ef, builtin_inputs, builtin_output, _) -> + set_branch i; add_input_mem i; List.iter (add_builtin_arg i) builtin_inputs; add_builtin_res i builtin_output; add_output_mem i | Icond(cond, inputs, _, _, _) -> + set_branch i; add_input_regs i inputs | Ijumptable(input, _) -> + set_branch i; add_input_reg i input | Ireturn(Some input) -> + set_branch i; add_input_reg i input - | Ireturn(None) -> () + | Ireturn(None) -> + set_branch i end seqa; !latency_constraints;; @@ -377,7 +393,7 @@ let define_problem seqa = let schedule_sequence (seqa : instruction array) = try if (Array.length seqa) <= 1 - then seqa + then None else begin let nr_instructions = Array.length seqa in @@ -388,7 +404,7 @@ let schedule_sequence (seqa : instruction array) = match reverse_list_scheduler problem (* scheduler_by_name !Clflags.option_fprepass_sched problem *) with | None -> Printf.printf "no solution in prepass scheduling\n"; - seqa + None | Some solution -> let positions = Array.init nr_instructions (fun i -> i) in Array.sort (fun i j -> @@ -396,11 +412,9 @@ let schedule_sequence (seqa : instruction array) = if si < sj then -1 else if si > sj then 1 else i - j) positions; - let reordered = Array.init nr_instructions - (fun i -> seqa.(positions.(i))) in - reordered + Some positions end with (Failure s) -> Printf.printf "failure in prepass scheduling: %s\n" s; - seqa;; + None;; diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index c304c6d3..3008543c 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -97,12 +97,23 @@ end let schedule_superblock sb code = let old_flag = !debug_flag in debug_flag := true; + print_endline "ORIGINAL SUPERBLOCK"; print_superblock sb code; - let _ = PrepassSchedulingOracle.schedule_sequence + debug_flag := old_flag; + match PrepassSchedulingOracle.schedule_sequence (Array.map (fun i -> match PTree.get i code with Some ii -> ii | None -> failwith "RTLpathScheduleraux.schedule_superblock") - sb.instructions) in - debug_flag := old_flag; + sb.instructions) with + | None -> sb.instructions + | Some order -> + let ins' = Array.map (fun i -> sb.instructions.(i)) order in + Printf.printf "REORDERED SUPERBLOCK %d\n" (Array.length ins'); + debug_flag := true; + print_instructions (Array.to_list ins') code; + debug_flag := old_flag; + (*sb.instructions; *) + ins';; + (* stub2: reverse function *) (* let reversed = Array.of_list @@ List.rev @@ Array.to_list (sb.instructions) in @@ -114,7 +125,6 @@ let schedule_superblock sb code = reversed end *) (* stub: identity function *) - sb.instructions let change_successors i = function | [] -> ( -- cgit From d37a9db8edb9c38c79940d9a3d647430a4c4d3e5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Jul 2020 09:31:10 +0200 Subject: use a command-line option --- driver/Clflags.ml | 1 + driver/Driver.ml | 4 +++- kvx/lib/RTLpathScheduleraux.ml | 39 +++++++++++++++++++++------------------ 3 files changed, 25 insertions(+), 19 deletions(-) diff --git a/driver/Clflags.ml b/driver/Clflags.ml index eb21b3f8..fc5d9e45 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -36,6 +36,7 @@ let option_fredundancy = ref true let option_fduplicate = ref (-1) let option_finvertcond = ref true let option_ftracelinearize = ref false +let option_fprepass = ref false let option_fpostpass = ref true let option_fpostpass_sched = ref "list" let option_fifconversion = ref true diff --git a/driver/Driver.ml b/driver/Driver.ml index 90afb812..6accc22b 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -204,7 +204,8 @@ Processing options: -fcse3-glb Refine CSE3 information using greatest lower bounds [on] -fmove-loop-invariants Perform loop-invariant code motion [off] -fredundancy Perform redundancy elimination [on] - -fpostpass Perform postpass scheduling (only for K1 architecture) [on] + -fprepass Perform prepass scheduling (only for K1 architecture) [off] + -fpostpass Perform postpass scheduling (only for K1 architecture) [on] -fpostpass= Perform postpass scheduling with the specified optimization [list] (=list: list scheduling, =ilp: ILP, =greedy: just packing bundles) -fduplicate Perform tail duplication to form superblocks on predicted traces @@ -419,6 +420,7 @@ let cmdline_actions = @ f_opt "cse3-glb" option_fcse3_glb @ f_opt "move-loop-invariants" option_fmove_loop_invariants @ f_opt "redundancy" option_fredundancy + @ f_opt "prepass" option_fprepass @ f_opt "postpass" option_fpostpass @ [ Exact "-fduplicate", Integer (fun n -> option_fduplicate := n) ] @ f_opt "invertcond" option_finvertcond diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 3008543c..32fb2c2a 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -95,24 +95,27 @@ end (* TODO David *) let schedule_superblock sb code = - let old_flag = !debug_flag in - debug_flag := true; - print_endline "ORIGINAL SUPERBLOCK"; - print_superblock sb code; - debug_flag := old_flag; - match PrepassSchedulingOracle.schedule_sequence - (Array.map (fun i -> - match PTree.get i code with Some ii -> ii | None -> failwith "RTLpathScheduleraux.schedule_superblock") - sb.instructions) with - | None -> sb.instructions - | Some order -> - let ins' = Array.map (fun i -> sb.instructions.(i)) order in - Printf.printf "REORDERED SUPERBLOCK %d\n" (Array.length ins'); - debug_flag := true; - print_instructions (Array.to_list ins') code; - debug_flag := old_flag; - (*sb.instructions; *) - ins';; + if not !Clflags.option_fprepass + then sb.instructions + else + let old_flag = !debug_flag in + debug_flag := true; + print_endline "ORIGINAL SUPERBLOCK"; + print_superblock sb code; + debug_flag := old_flag; + match PrepassSchedulingOracle.schedule_sequence + (Array.map (fun i -> + match PTree.get i code with Some ii -> ii | None -> failwith "RTLpathScheduleraux.schedule_superblock") + sb.instructions) with + | None -> sb.instructions + | Some order -> + let ins' = Array.map (fun i -> sb.instructions.(i)) order in + Printf.printf "REORDERED SUPERBLOCK %d\n" (Array.length ins'); + debug_flag := true; + print_instructions (Array.to_list ins') code; + debug_flag := old_flag; + (*sb.instructions; *) + ins';; (* stub2: reverse function *) (* -- cgit From 9dbceafe5c4d12fa6aead15b5cda565d60a509fb Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 8 Jul 2020 13:44:13 +0200 Subject: weaker sfval_simu --- kvx/lib/RTLpathSE_impl.v | 13 ++- kvx/lib/RTLpathSE_theory.v | 250 ++++++++++++++++++++++------------------ kvx/lib/RTLpathSchedulerproof.v | 54 +++++---- 3 files changed, 183 insertions(+), 134 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 2c4ae088..416fe422 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -560,6 +560,8 @@ Proof. apply hsexec_final_correct; auto. Qed. +(** TODO: all this needs to be adapted + Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. Definition hsilocal_simu dm f (hst1 hst2: hsistate_local): Prop := forall st1 st2, @@ -800,14 +802,21 @@ Proof. + subst. eapply simu_check_single_correct; eauto. + eapply IHlm; assumption. Qed. +*) -Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := simu_check_rec dm f tf (PTree.elements dm). +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := + (* TODO: adapt this + simu_check_rec dm f tf (PTree.elements dm) + *) + OK tt. Lemma simu_check_correct dm f tf: simu_check dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2. +Admitted. +(* TODO: adapt this. Proof. unfold simu_check. intros. eapply PTree.elements_correct in H0. eapply simu_check_rec_correct; eassumption. -Qed. \ No newline at end of file +Qed.*) \ No newline at end of file diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 7e83f772..2be567d6 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1634,9 +1634,7 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 /\ srce!(is2.(ipc)) = Some is1.(ipc). -(* Definition is_injective (dm: PTree.t node): Prop := - forall x y, dm!x = dm!y -> x = y. *) - +(* TODO: code à adapter pour plus tard Fixpoint sval_simub (sv1 sv2: sval) := match sv1 with | Sinput r => @@ -1749,6 +1747,7 @@ Proof. apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. congruence. Qed. +*) Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := match dm ! tn with @@ -1764,21 +1763,135 @@ Proof. unfold revmap_check_single. explore; try discriminate. congruence. Qed. +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_ge2 s = Genv.find_symbol the_ge1 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 (f: RTLpath.function) (srce: PTree.t node) (st1 st2: sistate): Prop := - forall sp ge1 ge2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - liveness_ok_function f -> -(* is_injective srce -> *) - forall rs m is1, ssem_internal ge1 sp st1 rs m is1 -> - exists is2, ssem_internal ge2 sp st2 rs m is2 /\ istate_simu f srce is1 is2. +Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) (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 is1 is2. + +Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) (ctx: simu_proof_context f): Prop := + forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) sl1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> + exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) sl2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) + /\ istate_simu f dm is1 is2. + +Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f):= + forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> + exists is2, + ssem_exit (the_ge1 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) + /\ istate_simu f dm is1 is2. + +Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f) := + siexit_simu_ssem dm f se1 se2 ctx /\ (si_cond se1) = (si_cond se2). (* NB on si_cond: a pure syntactic definition should suffice here *) + +Definition opt_simu {A} (o1: option A) (o2: option A) := + o1 <> None -> o2 = o1. + +Lemma opt_simu_Some {A} (a:A) (o1: option A) (o2: option A): + opt_simu o1 o2 -> o1 = Some a -> o2 = Some a. +Proof. + unfold opt_simu. intros H1 H2; subst; rewrite H1; congruence. +Qed. + +Inductive svident_simu (f: RTLpath.function) (ctx: simu_proof_context f): (sval + ident) -> (sval + ident) -> Prop := + | Sleft_simu sv1 sv2: + opt_simu (seval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) + (seval_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 -> + opt_simu (seval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) + (seval_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 -> + opt_simu (seval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx)) + (seval_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 -> + lbs1 = lbs2 -> (* FIXME: TOO STRONG *) + 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 -> + opt_simu (seval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) + (seval_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 os: + sfval_simu dm f opc1 opc2 ctx (Sreturn os) (Sreturn os). + +Definition sstate_simu dm f (s1 s2: sstate) (ctx: simu_proof_context f): Prop := + sistate_simu dm f s1.(internal) s2.(internal) ctx + /\ 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 st2, sexec f2 pc2 = Some st2 /\ forall ctx, sstate_simu dm f1 st1 st2 ctx. + +(* TODO: all this needs to be adapted .... -Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local): Prop := - forall sp ge1 ge2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - liveness_ok_function f -> - forall rs m is1, ssem_local ge1 sp sl1 rs m (irs is1) (imem is1) -> - exists is2, ssem_local ge2 sp sl2 rs m (irs is2) (imem is2) /\ istate_simu f dm is1 is2. (* TODO - generalize it with a list of registers to test *) Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (* (regs: list reg) *) (sl1 sl2: sistate_local) := @@ -1797,32 +1910,6 @@ Lemma silocal_simub_correct dm f sl1 sl2: Proof. Admitted. -Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := - forall sp ge1 ge2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - liveness_ok_function f -> - forall rs m is1, ssem_exit ge1 sp se1 rs m (irs is1) (imem is1) (ipc is1) -> - exists is2, - ssem_exit ge2 sp se2 rs m (irs is2) (imem is2) (ipc is2) - /\ istate_simu f dm is1 is2. - -Definition siexit_simu_cond (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := - forall sp ge1 ge2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - forall rs m, - seval_condition ge1 sp (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) rs m - = seval_condition ge2 sp (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) rs m. - -Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := - siexit_simu_ssem dm f se1 se2 /\ siexit_simu_cond dm f se1 se2. - -Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := - if (eq_condition (si_cond se1) (si_cond se2)) then - do _ <- list_sval_simub (si_scondargs se1) (si_scondargs se2); - do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); (* TODO - generalize it with a list of registers to test *) - revmap_check_single dm (si_ifso se1) (si_ifso se2) - else Error (msg "siexit_simub: conditions do not match") -. Lemma siexit_simub_correct dm f se1 se2: siexit_simub dm f se1 se2 = OK tt -> @@ -1888,6 +1975,15 @@ Proof. eapply all_fallthrough_upto_exit_cons; eauto. Qed. + +Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := + if (eq_condition (si_cond se1) (si_cond se2)) then + do _ <- list_sval_simub (si_scondargs se1) (si_scondargs se2); + do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); (* TODO - generalize it with a list of registers to test *) + revmap_check_single dm (si_ifso se1) (si_ifso se2) + else Error (msg "siexit_simub: conditions do not match") +. + Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := match lse1 with | nil => @@ -2000,46 +2096,6 @@ Proof. + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). Qed. -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. - Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := match ln with | nil => @@ -2065,25 +2121,6 @@ Proof. simpl. rewrite EQ. rewrite EQ0. reflexivity. Qed. -(* NOTE: a syntactic definition on [sfval] in order to abstract the [match_states] *) -Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node): sfval -> sfval -> Prop := - | Snone_simu: - dm!opc2 = Some opc1 -> - sfval_simu dm f opc1 opc2 Snone Snone - | Scall_simu sig svos lsv res pc1 pc2: - dm!pc2 = Some pc1 -> - sfval_simu dm f opc1 opc2 (Scall sig svos lsv res pc1) (Scall sig svos lsv res pc2) - | Stailcall_simu sig svos lsv: - sfval_simu dm f opc1 opc2 (Stailcall sig svos lsv) (Stailcall sig svos lsv) - | Sbuiltin_simu ef lbs br pc1 pc2: - dm!pc2 = Some pc1 -> - sfval_simu dm f opc1 opc2 (Sbuiltin ef lbs br pc1) (Sbuiltin ef lbs br pc2) - | Sjumptable_simu sv lpc1 lpc2: - ptree_get_list dm lpc2 = Some lpc1 -> - sfval_simu dm f opc1 opc2 (Sjumptable sv lpc1) (Sjumptable sv lpc2) - | Sreturn_simu os: - sfval_simu dm f opc1 opc2 (Sreturn os) (Sreturn os). - Definition svos_simub (svos1 svos2: sval + ident) := match svos1 with | inl sv1 => @@ -2311,10 +2348,6 @@ Proof. + constructor; auto. Qed. -Definition sstate_simu dm f (s1 s2: sstate): Prop := - sistate_simu f dm s1.(internal) s2.(internal) - /\ sfval_simu dm f s1.(si_pc) s2.(si_pc) s1.(final) s2.(final). - (* Computable check of sstate_simu *) Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := do u1 <- sistate_simub dm f (internal st1) (internal st2); @@ -2323,14 +2356,11 @@ Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstat Lemma sstate_simub_correct dm f st1 st2: sstate_simub dm f st1 st2 = OK tt -> - sstate_simu dm f st1 st2. + forall ctx, sstate_simu dm f st1 st2 ctx. Proof. unfold sstate_simub. intros. explore. constructor. apply sistate_simub_correct; assumption. apply sfval_simub_correct; assumption. Qed. - -Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := - forall st1, sexec f1 pc1 = Some st1 -> - exists st2, sexec f2 pc2 = Some st2 /\ sstate_simu dm f1 st1 st2. +*) diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 1e8adc71..892fe2fa 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -136,48 +136,58 @@ Proof. erewrite RTLpath.symbols_preserved; eauto. Qed. -Lemma s_find_function_preserved sp svos rs0 m0 fd: - sfind_function pge ge sp svos rs0 m0 = Some fd -> - exists fd', sfind_function tpge tge sp svos rs0 m0 = Some fd' +Program Definition mkctx sp rs0 m0 {f1: RTLpath.function} (hyp: liveness_ok_function f1) + : simu_proof_context f1 + := {| the_ge1:= ge; the_ge2 := tge; the_sp:=sp; the_rs0:=rs0; the_m0:=m0 |}. +Obligation 2. + eapply symbols_preserved_RTL. +Qed. + +Lemma s_find_function_preserved f sp svos1 svos2 rs0 m0 fd + (LIVE: liveness_ok_function f): + (svident_simu f (mkctx sp rs0 m0 LIVE) svos1 svos2) -> + sfind_function pge ge sp svos1 rs0 m0 = Some fd -> + exists fd', sfind_function tpge tge sp svos2 rs0 m0 = Some fd' /\ transf_fundef fd = OK fd' /\ liveness_ok_fundef fd. Proof. Local Hint Resolve symbols_preserved_RTL: core. - unfold sfind_function. destruct svos; simpl. - + rewrite (seval_preserved ge tge); auto. + unfold sfind_function. intros [sv1 sv2 SIMU|]; simpl in *. + + rewrite !(seval_preserved ge tge) in *; eauto. destruct (seval_sval _ _ _ _); try congruence. + erewrite SIMU; try congruence. clear SIMU. intros; exploit functions_preserved; eauto. - intros (fd' & cunit & X). eexists. intuition eauto. + intros (fd' & cunit & (X1 & X2 & X3)). eexists. + repeat split; eauto. eapply find_funct_liveness_ok; eauto. (* intros. eapply all_fundef_liveness_ok; eauto. *) - + rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. + + subst. rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence. intros; exploit function_ptr_preserved; eauto. intros (fd' & X). eexists. intuition eauto. (* intros. eapply all_fundef_liveness_ok; eauto. *) Qed. -Lemma sistate_simu f dupmap sp st st' rs m is: +Lemma sistate_simu f dupmap sp st st' rs m is + (LIVE: liveness_ok_function f): ssem_internal ge sp st rs m is -> - liveness_ok_function f -> -(* is_injective dupmap -> *) - sistate_simu f dupmap st st' -> + sistate_simu dupmap f st st' (mkctx sp rs m LIVE)-> exists is', ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap is is'. Proof. - intros SEM LIVE (* INJ *) X; eapply X; eauto. + intros SEM X; eapply X; eauto. Qed. -Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s: +Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s + (LIVE: liveness_ok_function f): match_function dm f f' -> - liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> (* s_istate_simu f dupmap st st' -> *) - sfval_simu dm f st.(si_pc) st'.(si_pc) sv sv' -> + sfval_simu dm f st.(si_pc) st'.(si_pc) (mkctx sp rs0 m0 LIVE) sv sv' -> ssem_final pge ge sp st.(si_pc) stk f rs0 m0 sv rs m t s -> exists s', ssem_final tpge tge sp st'.(si_pc) stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'. Proof. Local Hint Resolve transf_fundef_correct: core. - intros FUN LIVE STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl. + intros FUN STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl in *. - (* Snone *) exploit initialize_path. { eapply dupmap_path_entry1; eauto. } intros (path & PATH). @@ -190,7 +200,7 @@ Proof. exploit initialize_path. { eapply dupmap_path_entry1; eauto. } intros (path & PATH). eexists; split; econstructor; eauto. - + erewrite <- (list_sval_eval_preserved ge tge); auto. + + eapply opt_simu_Some; eauto. + simpl. repeat (econstructor; eauto). - (* Stailcall *) exploit s_find_function_preserved; eauto. @@ -198,7 +208,7 @@ Proof. erewrite <- function_sig_preserved; eauto. eexists; split; econstructor; eauto. + erewrite <- preserv_fnstacksize; eauto. - + erewrite <- (list_sval_eval_preserved ge tge); auto. + + eapply opt_simu_Some; eauto. - (* Sbuiltin *) pose senv_preserved_RTL as SRTL. exploit initialize_path. { eapply dupmap_path_entry1; eauto. } @@ -212,7 +222,7 @@ Proof. exploit initialize_path. { eapply dupmap_path_entry1; eauto. } intros (path & PATH). eexists; split; econstructor; eauto. - + erewrite <- (seval_preserved ge tge); eauto. + + eapply opt_simu_Some; eauto. + eapply eqlive_reg_refl. - (* Sreturn *) eexists; split; econstructor; eauto. @@ -227,11 +237,11 @@ Theorem ssem_sstate_simu dm f f' stk stk' sp st st' rs m t s: liveness_ok_function f -> list_forall2 match_stackframes stk stk' -> ssem pge ge sp st stk f rs m t s -> - sstate_simu dm f st st' -> + (forall ctx: simu_proof_context f, sstate_simu dm f st st' ctx) -> exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'. Proof. - intros MFUNC LIVE STACKS SEM (SIMU1 & SIMU2). -(* assert (INJ: is_injective dm) by (inv MFUNC; assumption). *) + intros MFUNC LIVE STACKS SEM SIMU. + destruct (SIMU (mkctx sp rs m LIVE)) as (SIMU1 & SIMU2); clear SIMU. destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl. - (* sem_early *) exploit sistate_simu; eauto. -- cgit From 5b1071d9f46bbb6083b1e6dd6f7975aaa64d0a9a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Jul 2020 15:26:11 +0200 Subject: making progress on prepass --- kvx/lib/PrepassSchedulingOracle.ml | 15 ++++++++++----- kvx/lib/RTLpathScheduleraux.ml | 19 ++++++++++++++++--- test/monniaux/rules.mk | 2 +- 3 files changed, 27 insertions(+), 9 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index eec2d9b5..5a48924e 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -19,7 +19,7 @@ let length_of_chunk = function | Mfloat64 | Many64 -> 8;; -let get_simple_dependencies (seqa : instruction array) = +let get_simple_dependencies (seqa : (instruction*Regset.t) array) = 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 [] @@ -123,7 +123,12 @@ let get_simple_dependencies (seqa : instruction array) = in Array.iteri begin - fun i insn -> + fun i (insn, other_uses) -> + List.iter (fun use -> + print_endline "add other use"; + add_input_reg i use) + (Regset.elements other_uses); + match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> @@ -384,13 +389,13 @@ let define_problem seqa = let simple_deps = get_simple_dependencies seqa in { max_latency = -1; resource_bounds = OpWeights.resource_bounds; - instruction_usages = Array.map resources_of_instruction seqa; + instruction_usages = Array.map resources_of_instruction (Array.map fst seqa); latency_constraints = (* if (use_alias_analysis ()) then (get_alias_dependencies seqa) @ simple_deps else *) simple_deps };; -let schedule_sequence (seqa : instruction array) = +let schedule_sequence (seqa : (instruction*Regset.t) array) = try if (Array.length seqa) <= 1 then None @@ -399,7 +404,7 @@ let schedule_sequence (seqa : instruction array) = let nr_instructions = Array.length seqa in Printf.printf "prepass scheduling length = %d\n" (Array.length seqa); let problem = define_problem seqa in - print_sequence stdout seqa; + print_sequence stdout (Array.map fst seqa); print_problem stdout problem; match reverse_list_scheduler problem (* scheduler_by_name !Clflags.option_fprepass_sched problem *) with diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 32fb2c2a..97de1a12 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -103,13 +103,26 @@ let schedule_superblock sb code = print_endline "ORIGINAL SUPERBLOCK"; print_superblock sb code; debug_flag := old_flag; + let nr_instr = Array.length sb.instructions in + let trailer_length = + match PTree.get (sb.instructions.(nr_instr-1)) code with + | Some (Ireturn _ | Itailcall _ | Ijumptable _) -> 1 + | _ -> 0 in match PrepassSchedulingOracle.schedule_sequence (Array.map (fun i -> - match PTree.get i code with Some ii -> ii | None -> failwith "RTLpathScheduleraux.schedule_superblock") - sb.instructions) with + (match PTree.get i code with + | Some ii -> ii + | None -> failwith "RTLpathScheduleraux.schedule_superblock"), + (match PTree.get i sb.liveins with + | Some s -> s + | None -> Regset.empty)) + (Array.sub sb.instructions 0 (nr_instr-trailer_length))) with | None -> sb.instructions | Some order -> - let ins' = Array.map (fun i -> sb.instructions.(i)) order in + let ins' = + Array.append + (Array.map (fun i -> sb.instructions.(i)) order) + (Array.sub sb.instructions (nr_instr-trailer_length) trailer_length) in Printf.printf "REORDERED SUPERBLOCK %d\n" (Array.length ins'); debug_flag := true; print_instructions (Array.to_list ins') code; diff --git a/test/monniaux/rules.mk b/test/monniaux/rules.mk index c0594ef9..95d57f80 100644 --- a/test/monniaux/rules.mk +++ b/test/monniaux/rules.mk @@ -21,7 +21,7 @@ MEASURES?=time ALL_CFLAGS+=-Wall -D__KVX_COS__ -DMAX_MEASURES=$(MAX_MEASURES) #ALL_CFLAGS+=-g ALL_GCCFLAGS+=$(ALL_CFLAGS) -std=c99 -Wextra -Werror=implicit -ALL_CCOMPFLAGS+=$(ALL_CFLAGS) +ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fprepass # The compilers KVX_CC?=kvx-cos-gcc -- cgit From b67c28dd4dc6df5a5d3827f3a5827e437f3fcce4 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Jul 2020 15:30:09 +0200 Subject: Output regs in RTLpath --- kvx/lib/RTLpath.v | 7 +++++-- kvx/lib/RTLpathLivegenaux.ml | 30 ++++++++++++++++++++++++++---- 2 files changed, 31 insertions(+), 6 deletions(-) diff --git a/kvx/lib/RTLpath.v b/kvx/lib/RTLpath.v index 64f3917e..82991a4d 100644 --- a/kvx/lib/RTLpath.v +++ b/kvx/lib/RTLpath.v @@ -82,8 +82,11 @@ Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, *) Record path_info := { - psize: nat; (* number minus 1 of instructions in the path *) - input_regs: Regset.t + psize: nat; (* number minus 1 of instructions in the path *) + input_regs: Regset.t; + (** Registers that are used (as input_regs) by the "fallthrough successors" of the path *) + (** This field is not used by the verificator, but is helpful for the superblock scheduler *) + output_regs: Regset.t }. Definition path_map: Type := PTree.t path_info. diff --git a/kvx/lib/RTLpathLivegenaux.ml b/kvx/lib/RTLpathLivegenaux.ml index 09684f22..6875f155 100644 --- a/kvx/lib/RTLpathLivegenaux.ml +++ b/kvx/lib/RTLpathLivegenaux.ml @@ -83,7 +83,7 @@ let get_join_points code entry = | pc :: l -> traverse pc; traverse_succs l in traverse entry; !reached_twice -(* Does not set the input_regs field *) +(* 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 let path_map = ref PTree.empty in @@ -104,7 +104,9 @@ let get_path_map code entry join_points = path_successors := !path_successors @ non_predicted_successors inst; dig_path_rec n' end - | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); input_regs = Regset.empty }, !path_successors @ successors_inst inst) + | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize); + input_regs = Regset.empty; output_regs = Regset.empty }, + !path_successors @ successors_inst inst) end else None in match dig_path_rec e with @@ -235,14 +237,32 @@ let analyze f = end *) +let rec traverse code n size = + let inst = get_some @@ PTree.get n code in + if (size == 0) then inst + else + let n' = get_some @@ predicted_successor inst in + traverse code n' (size-1) + +let get_outputs liveness code n pi = + let last_instruction = traverse code n (Camlcoq.Nat.to_int pi.psize) in + let path_last_successors = successors_inst last_instruction in + let list_input_regs = List.map ( + fun n -> get_some @@ PTree.get n liveness + ) path_last_successors in + List.fold_left Regset.union Regset.empty list_input_regs + let set_pathmap_liveness f pm = let liveness = analyze f in let new_pm = ref PTree.empty in + let code = f.fn_code in begin dprintf "Liveness: "; print_ptree_regset liveness; dprintf "\n"; List.iter (fun (n, pi) -> - let rs = get_some @@ PTree.get n liveness in - new_pm := PTree.set n {psize=pi.psize; input_regs=rs} !new_pm + let inputs = get_some @@ PTree.get n liveness in + let outputs = get_outputs liveness code n pi in + new_pm := PTree.set n + {psize=pi.psize; input_regs=inputs; output_regs=outputs} !new_pm ) (PTree.elements pm); !new_pm end @@ -259,6 +279,8 @@ let print_path_info pi = begin dprintf "(psize=%d; " (Camlcoq.Nat.to_int pi.psize); dprintf "input_regs="; print_regset pi.input_regs; + dprintf "output_regs="; + print_regset pi.output_regs; dprintf ")" end -- cgit From 94f9a9c36db60971ba347779283895c96f05cb78 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Jul 2020 15:48:02 +0200 Subject: print_path_info fix --- kvx/lib/RTLpathLivegenaux.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathLivegenaux.ml b/kvx/lib/RTLpathLivegenaux.ml index 6875f155..dd971db8 100644 --- a/kvx/lib/RTLpathLivegenaux.ml +++ b/kvx/lib/RTLpathLivegenaux.ml @@ -279,7 +279,7 @@ let print_path_info pi = begin dprintf "(psize=%d; " (Camlcoq.Nat.to_int pi.psize); dprintf "input_regs="; print_regset pi.input_regs; - dprintf "output_regs="; + dprintf "; output_regs="; print_regset pi.output_regs; dprintf ")" end -- cgit From 823d1bd6e3db8e771e3f2aeacfdf86c01648c1a9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Jul 2020 15:58:16 +0200 Subject: progress on prepass scheduling --- kvx/lib/RTLpathScheduleraux.ml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 97de1a12..ad6f28b4 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -127,6 +127,8 @@ let schedule_superblock sb code = debug_flag := true; print_instructions (Array.to_list ins') code; debug_flag := old_flag; + flush stdout; + assert ((Array.length sb.instructions) = (Array.length ins')); (*sb.instructions; *) ins';; -- cgit From 3c1050e1a2704383a49eb39b8aee08e5dcdd983e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Jul 2020 16:06:00 +0200 Subject: Output regs in superblocks --- kvx/lib/RTLpathScheduleraux.ml | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 97de1a12..3bbef9ec 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -10,6 +10,8 @@ type superblock = { (* each predicted Pcb has its attached liveins *) (* This is indexed by the pc value *) liveins: Regset.t PTree.t; + (* Union of the input_regs of the last successors *) + output_regs: Regset.t } let print_instructions insts code = @@ -23,9 +25,11 @@ let print_instructions insts code = let print_superblock sb code = let insts = sb.instructions in let li = sb.liveins in + let outs = sb.output_regs in begin dprintf "{ instructions = "; print_instructions (Array.to_list insts) code; dprintf "\n"; - dprintf " liveins = "; print_ptree_regset li; dprintf "}" + dprintf " liveins = "; print_ptree_regset li; dprintf "\n"; + dprintf " output_regs = "; print_regset outs; dprintf "}" end let print_superblocks lsb code = @@ -81,9 +85,10 @@ let get_superblocks code entry pm = in if (get_some @@ PTree.get pc !visited) then [] else begin visited := PTree.set pc true !visited; - let n = (get_some @@ PTree.get pc pm).psize in - let (insts, nexts) = follow pc (Camlcoq.Nat.to_int n) in - let superblock = { instructions = Array.of_list insts; liveins = !liveins } in + let pi = get_some @@ PTree.get pc pm in + let (insts, nexts) = follow pc (Camlcoq.Nat.to_int pi.psize) in + let superblock = { instructions = Array.of_list insts; liveins = !liveins; + output_regs = pi.output_regs } in superblock :: (List.concat @@ List.map get_superblocks_rec nexts) end in let lsb = get_superblocks_rec entry in begin -- cgit From c0c9a5981046ad417e70f46370d3ee44f87b852d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Jul 2020 16:38:16 +0200 Subject: Typing information --- kvx/lib/RTLpathScheduleraux.ml | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index f332507c..bd6784d0 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -11,7 +11,8 @@ type superblock = { (* This is indexed by the pc value *) liveins: Regset.t PTree.t; (* Union of the input_regs of the last successors *) - output_regs: Regset.t + output_regs: Regset.t; + typing: RTLtyping.regenv } let print_instructions insts code = @@ -62,7 +63,7 @@ let print_arrayp arr = begin dprintf "|]" end -let get_superblocks code entry pm = +let get_superblocks code entry pm typing = let visited = ref (PTree.map (fun n i -> false) code) in let rec get_superblocks_rec pc = let liveins = ref (PTree.empty) in @@ -88,7 +89,7 @@ let get_superblocks code entry pm = let pi = get_some @@ PTree.get pc pm in let (insts, nexts) = follow pc (Camlcoq.Nat.to_int pi.psize) in let superblock = { instructions = Array.of_list insts; liveins = !liveins; - output_regs = pi.output_regs } in + output_regs = pi.output_regs; typing = typing } in superblock :: (List.concat @@ List.map get_superblocks_rec nexts) end in let lsb = get_superblocks_rec entry in begin @@ -227,11 +228,14 @@ let rec do_schedule code = function do_schedule new_code lsb end +let get_ok r = match r with Errors.OK x -> x | _ -> failwith "Did not get OK" + let scheduler f = let code = f.fn_RTL.fn_code in let id_ptree = PTree.map (fun n i -> n) (f.fn_path) in let entry = f.fn_RTL.fn_entrypoint in let pm = f.fn_path in - let lsb = get_superblocks code entry pm in + let typing = get_ok @@ RTLtyping.type_function f.fn_RTL in + let lsb = get_superblocks code entry pm typing in let tc = do_schedule code lsb in (((tc, entry), pm), id_ptree) -- cgit From 26525684b6347ce71aeb5494415d99409a1211c5 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Jul 2020 17:21:25 +0200 Subject: More explicit failwith messages for change_predicted_successor --- kvx/lib/RTLpathScheduleraux.ml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index bd6784d0..5e9111d2 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -179,8 +179,8 @@ let change_successors i = function | _ -> failwith "Wrong instruction (4)") let change_predicted_successor i s = match i with - | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5)" - | Ijumptable _ -> failwith "Wrong instruction (6) (shouldn't be predicted successor for jumptable)" + | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5) - Tailcalls and returns should not be moved in the middle of a superblock" + | Ijumptable _ -> failwith "Wrong instruction (6) - Jumptables should not be moved in the middle of a superblock" | Inop n -> Inop s | Iop (a,b,c,n) -> Iop (a,b,c,s) | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) @@ -191,7 +191,7 @@ let change_predicted_successor i s = match i with match p with | Some true -> Icond (a,b,s,n2,p) | Some false -> Icond (a,b,n1,s,p) - | None -> failwith "Predicted a successor for an Icond with p=None" + | None -> failwith "Predicted a successor for an Icond with p=None - unpredicted CB should not be moved in the middle of the superblock" ) let apply_schedule code sb new_order = -- cgit From 51b693a1c29819f3ecb463c37853de5216b625ad Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 8 Jul 2020 19:59:04 +0200 Subject: adapting Cyril's proofs in RTLpathSE_impl for the new definitions of simu. --- kvx/lib/RTLpathSE_impl.v | 544 +++++++++++++++++++++++++++++++--- kvx/lib/RTLpathSE_theory.v | 636 ++-------------------------------------- kvx/lib/RTLpathSchedulerproof.v | 2 +- 3 files changed, 525 insertions(+), 657 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 416fe422..7b6b0cc9 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -14,7 +14,7 @@ Local Open Scope option_monad_scope. (** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) -(* Implementation of local symbolic internal states +(** * Implementation of local symbolic internal states TODO: use refined symbolic values instead of those from RTLpathSE_theory. *) @@ -560,28 +560,157 @@ Proof. apply hsexec_final_correct; auto. Qed. -(** TODO: all this needs to be adapted +(** * The simulation test of concrete symbolic execution *) +(* TODO - generalize it with a list of registers to test ? *) Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. -Definition hsilocal_simu dm f (hst1 hst2: hsistate_local): Prop := forall st1 st2, +Definition hsilocal_simu dm f (hst1 hst2: hsistate_local) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_local_refines hst1 st1 -> hsistate_local_refines hst2 st2 -> - silocal_simu dm f st1 st2. + silocal_simu dm f st1 st2 ctx. Lemma hsilocal_simub_correct dm f hst1 hst2: hsilocal_simub dm f hst1 hst2 = OK tt -> - hsilocal_simu dm f hst1 hst2. + forall ctx, hsilocal_simu dm f hst1 hst2 ctx. Proof. Admitted. Lemma hsilocal_simub_eq dm f hsl1 hsl2: hsilocal_simub dm f hsl1 hsl2 = OK tt -> - hsi_sreg_get (hsi_sreg hsl1) = hsi_sreg_get (hsi_sreg hsl2) + hsi_sreg_get (hsi_sreg hsl1) = hsi_sreg_get (hsi_sreg hsl2) (* FIXME: a bit too strong ? *) /\ hsi_smem_get (hsi_lsmem hsl1) = hsi_smem_get (hsi_lsmem hsl2). Proof. Admitted. +Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := + match dm ! tn with + | None => Error (msg "revmap_check_single: no mapping for tn") + | Some res => if (Pos.eq_dec n res) then OK tt + else Error (msg "revmap_check_single: n and res do not match") + end. + +Lemma revmap_check_single_correct dm n tn: + revmap_check_single dm n tn = OK tt -> + dm ! tn = Some n. +Proof. + unfold revmap_check_single. explore; try discriminate. congruence. +Qed. + +(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) +Fixpoint sval_simub (sv1 sv2: sval) := + match sv1 with + | Sinput r => + match sv2 with + | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") + | _ => Error (msg "sval_simub: Sinput expected") + end + | Sop op lsv sm => + match sv2 with + | Sop op' lsv' sm' => + if (eq_operation op op') then + do _ <- list_sval_simub lsv lsv'; + smem_simub sm sm' + else Error (msg "sval_simub: different operations in Sop") + | _ => Error (msg "sval_simub: Sop expected") + end + | Sload sm trap chk addr lsv => + match sv2 with + | Sload sm' trap' chk' addr' lsv' => + if (trapping_mode_eq trap trap') then + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + list_sval_simub lsv lsv' + else Error (msg "sval_simub: addressings do not match") + else Error (msg "sval_simub: chunks do not match") + else Error (msg "sval_simub: trapping modes do not match") + (* FIXME - should be refined once we introduce non trapping loads *) + | _ => Error (msg "sval_simub: Sload expected") + end + end +with list_sval_simub (lsv1 lsv2: list_sval) := + match lsv1 with + | Snil => + match lsv2 with + | Snil => OK tt + | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") + end + | Scons sv1 lsv1 => + match lsv2 with + | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") + | Scons sv2 lsv2 => + do _ <- sval_simub sv1 sv2; + list_sval_simub lsv1 lsv2 + end + end +with smem_simub (sm1 sm2: smem) := + match sm1 with + | Sinit => + match sm2 with + | Sinit => OK tt + | _ => Error (msg "smem_simub: Sinit expected") + end + | Sstore sm chk addr lsv sv => + match sm2 with + | Sstore sm' chk' addr' lsv' sv' => + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + do _ <- list_sval_simub lsv lsv'; + sval_simub sv sv' + else Error (msg "smem_simub: addressings do not match") + else Error (msg "smem_simub: chunks not match") + | _ => Error (msg "smem_simub: Sstore expected") + end + end. + +Lemma sval_simub_correct sv1: forall sv2, + sval_simub sv1 sv2 = OK tt -> sv1 = sv2. +Proof. + induction sv1 using sval_mut with + (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) + (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. + (* Sinput *) + - destruct sv2; try discriminate. + destruct (Pos.eq_dec r r0); [congruence|discriminate]. + (* Sop *) + - destruct sv2; try discriminate. + destruct (eq_operation _ _); [|discriminate]. subst. + intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sload *) + - destruct sv2; try discriminate. + destruct (trapping_mode_eq _ _); [|discriminate]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. + congruence. + (* Snil *) + - destruct lsv2; [|discriminate]. congruence. + (* Scons *) + - destruct lsv2; [discriminate|]. intro. explore. + apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sinit *) + - destruct sm2; [|discriminate]. congruence. + (* Sstore *) + - destruct sm2; [discriminate|]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. + assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. + congruence. +Qed. + +Lemma list_sval_simub_correct lsv1: forall lsv2, + list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. +Proof. + induction lsv1; simpl; auto. + - destruct lsv2; try discriminate. reflexivity. + - destruct lsv2; try discriminate. intro. explore. + apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. + congruence. +Qed. + Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); @@ -590,10 +719,10 @@ Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hs else Error (msg "siexit_simub: conditions do not match") . -Definition hsiexit_simu dm f hse1 hse2: Prop := forall se1 se2, +Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, hsistate_exit_refines hse1 se1 -> hsistate_exit_refines hse2 se2 -> - siexit_simu dm f se1 se2. + siexit_simu dm f ctx se1 se2. (* Lemma seval_condition_refines hst hst' ge sp cond args rs m: hsi_smem_get hst = hsi_smem_get hst' -> @@ -606,20 +735,22 @@ Lemma seval_condition_refines hst st ge sp cond args rs m: Proof. Admitted. +Local Hint Resolve genv_match: core. + Lemma hsiexit_simub_correct dm f hse1 hse2: hsiexit_simub dm f hse1 hse2 = OK tt -> - hsiexit_simu dm f hse1 hse2. + forall ctx, hsiexit_simu dm f hse1 hse2 ctx. Proof. unfold hsiexit_simub. intro. explore. apply list_sval_simub_correct in EQ0. exploit hsilocal_simub_eq; eauto. - intros (SREGEQ & SMEMEQ). - apply hsilocal_simub_correct in EQ2. + intros (SREGEQ & SMEMEQ) ctx. + eapply hsilocal_simub_correct in EQ2. apply revmap_check_single_correct in EQ3. intros se1 se2 HEREF1 HEREF2. destruct HEREF1 as (HCOND1 & HARGS1 & HLREF1 & HIFSO1). destruct HEREF2 as (HCOND2 & HARGS2 & HLREF2 & HIFSO2). constructor. - - unfold siexit_simu_ssem. intros sp ge ge' GFS LOK rs m is SEMEXIT. + - unfold siexit_simu_ssem. intros is SEMEXIT. destruct SEMEXIT as (SCOND & SLOCAL & SIFSO). let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). @@ -636,10 +767,10 @@ Proof. * destruct ISIMU as (A & B & C). constructor; auto. * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. repeat (constructor; auto). simpl in *. congruence. - - unfold siexit_simu_cond. intros sp ge ge' GFS rs m. erewrite seval_condition_preserved; eauto. - rewrite <- HARGS2. rewrite <- HCOND2. erewrite <- seval_condition_refines; eauto. + - rewrite <- HARGS2. rewrite <- HCOND2. erewrite <- seval_condition_refines; eauto. rewrite <- HARGS1. rewrite <- HCOND1. erewrite <- seval_condition_refines; eauto. - congruence. + rewrite SMEMEQ, EQ0, e. + eapply seval_condition_preserved; eauto. Qed. Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := @@ -659,21 +790,21 @@ Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: l end end. -Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit): Prop := forall se1 se2, +Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit) (ctx: simu_proof_context f): Prop := forall se1 se2, list_forall2 hsistate_exit_refines hse1 se1 -> list_forall2 hsistate_exit_refines hse2 se2 -> - siexits_simu dm f se1 se2. + siexits_simu dm f se1 se2 ctx. -Lemma hsiexits_simub_correct dm f lhse1: forall lhse2, +Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, hsiexits_simub dm f lhse1 lhse2 = OK tt -> - hsiexits_simu dm f lhse1 lhse2. + hsiexits_simu dm f lhse1 lhse2 ctx. Proof. induction lhse1. - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREF1 HEREF2. inv HEREF1. inv HEREF2. constructor. - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. fold hsiexits_simub in EQ1. - apply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. + eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. intros se1 se2 HEREF1 HEREF2. inv HEREF1. inv HEREF2. constructor; auto. apply EQ1; assumption. Qed. @@ -683,23 +814,98 @@ Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: h do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); OK tt. -Definition hsistate_simu dm f (hst1 hst2: hsistate): Prop := forall st1 st2, +Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_refines hst1 st1 -> hsistate_refines hst2 st2 -> - sistate_simu dm f st1 st2. + sistate_simu dm f st1 st2 ctx. + +Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, + siexits_simu dm f lse1 lse2 ctx -> + all_fallthrough (the_ge1 ctx) (the_sp ctx) lse1 (the_rs0 ctx) (the_m0 ctx) -> + all_fallthrough (the_ge2 ctx) (the_sp ctx) lse2 (the_rs0 ctx) (the_m0 ctx). +Proof. + induction 1; [unfold all_fallthrough; contradiction|]. + intros X ext INEXT. eapply all_fallthrough_revcons in X. destruct X as (SEVAL & ALLFU). + apply IHlist_forall2 in ALLFU. + destruct H as (_ & CONDSIMU). + inv INEXT; [|eauto]. + erewrite <- CONDSIMU; eauto. +Qed. + +Lemma siexits_simu_all_fallthrough_upto dm f ctx lse1 lse2: + siexits_simu dm f lse1 lse2 ctx -> forall ext1 lx1, + all_fallthrough_upto_exit (the_ge1 ctx) (the_sp ctx) ext1 lx1 lse1 (the_rs0 ctx) (the_m0 ctx) -> + exists ext2 lx2, + all_fallthrough_upto_exit (the_ge2 ctx) (the_sp ctx) ext2 lx2 lse2 (the_rs0 ctx) (the_m0 ctx) + /\ length lx1 = length lx2. +Proof. + induction 1. + - intros. destruct H as (ITAIL & ALLFU). eapply is_tail_false in ITAIL. contradiction. + - intros ext1 lx1 ALLFUE. + destruct ALLFUE as (ITAIL & ALLFU). inv ITAIL. + + eexists; eexists. + constructor; [| eapply list_forall2_length; eauto]. + constructor; [econstructor | eapply siexits_simu_all_fallthrough; eauto]. + + exploit IHlist_forall2; [constructor; eauto|]. + intros (ext2 & lx2 & ALLFUE2 & LENEQ). + eexists; eexists. constructor; eauto. + eapply all_fallthrough_upto_exit_cons; eauto. +Qed. + +Lemma list_forall2_nth_error {A} (l1 l2: list A) P: + list_forall2 P l1 l2 -> + forall x1 x2 n, + nth_error l1 n = Some x1 -> + nth_error l2 n = Some x2 -> + P x1 x2. +Proof. + induction 1. + - intros. rewrite nth_error_nil in H. discriminate. + - intros x1 x2 n. destruct n as [|n]; simpl. + + intros. inv H1. inv H2. assumption. + + apply IHlist_forall2. +Qed. + +Lemma is_tail_length {A} (l1 l2: list A): + is_tail l1 l2 -> + (length l1 <= length l2)%nat. +Proof. + induction l2. + - intro. destruct l1; auto. apply is_tail_false in H. contradiction. + - intros ITAIL. inv ITAIL; auto. + apply IHl2 in H1. clear IHl2. simpl. omega. +Qed. + +Lemma is_tail_nth_error {A} (l1 l2: list A) x: + is_tail (x::l1) l2 -> + nth_error l2 ((length l2) - length l1 - 1) = Some x. +Proof. + induction l2. + - intro ITAIL. apply is_tail_false in ITAIL. contradiction. + - intros ITAIL. assert (length (a::l2) = S (length l2)) by auto. rewrite H. clear H. + assert (forall n n', ((S n) - n' - 1)%nat = (n - n')%nat) by (intros; omega). rewrite H. clear H. + inv ITAIL. + + assert (forall n, (n - n)%nat = 0%nat) by (intro; omega). rewrite H. + simpl. reflexivity. + + exploit IHl2; eauto. intros. clear IHl2. + assert (forall n n', (n > n')%nat -> (n - n')%nat = S (n - n' - 1)%nat) by (intros; omega). + exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2. + assert ((length l2 > length l1)%nat) by omega. clear H2. + rewrite H0; auto. +Qed. (* Very adapted from sistate_simub_correct ---> should delete sistate_simub_correct after *) Lemma hsistate_simub_correct dm f hst1 hst2: hsistate_simub dm f hst1 hst2 = OK tt -> - hsistate_simu f dm hst1 hst2. + forall ctx, hsistate_simu dm f hst1 hst2 ctx. Proof. unfold hsistate_simub. intro. explore. unfold hsistate_simu. - intros st1 st2 HREF1 HREF2 sp ge1 ge2 GFS LOK rs m is1 SEMI. + intros ctx st1 st2 HREF1 HREF2 is1 SEMI. destruct HREF1 as (PCEQ1 & HEREF1 & HLREF1). destruct HREF2 as (PCEQ2 & HEREF2 & HLREF2). exploit hsiexits_simub_correct; eauto. intro ESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). - exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). + exploit hsilocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. + unfold ssem_internal. simpl. @@ -711,7 +917,7 @@ Proof. destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - destruct SEMI as (ext & lx & SSEME & ALLFU). exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). - assert (EXTSIMU: siexit_simu dm f ext ext2). { + assert (EXTSIMU: siexit_simu dm f ctx ext ext2). { eapply list_forall2_nth_error; eauto. - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. @@ -725,32 +931,288 @@ Proof. rewrite <- CONTEQ. rewrite ICONT. constructor; auto. + eexists; eexists; constructor; eauto. + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). - Unshelve. apply dm. (* Very weird to have this *) Qed. -Definition hsstate_simu dm f (hst1 hst2: hsstate): Prop := +Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := forall st1 st2, - (* CS: this hypothesis is not needed? *) -(* (forall s : ident, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> *) hsstate_refines hst1 st1 -> - hsstate_refines hst2 st2 -> sstate_simu f dm st1 st2. + hsstate_refines hst2 st2 -> sstate_simu f dm st1 st2 ctx. + + +Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := + match ln with + | nil => + match ln' with + | nil => OK tt + | _ => Error (msg "revmap_check_list: lists have different lengths") + end + | n::ln => + match ln' with + | nil => Error (msg "revmap_check_list: lists have different lengths") + | n'::ln' => do _ <- revmap_check_single dm n n'; revmap_check_list dm ln ln' + end + end. + +Lemma revmap_check_list_correct dm ln: forall ln', + revmap_check_list dm ln ln' = OK tt -> + ptree_get_list dm ln' = Some ln. +Proof. + induction ln. + - simpl. intros. destruct ln'; try discriminate. constructor; auto. + - simpl. intros; destruct ln'; try discriminate. explore. + apply IHln in EQ0. apply revmap_check_single_correct in EQ. + simpl. rewrite EQ. rewrite EQ0. reflexivity. +Qed. + +Definition svos_simub (svos1 svos2: sval + ident) := + match svos1 with + | inl sv1 => + match svos2 with + | inl sv2 => sval_simub sv1 sv2 + | _ => Error (msg "svos_simub: expected sval") + end + | inr id1 => + match svos2 with + | inr id2 => + if (ident_eq id1 id2) then OK tt + else Error (msg "svos_simub: ids do not match") + | _ => Error (msg "svos_simub: expected id") + end + end. + +Lemma svos_simub_correct svos1 svos2: + svos_simub svos1 svos2 = OK tt -> + svos1 = svos2. +Proof. + destruct svos1. + - simpl. destruct svos2; [|discriminate]. + intro. exploit sval_simub_correct; eauto. congruence. + - simpl. destruct svos2; [discriminate|]. + intro. explore. reflexivity. +Qed. + +Fixpoint builtin_arg_simub (bs bs': builtin_arg sval) := + match bs with + | BA sv => + match bs' with + | BA sv' => sval_simub sv sv' + | _ => Error (msg "builtin_arg_simub: BA expected") + end + | BA_int n => + match bs' with + | BA_int n' => if (Integers.int_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_int expected") + end + | BA_long n => + match bs' with + | BA_long n' => if (int64_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_long expected") + end + | BA_float f => + match bs' with + | BA_float f' => if (float_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_float expected") + end + | BA_single f => + match bs' with + | BA_single f' => if (float32_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_single expected") + end + | BA_loadstack chk ptr => + match bs' with + | BA_loadstack chk' ptr' => + if (chunk_eq chk chk') then + if (ptrofs_eq ptr ptr') then OK tt + else Error (msg "builtin_arg_simub: ptr do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadstack expected") + end + | BA_addrstack ptr => + match bs' with + | BA_addrstack ptr' => if (ptrofs_eq ptr ptr') then OK tt else Error (msg "builtin_arg_simub: ptr do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrstack expected") + end + | BA_loadglobal chk id ofs => + match bs' with + | BA_loadglobal chk' id' ofs' => + if (chunk_eq chk chk') then + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadglobal expected") + end + | BA_addrglobal id ofs => + match bs' with + | BA_addrglobal id' ofs' => + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrglobal expected") + end + | BA_splitlong lo hi => + match bs' with + | BA_splitlong lo' hi' => do _ <- builtin_arg_simub lo lo'; builtin_arg_simub hi hi' + | _ => Error (msg "builtin_arg_simub: BA_splitlong expected") + end + | BA_addptr b1 b2 => + match bs' with + | BA_addptr b1' b2' => do _ <- builtin_arg_simub b1 b1'; builtin_arg_simub b2 b2' + | _ => Error (msg "builtin_arg_simub: BA_addptr expected") + end + end. + +Lemma builtin_arg_simub_correct b1: forall b2, + builtin_arg_simub b1 b2 = OK tt -> b1 = b2. +Proof. + induction b1; simpl; destruct b2; try discriminate; auto; intros; try (explore; congruence). + - apply sval_simub_correct in H. congruence. + - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. + - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. +Qed. + +Fixpoint list_builtin_arg_simub lbs1 lbs2 := + match lbs1 with + | nil => + match lbs2 with + | nil => OK tt + | _ => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs2 is bigger)") + end + | bs1::lbs1 => + match lbs2 with + | nil => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs1 is bigger)") + | bs2::lbs2 => + do _ <- builtin_arg_simub bs1 bs2; + list_builtin_arg_simub lbs1 lbs2 + end + end. + +Lemma list_builtin_arg_simub_correct lsb1: forall lsb2, + list_builtin_arg_simub lsb1 lsb2 = OK tt -> lsb1 = lsb2. +Proof. + induction lsb1; intros; simpl; destruct lsb2; try discriminate; auto. + simpl in H. explore. apply builtin_arg_simub_correct in EQ. + assert (lsb1 = lsb2) by auto. congruence. +Qed. + +(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) +Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := + match fv1 with + | Snone => + match fv2 with + | Snone => revmap_check_single dm pc1 pc2 + | _ => Error (msg "sfval_simub: Snone expected") + end + | Scall sig svos lsv res pc1 => + match fv2 with + | Scall sig2 svos2 lsv2 res2 pc2 => + do _ <- revmap_check_single dm pc1 pc2; + if (signature_eq sig sig2) then + if (Pos.eq_dec res res2) then + do _ <- svos_simub svos svos2; + list_sval_simub lsv lsv2 + else Error (msg "sfval_simub: Scall res do not match") + else Error (msg "sfval_simub: Scall different signatures") + | _ => Error (msg "sfval_simub: Scall expected") + end + | Stailcall sig svos lsv => + match fv2 with + | Stailcall sig' svos' lsv' => + if (signature_eq sig sig') then + do _ <- svos_simub svos svos'; + list_sval_simub lsv lsv' + else Error (msg "sfval_simub: signatures do not match") + | _ => Error (msg "sfval_simub: Stailcall expected") + end + | Sbuiltin ef lbs br pc => + match fv2 with + | Sbuiltin ef' lbs' br' pc' => + if (external_function_eq ef ef') then + if (builtin_res_eq_pos br br') then + do _ <- revmap_check_single dm pc pc'; + list_builtin_arg_simub lbs lbs' + else Error (msg "sfval_simub: builtin res do not match") + else Error (msg "sfval_simub: external functions do not match") + | _ => Error (msg "sfval_simub: Sbuiltin expected") + end + | Sjumptable sv ln => + match fv2 with + | Sjumptable sv' ln' => + do _ <- revmap_check_list dm ln ln'; sval_simub sv sv' + | _ => Error (msg "sfval_simub: Sjumptable expected") + end + | Sreturn osv => + match fv2 with + | Sreturn osv' => + match osv with + | Some sv => + match osv' with + | Some sv' => sval_simub sv sv' + | None => Error (msg "sfval_simub sv' expected") + end + | None => + match osv' with + | Some sv' => Error (msg "None expected") + | None => OK tt + end + end + | _ => Error (msg "sfval_simub: Sreturn expected") + end + end. + +Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2 ctx: + sfval_simub dm f pc1 pc2 fv1 fv2 = OK tt -> + sfval_simu dm f pc1 pc2 ctx fv1 fv2. +Proof. + unfold sfval_simub. destruct fv1. + (* Snone *) + - destruct fv2; try discriminate. intro. + apply revmap_check_single_correct in H. constructor; auto. + (* Scall *) + - destruct fv2; try discriminate. intro. explore. + apply svos_simub_correct in EQ3. apply list_sval_simub_correct in EQ4. + subst. apply revmap_check_single_correct in EQ. constructor; auto. + + admit. + + admit. + (* Stailcall *) + - destruct fv2; try discriminate. intro. explore. + apply svos_simub_correct in EQ0. apply list_sval_simub_correct in EQ1. + subst. constructor; auto. + + admit. + + admit. + (* Sbuiltin *) + - destruct fv2; try discriminate. intro. explore. + apply revmap_check_single_correct in EQ1. apply list_builtin_arg_simub_correct in EQ2. + subst. constructor; auto. + (* Sjumptable *) + - destruct fv2; try discriminate. intro. explore. + apply revmap_check_list_correct in EQ. apply sval_simub_correct in EQ0. subst. + constructor; auto. + admit. + (* Sreturn *) + - destruct fv2; try discriminate. destruct o; destruct o0; try discriminate. + + intro. apply sval_simub_correct in H. subst. constructor; auto. + + constructor; auto. +Admitted. Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); do u2 <- sfval_simub dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2); OK tt. -Lemma hsstate_simub_correct dm f hst1 hst2: +Lemma hsstate_simub_correct dm f hst1 hst2 ctx: hsstate_simub dm f hst1 hst2 = OK tt -> - hsstate_simu f dm hst1 hst2. + hsstate_simu f dm hst1 hst2 ctx. Proof. unfold hsstate_simub. intro. explore. - apply sfval_simub_correct in EQ1. apply hsistate_simub_correct in EQ. + eapply sfval_simub_correct in EQ1. eapply hsistate_simub_correct in EQ. intros st1 st2 (LREF & FREF) (LREF' & FREF'). constructor; [auto|]. destruct LREF as (PCEQ & _ & _). destruct LREF' as (PCEQ' & _ & _). rewrite <- PCEQ. rewrite <- PCEQ'. rewrite <- FREF. rewrite <- FREF'. - assumption. + eauto. (* FIXME - that proof will have to be changed once hfinal_refines is more complex *) Qed. @@ -780,7 +1242,7 @@ Proof. intros (st0 & SEXEC1 & REF1). rewrite SEXEC1 in SEXEC; inversion SEXEC; subst. eexists; split; eauto. - eapply hsstate_simub_correct; eauto. + intros; eapply hsstate_simub_correct; eauto. Qed. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := @@ -802,21 +1264,15 @@ Proof. + subst. eapply simu_check_single_correct; eauto. + eapply IHlm; assumption. Qed. -*) Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := - (* TODO: adapt this - simu_check_rec dm f tf (PTree.elements dm) - *) - OK tt. + simu_check_rec dm f tf (PTree.elements dm). Lemma simu_check_correct dm f tf: simu_check dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2. -Admitted. -(* TODO: adapt this. Proof. unfold simu_check. intros. eapply PTree.elements_correct in H0. eapply simu_check_rec_correct; eassumption. -Qed.*) \ No newline at end of file +Qed. \ No newline at end of file diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 2be567d6..e81a107e 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1,6 +1,6 @@ (* A theory of symbolic execution on RTLpath -NB: an efficient implementation with hash-consing will be defined in ... +NB: an efficient implementation with hash-consing will be defined in RTLpathSE_impl.v *) @@ -1617,6 +1617,9 @@ End SymbValPreserved. 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) @@ -1634,140 +1637,11 @@ Definition istate_simu f (srce: PTree.t node) is1 is2: Prop := /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2 /\ srce!(is2.(ipc)) = Some is1.(ipc). -(* TODO: code à adapter pour plus tard -Fixpoint sval_simub (sv1 sv2: sval) := - match sv1 with - | Sinput r => - match sv2 with - | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") - | _ => Error (msg "sval_simub: Sinput expected") - end - | Sop op lsv sm => - match sv2 with - | Sop op' lsv' sm' => - if (eq_operation op op') then - do _ <- list_sval_simub lsv lsv'; - smem_simub sm sm' - else Error (msg "sval_simub: different operations in Sop") - | _ => Error (msg "sval_simub: Sop expected") - end - | Sload sm trap chk addr lsv => - match sv2 with - | Sload sm' trap' chk' addr' lsv' => - if (trapping_mode_eq trap trap') then - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - list_sval_simub lsv lsv' - else Error (msg "sval_simub: addressings do not match") - else Error (msg "sval_simub: chunks do not match") - else Error (msg "sval_simub: trapping modes do not match") - (* FIXME - should be refined once we introduce non trapping loads *) - | _ => Error (msg "sval_simub: Sload expected") - end - end -with list_sval_simub (lsv1 lsv2: list_sval) := - match lsv1 with - | Snil => - match lsv2 with - | Snil => OK tt - | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") - end - | Scons sv1 lsv1 => - match lsv2 with - | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") - | Scons sv2 lsv2 => - do _ <- sval_simub sv1 sv2; - list_sval_simub lsv1 lsv2 - end - end -with smem_simub (sm1 sm2: smem) := - match sm1 with - | Sinit => - match sm2 with - | Sinit => OK tt - | _ => Error (msg "smem_simub: Sinit expected") - end - | Sstore sm chk addr lsv sv => - match sm2 with - | Sstore sm' chk' addr' lsv' sv' => - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - do _ <- list_sval_simub lsv lsv'; - sval_simub sv sv' - else Error (msg "smem_simub: addressings do not match") - else Error (msg "smem_simub: chunks not match") - | _ => Error (msg "smem_simub: Sstore expected") - end - end. - -Lemma sval_simub_correct sv1: forall sv2, - sval_simub sv1 sv2 = OK tt -> sv1 = sv2. -Proof. - induction sv1 using sval_mut with - (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) - (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. - (* Sinput *) - - destruct sv2; try discriminate. - destruct (Pos.eq_dec r r0); [congruence|discriminate]. - (* Sop *) - - destruct sv2; try discriminate. - destruct (eq_operation _ _); [|discriminate]. subst. - intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sload *) - - destruct sv2; try discriminate. - destruct (trapping_mode_eq _ _); [|discriminate]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. - congruence. - (* Snil *) - - destruct lsv2; [|discriminate]. congruence. - (* Scons *) - - destruct lsv2; [discriminate|]. intro. explore. - apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sinit *) - - destruct sm2; [|discriminate]. congruence. - (* Sstore *) - - destruct sm2; [discriminate|]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. - assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. - congruence. -Qed. - -Lemma list_sval_simub_correct lsv1: forall lsv2, - list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. -Proof. - induction lsv1; simpl; auto. - - destruct lsv2; try discriminate. reflexivity. - - destruct lsv2; try discriminate. intro. explore. - apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. - congruence. -Qed. -*) - -Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := - match dm ! tn with - | None => Error (msg "revmap_check_single: no mapping for tn") - | Some res => if (Pos.eq_dec n res) then OK tt - else Error (msg "revmap_check_single: n and res do not match") - end. - -Lemma revmap_check_single_correct dm n tn: - revmap_check_single dm n tn = OK tt -> - dm ! tn = Some n. -Proof. - unfold revmap_check_single. explore; try discriminate. congruence. -Qed. - 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_ge2 s = Genv.find_symbol the_ge1 s; + 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 @@ -1780,20 +1654,6 @@ Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sista exists is2, ssem_internal (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) is2 /\ istate_simu f dm is1 is2. -Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) (ctx: simu_proof_context f): Prop := - forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) sl1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> - exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) sl2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) - /\ istate_simu f dm is1 is2. - -Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f):= - forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> - exists is2, - ssem_exit (the_ge1 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) - /\ istate_simu f dm is1 is2. - -Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f) := - siexit_simu_ssem dm f se1 se2 ctx /\ (si_cond se1) = (si_cond se2). (* NB on si_cond: a pure syntactic definition should suffice here *) - Definition opt_simu {A} (o1: option A) (o2: option A) := o1 <> None -> o2 = o1. @@ -1890,477 +1750,29 @@ Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop := forall st1, sexec f1 pc1 = Some st1 -> exists st2, sexec f2 pc2 = Some st2 /\ forall ctx, sstate_simu dm f1 st1 st2 ctx. -(* TODO: all this needs to be adapted .... - - -(* TODO - generalize it with a list of registers to test *) -Definition silocal_simub (dm: PTree.t node) (f: RTLpath.function) (* (regs: list reg) *) (sl1 sl2: sistate_local) := - OK tt -. - -Lemma silocal_simub_eq dm f sl1 sl2: - silocal_simub dm f sl1 sl2 = OK tt -> - si_sreg sl1 = si_sreg sl2 /\ si_smem sl1 = si_smem sl2. -Proof. -Admitted. - -Lemma silocal_simub_correct dm f sl1 sl2: - silocal_simub dm f sl1 sl2 = OK tt -> - silocal_simu dm f sl1 sl2. -Proof. -Admitted. - +(** Maybe it could be useful to develop a small theory here to help in decomposing the simulation test ? -Lemma siexit_simub_correct dm f se1 se2: - siexit_simub dm f se1 se2 = OK tt -> - siexit_simu dm f se1 se2. -Proof. - unfold siexit_simub. intro. explore. - apply list_sval_simub_correct in EQ0. exploit silocal_simub_eq; eauto. - intros (SREGEQ & SMEMEQ). - apply silocal_simub_correct in EQ2. - apply revmap_check_single_correct in EQ3. - constructor. - - unfold siexit_simu_ssem. intros sp ge ge' GFS LOK rs m is SEMEXIT. - destruct SEMEXIT as (SCOND & SLOCAL & SIFSO). - let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). - eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). - destruct is2 as [icont ipc irs imem]. simpl in *. - exists (mk_istate icont (si_ifso se2) irs imem). simpl. - constructor; auto. - + repeat (constructor; auto). - rewrite <- EQ0. rewrite <- e. rewrite <- SMEMEQ. erewrite seval_condition_preserved; eauto. - + unfold istate_simu in *. destruct (icontinue is) eqn:ICONT. - * destruct ISIMU as (A & B & C). constructor; auto. - * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. - repeat (constructor; auto). simpl in *. congruence. - - unfold siexit_simu_cond. intros sp ge ge' GFS rs m. erewrite seval_condition_preserved; eauto. congruence. -Qed. - -Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := - list_forall2 (siexit_simu dm f) lse1 lse2. - -Lemma siexits_simu_all_fallthrough dm f ge1 ge2 sp rs m: forall lse1 lse2, - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - all_fallthrough ge1 sp lse1 rs m -> - siexits_simu dm f lse1 lse2 -> - all_fallthrough ge2 sp lse2 rs m. -Proof. - induction 3; unfold all_fallthrough; [contradiction|]. - intros ext INEXT. apply all_fallthrough_revcons in H0. destruct H0 as (SEVAL & ALLFU). - apply IHlist_forall2 in ALLFU. - destruct H1 as (_ & CONDSIMU). - inv INEXT; [|eauto]. - erewrite <- CONDSIMU; eauto. -Qed. - -Lemma siexits_simu_all_fallthrough_upto dm f ge1 ge2 sp rs m lse1 lse2: - (forall s, Genv.find_symbol ge2 s = Genv.find_symbol ge1 s) -> - siexits_simu dm f lse1 lse2 -> forall ext1 lx1, - all_fallthrough_upto_exit ge1 sp ext1 lx1 lse1 rs m -> - exists ext2 lx2, - all_fallthrough_upto_exit ge2 sp ext2 lx2 lse2 rs m - /\ length lx1 = length lx2. -Proof. - induction 2. - - intros. destruct H0 as (ITAIL & ALLFU). apply is_tail_false in ITAIL. contradiction. - - intros ext1 lx1 ALLFUE. - destruct ALLFUE as (ITAIL & ALLFU). inv ITAIL. - + eexists; eexists. - constructor; [| eapply list_forall2_length; eauto]. - constructor; auto. eapply siexits_simu_all_fallthrough; eauto. - + exploit IHlist_forall2; [constructor; eauto|]. - intros (ext2 & lx2 & ALLFUE2 & LENEQ). - eexists; eexists. constructor; eauto. - eapply all_fallthrough_upto_exit_cons; eauto. -Qed. - - -Definition siexit_simub (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) := - if (eq_condition (si_cond se1) (si_cond se2)) then - do _ <- list_sval_simub (si_scondargs se1) (si_scondargs se2); - do _ <- silocal_simub dm f (si_elocal se1) (si_elocal se2); (* TODO - generalize it with a list of registers to test *) - revmap_check_single dm (si_ifso se1) (si_ifso se2) - else Error (msg "siexit_simub: conditions do not match") -. +Here are intermediate definitions. -Fixpoint siexits_simub (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) := - match lse1 with - | nil => - match lse2 with - | nil => OK tt - | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") - end - | se1 :: lse1 => - match lse2 with - | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") - | se2 :: lse2 => - do _ <- siexit_simub dm f se1 se2; - do _ <- siexits_simub dm f lse1 lse2; - OK tt - end - end. - -Lemma siexits_simub_correct dm f lse1: forall lse2, - siexits_simub dm f lse1 lse2 = OK tt -> - siexits_simu dm f lse1 lse2. -Proof. - induction lse1. - - simpl. intros. destruct lse2; try discriminate. constructor. - - simpl. intros. destruct lse2; try discriminate. explore. - apply siexit_simub_correct in EQ. apply IHlse1 in EQ1. constructor; auto. -Qed. - -Definition sistate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) := - do _ <- silocal_simub dm f (si_local st1) (si_local st2); - do _ <- siexits_simub dm f (si_exits st1) (si_exits st2); - OK tt. - -Lemma list_forall2_nth_error {A} (l1 l2: list A) P: - list_forall2 P l1 l2 -> - forall x1 x2 n, - nth_error l1 n = Some x1 -> - nth_error l2 n = Some x2 -> - P x1 x2. -Proof. - induction 1. - - intros. rewrite nth_error_nil in H. discriminate. - - intros x1 x2 n. destruct n as [|n]; simpl. - + intros. inv H1. inv H2. assumption. - + apply IHlist_forall2. -Qed. - -Lemma is_tail_length {A} (l1 l2: list A): - is_tail l1 l2 -> - (length l1 <= length l2)%nat. -Proof. - induction l2. - - intro. destruct l1; auto. apply is_tail_false in H. contradiction. - - intros ITAIL. inv ITAIL; auto. - apply IHl2 in H1. clear IHl2. simpl. omega. -Qed. - -Lemma is_tail_nth_error {A} (l1 l2: list A) x: - is_tail (x::l1) l2 -> - nth_error l2 ((length l2) - length l1 - 1) = Some x. -Proof. - induction l2. - - intro ITAIL. apply is_tail_false in ITAIL. contradiction. - - intros ITAIL. assert (length (a::l2) = S (length l2)) by auto. rewrite H. clear H. - assert (forall n n', ((S n) - n' - 1)%nat = (n - n')%nat) by (intros; omega). rewrite H. clear H. - inv ITAIL. - + assert (forall n, (n - n)%nat = 0%nat) by (intro; omega). rewrite H. - simpl. reflexivity. - + exploit IHl2; eauto. intros. clear IHl2. - assert (forall n n', (n > n')%nat -> (n - n')%nat = S (n - n' - 1)%nat) by (intros; omega). - exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2. - assert ((length l2 > length l1)%nat) by omega. clear H2. - rewrite H0; auto. -Qed. - -Lemma sistate_simub_correct dm f st1 st2: - sistate_simub dm f st1 st2 = OK tt -> - sistate_simu f dm st1 st2. -Proof. - unfold sistate_simub. intro. explore. unfold sistate_simu. - intros sp ge1 ge2 GFS LOK (* INJ *) rs m is1 SEMI. - exploit siexits_simub_correct; eauto. intro ESIMU. - unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - - destruct SEMI as (SSEML & PCEQ & ALLFU). - exploit silocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). - destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. - exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. - + unfold ssem_internal. simpl. - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. - destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. - rewrite <- CONTEQ. rewrite ICONT. - repeat (constructor; auto). - eapply siexits_simu_all_fallthrough; eauto. - + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. - destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - - destruct SEMI as (ext & lx & SSEME & ALLFU). - exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). - assert (EXTSIMU: siexit_simu dm f ext ext2). { - eapply list_forall2_nth_error; eauto. - - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. - assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). - congruence. } - destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. - destruct SSEME as (is2 & SSEME2 & ISIMU). - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). - destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). - unfold ssem_internal. exists is2. - rewrite <- CONTEQ. rewrite ICONT. constructor; auto. - + eexists; eexists; constructor; eauto. - + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). -Qed. - -Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := - match ln with - | nil => - match ln' with - | nil => OK tt - | _ => Error (msg "revmap_check_list: lists have different lengths") - end - | n::ln => - match ln' with - | nil => Error (msg "revmap_check_list: lists have different lengths") - | n'::ln' => do _ <- revmap_check_single dm n n'; revmap_check_list dm ln ln' - end - end. - -Lemma revmap_check_list_correct dm ln: forall ln', - revmap_check_list dm ln ln' = OK tt -> - ptree_get_list dm ln' = Some ln. -Proof. - induction ln. - - simpl. intros. destruct ln'; try discriminate. constructor; auto. - - simpl. intros; destruct ln'; try discriminate. explore. - apply IHln in EQ0. apply revmap_check_single_correct in EQ. - simpl. rewrite EQ. rewrite EQ0. reflexivity. -Qed. - -Definition svos_simub (svos1 svos2: sval + ident) := - match svos1 with - | inl sv1 => - match svos2 with - | inl sv2 => sval_simub sv1 sv2 - | _ => Error (msg "svos_simub: expected sval") - end - | inr id1 => - match svos2 with - | inr id2 => - if (ident_eq id1 id2) then OK tt - else Error (msg "svos_simub: ids do not match") - | _ => Error (msg "svos_simub: expected id") - end - end. - -Lemma svos_simub_correct svos1 svos2: - svos_simub svos1 svos2 = OK tt -> - svos1 = svos2. -Proof. - destruct svos1. - - simpl. destruct svos2; [|discriminate]. - intro. exploit sval_simub_correct; eauto. congruence. - - simpl. destruct svos2; [discriminate|]. - intro. explore. reflexivity. -Qed. - -Fixpoint builtin_arg_simub (bs bs': builtin_arg sval) := - match bs with - | BA sv => - match bs' with - | BA sv' => sval_simub sv sv' - | _ => Error (msg "builtin_arg_simub: BA expected") - end - | BA_int n => - match bs' with - | BA_int n' => if (Integers.int_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") - | _ => Error (msg "buitin_arg_simub: BA_int expected") - end - | BA_long n => - match bs' with - | BA_long n' => if (int64_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") - | _ => Error (msg "buitin_arg_simub: BA_long expected") - end - | BA_float f => - match bs' with - | BA_float f' => if (float_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") - | _ => Error (msg "builtin_arg_simub: BA_float expected") - end - | BA_single f => - match bs' with - | BA_single f' => if (float32_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") - | _ => Error (msg "builtin_arg_simub: BA_single expected") - end - | BA_loadstack chk ptr => - match bs' with - | BA_loadstack chk' ptr' => - if (chunk_eq chk chk') then - if (ptrofs_eq ptr ptr') then OK tt - else Error (msg "builtin_arg_simub: ptr do not match") - else Error (msg "builtin_arg_simub: chunks do not match") - | _ => Error (msg "builtin_arg_simub: BA_loadstack expected") - end - | BA_addrstack ptr => - match bs' with - | BA_addrstack ptr' => if (ptrofs_eq ptr ptr') then OK tt else Error (msg "builtin_arg_simub: ptr do not match") - | _ => Error (msg "builtin_arg_simub: BA_addrstack expected") - end - | BA_loadglobal chk id ofs => - match bs' with - | BA_loadglobal chk' id' ofs' => - if (chunk_eq chk chk') then - if (ident_eq id id') then - if (ptrofs_eq ofs ofs') then OK tt - else Error (msg "builtin_arg_simub: offsets do not match") - else Error (msg "builtin_arg_simub: identifiers do not match") - else Error (msg "builtin_arg_simub: chunks do not match") - | _ => Error (msg "builtin_arg_simub: BA_loadglobal expected") - end - | BA_addrglobal id ofs => - match bs' with - | BA_addrglobal id' ofs' => - if (ident_eq id id') then - if (ptrofs_eq ofs ofs') then OK tt - else Error (msg "builtin_arg_simub: offsets do not match") - else Error (msg "builtin_arg_simub: identifiers do not match") - | _ => Error (msg "builtin_arg_simub: BA_addrglobal expected") - end - | BA_splitlong lo hi => - match bs' with - | BA_splitlong lo' hi' => do _ <- builtin_arg_simub lo lo'; builtin_arg_simub hi hi' - | _ => Error (msg "builtin_arg_simub: BA_splitlong expected") - end - | BA_addptr b1 b2 => - match bs' with - | BA_addptr b1' b2' => do _ <- builtin_arg_simub b1 b1'; builtin_arg_simub b2 b2' - | _ => Error (msg "builtin_arg_simub: BA_addptr expected") - end - end. - -Lemma builtin_arg_simub_correct b1: forall b2, - builtin_arg_simub b1 b2 = OK tt -> b1 = b2. -Proof. - induction b1; simpl; destruct b2; try discriminate; auto; intros; try (explore; congruence). - - apply sval_simub_correct in H. congruence. - - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. - - explore. assert (b1_1 = b2_1) by auto. assert (b1_2 = b2_2) by auto. congruence. -Qed. - -Fixpoint list_builtin_arg_simub lbs1 lbs2 := - match lbs1 with - | nil => - match lbs2 with - | nil => OK tt - | _ => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs2 is bigger)") - end - | bs1::lbs1 => - match lbs2 with - | nil => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs1 is bigger)") - | bs2::lbs2 => - do _ <- builtin_arg_simub bs1 bs2; - list_builtin_arg_simub lbs1 lbs2 - end - end. - -Lemma list_builtin_arg_simub_correct lsb1: forall lsb2, - list_builtin_arg_simub lsb1 lsb2 = OK tt -> lsb1 = lsb2. -Proof. - induction lsb1; intros; simpl; destruct lsb2; try discriminate; auto. - simpl in H. explore. apply builtin_arg_simub_correct in EQ. - assert (lsb1 = lsb2) by auto. congruence. -Qed. +*) -Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := - match fv1 with - | Snone => - match fv2 with - | Snone => revmap_check_single dm pc1 pc2 - | _ => Error (msg "sfval_simub: Snone expected") - end - | Scall sig svos lsv res pc1 => - match fv2 with - | Scall sig2 svos2 lsv2 res2 pc2 => - do _ <- revmap_check_single dm pc1 pc2; - if (signature_eq sig sig2) then - if (Pos.eq_dec res res2) then - do _ <- svos_simub svos svos2; - list_sval_simub lsv lsv2 - else Error (msg "sfval_simub: Scall res do not match") - else Error (msg "sfval_simub: Scall different signatures") - | _ => Error (msg "sfval_simub: Scall expected") - end - | Stailcall sig svos lsv => - match fv2 with - | Stailcall sig' svos' lsv' => - if (signature_eq sig sig') then - do _ <- svos_simub svos svos'; - list_sval_simub lsv lsv' - else Error (msg "sfval_simub: signatures do not match") - | _ => Error (msg "sfval_simub: Stailcall expected") - end - | Sbuiltin ef lbs br pc => - match fv2 with - | Sbuiltin ef' lbs' br' pc' => - if (external_function_eq ef ef') then - if (builtin_res_eq_pos br br') then - do _ <- revmap_check_single dm pc pc'; - list_builtin_arg_simub lbs lbs' - else Error (msg "sfval_simub: builtin res do not match") - else Error (msg "sfval_simub: external functions do not match") - | _ => Error (msg "sfval_simub: Sbuiltin expected") - end - | Sjumptable sv ln => - match fv2 with - | Sjumptable sv' ln' => - do _ <- revmap_check_list dm ln ln'; sval_simub sv sv' - | _ => Error (msg "sfval_simub: Sjumptable expected") - end - | Sreturn osv => - match fv2 with - | Sreturn osv' => - match osv with - | Some sv => - match osv' with - | Some sv' => sval_simub sv sv' - | None => Error (msg "sfval_simub sv' expected") - end - | None => - match osv' with - | Some sv' => Error (msg "None expected") - | None => OK tt - end - end - | _ => Error (msg "sfval_simub: Sreturn expected") - end - end. +Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) (ctx: simu_proof_context f): Prop := + forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) sl1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> + exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) sl2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) + /\ istate_simu f dm is1 is2. -Lemma sfval_simub_correct dm f pc1 pc2 fv1 fv2: - sfval_simub dm f pc1 pc2 fv1 fv2 = OK tt -> - sfval_simu dm f pc1 pc2 fv1 fv2. -Proof. - unfold sfval_simub. destruct fv1. - (* Snone *) - - destruct fv2; try discriminate. intro. - apply revmap_check_single_correct in H. constructor; auto. - (* Scall *) - - destruct fv2; try discriminate. intro. explore. - apply svos_simub_correct in EQ3. apply list_sval_simub_correct in EQ4. - subst. apply revmap_check_single_correct in EQ. constructor; auto. - (* Stailcall *) - - destruct fv2; try discriminate. intro. explore. - apply svos_simub_correct in EQ0. apply list_sval_simub_correct in EQ1. - subst. constructor; auto. - (* Sbuiltin *) - - destruct fv2; try discriminate. intro. explore. - apply revmap_check_single_correct in EQ1. apply list_builtin_arg_simub_correct in EQ2. - subst. constructor; auto. - (* Sjumptable *) - - destruct fv2; try discriminate. intro. explore. - apply revmap_check_list_correct in EQ. apply sval_simub_correct in EQ0. subst. - constructor; auto. - (* Sreturn *) - - destruct fv2; try discriminate. destruct o; destruct o0; try discriminate. - + intro. apply sval_simub_correct in H. subst. constructor; auto. - + constructor; auto. -Qed. +Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f):= + forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> + exists is2, + ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) + /\ istate_simu f dm is1 is2. -(* Computable check of sstate_simu *) -Definition sstate_simub (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sstate) := - do u1 <- sistate_simub dm f (internal st1) (internal st2); - do u2 <- sfval_simub dm f (si_pc st1) (si_pc st2) (final st1) (final st2); - OK tt. +Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (ctx: simu_proof_context f) (se1 se2: sistate_exit) := + siexit_simu_ssem dm f se1 se2 ctx + /\ + seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx) + = seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx). -Lemma sstate_simub_correct dm f st1 st2: - sstate_simub dm f st1 st2 = OK tt -> - forall ctx, sstate_simu dm f st1 st2 ctx. -Proof. - unfold sstate_simub. intros. explore. constructor. - apply sistate_simub_correct; assumption. - apply sfval_simub_correct; assumption. -Qed. -*) +Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) (ctx: simu_proof_context f) := + list_forall2 (siexit_simu dm f ctx) lse1 lse2. diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v index 892fe2fa..5c32847e 100644 --- a/kvx/lib/RTLpathSchedulerproof.v +++ b/kvx/lib/RTLpathSchedulerproof.v @@ -140,7 +140,7 @@ Program Definition mkctx sp rs0 m0 {f1: RTLpath.function} (hyp: liveness_ok_func : simu_proof_context f1 := {| the_ge1:= ge; the_ge2 := tge; the_sp:=sp; the_rs0:=rs0; the_m0:=m0 |}. Obligation 2. - eapply symbols_preserved_RTL. + erewrite symbols_preserved_RTL. eauto. Qed. Lemma s_find_function_preserved f sp svos1 svos2 rs0 m0 fd -- cgit From 14886020bda9c5461d488ff316b95864f1e1789c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 9 Jul 2020 10:51:41 +0200 Subject: Added check on last instruction --- kvx/lib/RTLpathScheduleraux.ml | 20 +++++++++++++++++++- 1 file changed, 19 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 5e9111d2..11539994 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -194,13 +194,31 @@ let change_predicted_successor i s = match i with | None -> failwith "Predicted a successor for an Icond with p=None - unpredicted CB should not be moved in the middle of the superblock" ) +(** + * Perform basic checks on the new order : + * - must have the same length as the old order + * - non basic instructions (call, tailcall, return, jumptable, non predicted CB) must not move + *) +let check_order code old_order new_order = begin + assert ((Array.length old_order) == (Array.length new_order)); + let length = Array.length new_order in + if length > 0 then + let last_inst = Array.get old_order (length - 1) in + let instr = get_some @@ PTree.get last_inst code in + match predicted_successor instr with + | None -> + if (last_inst != Array.get new_order (length - 1)) then + failwith "The last instruction of the superblock is not basic, but was moved" + | _ -> () +end + let apply_schedule code sb new_order = let tc = ref code in let old_order = sb.instructions in let last_node = Array.get old_order (Array.length old_order - 1) in let last_successors = successors_inst @@ get_some @@ PTree.get last_node code in begin - assert ((Array.length old_order) == (Array.length new_order)); + check_order code old_order new_order; Array.iteri (fun i n' -> let inst' = get_some @@ PTree.get n' code in let new_inst = -- cgit From df6ccab1280be1f5a47c3def1dd9cba0c91016dd Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 9 Jul 2020 15:06:07 +0200 Subject: seval_condition_refines ---> refinement correct only for successful executions ? --- kvx/lib/RTLpathSE_impl.v | 66 ++++++++++++++++++++++++++++++++++-------------- 1 file changed, 47 insertions(+), 19 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index f5285070..294867c0 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -760,19 +760,25 @@ Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hs else Error (msg "siexit_simub: conditions do not match") . +Definition hsok_exit ge sp rs m hse := hsok_local ge sp rs m (hsi_elocal hse). + Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, hsistate_exit_refines_stat hse1 se1 -> hsistate_exit_refines_stat hse2 se2 -> hsistate_exit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> + hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> + hsok_exit (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 -> siexit_simu dm f ctx se1 se2. - Lemma seval_condition_refines hst st ge sp cond args rs m: + hsok_local ge sp rs m hst -> hsistate_local_refines ge sp rs m hst st -> seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. Proof. -Admitted. + intros HOK (OKEQ & MEMEQ & RSEQ). unfold seval_condition. + rewrite <- MEMEQ; auto. rewrite hsi_smem_eval_correct. reflexivity. +Qed. Local Hint Resolve genv_match: core. @@ -785,7 +791,7 @@ Proof. intros (SREGEQ & SMEMEQ) ctx. eapply hsilocal_simub_correct in EQ2. apply revmap_check_single_correct in EQ3. - intros se1 se2 (HCOND1 & HARGS1 & HIFSO1) (HCOND2 & HARGS2 & HIFSO2) HLREF1 HLREF2. + intros se1 se2 (HCOND1 & HARGS1 & HIFSO1) (HCOND2 & HARGS2 & HIFSO2) HLREF1 HLREF2 HOK1 HOK2. unfold hsistate_exit_refines_dyn in *. constructor. - unfold siexit_simu_ssem. intros is SEMEXIT. @@ -828,11 +834,18 @@ Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: l end end. +Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := + | hsok_exits_nil: hsok_exits ge sp rs m nil + | hsok_exits_cons: forall l ex, hsok_exits ge sp rs m l -> hsok_exit ge sp rs m ex -> hsok_exits ge sp rs m (ex::l) + . + Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit) (ctx: simu_proof_context f): Prop := forall se1 se2, list_forall2 hsistate_exit_refines_stat hse1 se1 -> list_forall2 hsistate_exit_refines_stat hse2 se2 -> list_forall2 (hsistate_exit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse1 se1 -> list_forall2 (hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse2 se2 -> + hsok_exits (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> + hsok_exits (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 -> siexits_simu dm f se1 se2 ctx. Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, @@ -840,12 +853,13 @@ Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, hsiexits_simu dm f lhse1 lhse2 ctx. Proof. induction lhse1. - - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. + - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _ _ _. inv HEREFS1. inv HEREFS2. constructor. - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. fold hsiexits_simub in EQ1. eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. - intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. + intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2 HOK1 HOK2. + inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. inv HOK1. inv HOK2. constructor; auto. apply EQ1; assumption. Qed. @@ -854,11 +868,16 @@ Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: h do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); OK tt. +Definition hsok_internal ge sp rs m hst := + hsok_exits ge sp rs m (hsi_exits hst) /\ hsok_local ge sp rs m (hsi_local hst). + Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_refines_stat hst1 st1 -> hsistate_refines_stat hst2 st2 -> hsistate_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + hsok_internal (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> + hsok_internal (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 -> sistate_simu dm f st1 st2 ctx. Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, @@ -942,7 +961,8 @@ Lemma hsistate_simub_correct dm f hst1 hst2: forall ctx, hsistate_simu dm f hst1 hst2 ctx. Proof. unfold hsistate_simub. intro. explore. unfold hsistate_simu. - intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) (HREF1 & HLREF1) (HREF2 & HLREF2) is1 SEMI. + intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) + (HREF1 & HLREF1) (HREF2 & HLREF2) (HOKE1 & HOKL1) (HOKE2 & HOKL2) is1 SEMI. exploit hsiexits_simub_correct; eauto. intro ESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). @@ -962,11 +982,13 @@ Proof. eapply list_forall2_nth_error; eauto. - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. - assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). + assert (LENEQ': length (si_exits st1) = length (si_exits st2)) + by (eapply list_forall2_length; eauto). congruence. } destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. destruct SSEME as (is2 & SSEME2 & ISIMU). - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. + destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). unfold ssem_internal. exists is2. rewrite <- CONTEQ. rewrite ICONT. constructor; auto. @@ -974,11 +996,15 @@ Proof. + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). Qed. +Definition hsok ge sp rs m hst := hsok_internal ge sp rs m (hinternal hst). + Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := forall st1 st2, hsstate_refines hst1 st1 -> - hsstate_refines hst2 st2 -> sstate_simu f dm st1 st2 ctx. - + hsstate_refines hst2 st2 -> + hsok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> + hsok (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 -> + sstate_simu f dm st1 st2 ctx. Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := match ln with @@ -1215,14 +1241,16 @@ Proof. - destruct fv2; try discriminate. intro. explore. apply svos_simub_correct in EQ3. apply list_sval_simub_correct in EQ4. subst. apply revmap_check_single_correct in EQ. constructor; auto. - + admit. - + admit. + + destruct svos0; constructor; [| reflexivity ]. + erewrite <- seval_preserved; eauto. constructor. + + erewrite <- list_sval_eval_preserved; eauto. constructor. (* Stailcall *) - destruct fv2; try discriminate. intro. explore. apply svos_simub_correct in EQ0. apply list_sval_simub_correct in EQ1. subst. constructor; auto. - + admit. - + admit. + + destruct s2; constructor; [| reflexivity]. + erewrite <- seval_preserved; eauto. constructor. + + erewrite <- list_sval_eval_preserved; eauto. constructor. (* Sbuiltin *) - destruct fv2; try discriminate. intro. explore. apply revmap_check_single_correct in EQ1. apply list_builtin_arg_simub_correct in EQ2. @@ -1230,13 +1258,12 @@ Proof. (* Sjumptable *) - destruct fv2; try discriminate. intro. explore. apply revmap_check_list_correct in EQ. apply sval_simub_correct in EQ0. subst. - constructor; auto. - admit. + constructor; auto. erewrite <- seval_preserved; eauto. constructor. (* Sreturn *) - destruct fv2; try discriminate. destruct o; destruct o0; try discriminate. + intro. apply sval_simub_correct in H. subst. constructor; auto. + constructor; auto. -Admitted. +Qed. Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); @@ -1249,7 +1276,7 @@ Lemma hsstate_simub_correct dm f hst1 hst2 ctx: Proof. unfold hsstate_simub. intro. explore. eapply sfval_simub_correct in EQ1. eapply hsistate_simub_correct in EQ. - intros st1 st2 (SREF1 & LREF1 & FREF1) (SREF2 & LREF2 & FREF2). + intros st1 st2 (SREF1 & LREF1 & FREF1) (SREF2 & LREF2 & FREF2) HOK1 HOK2. constructor; [eauto |]. destruct SREF1 as (PCEQ1 & _ ). destruct SREF2 as (PCEQ2 & _ ). rewrite <- PCEQ1. rewrite <- PCEQ2. rewrite <- FREF1. rewrite <- FREF2. @@ -1284,7 +1311,8 @@ Proof. rewrite SEXEC1 in SEXEC; inversion SEXEC; subst. eexists; split; eauto. intros; eapply hsstate_simub_correct; eauto. -Qed. + all: admit. (* FIXME - le raffinement n'est prouvé correct que dans le cas où on ne fail pas *) +Admitted. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with -- cgit From 15245d51cbb09133e8b0e6ad1fc8e687bd8c523b Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 10 Jul 2020 08:48:07 +0200 Subject: try to fix refinement definition on exits --- kvx/lib/RTLpathSE_impl.v | 95 +++++++++++++++++++++++++++++++++++----------- kvx/lib/RTLpathSE_theory.v | 16 +++----- 2 files changed, 78 insertions(+), 33 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index f5285070..d24efdb5 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -77,6 +77,22 @@ Definition sok_local (ge: RTL.genv) (sp:val) (rs0: regset) (m0: mem) (st: sistat /\ seval_smem ge sp st.(si_smem) rs0 m0 <> None /\ forall (r: reg), seval_sval ge sp (si_sreg st r) rs0 m0 <> None. +Lemma ssem_local_sok ge sp rs0 m0 st rs m: + ssem_local ge sp st rs0 m0 rs m -> sok_local ge sp rs0 m0 st. +Proof. + unfold sok_local, ssem_local. + intuition congruence. +Qed. + +(* +Lemma sok_local_no_abort ge sp rs0 m0 st: + sok_local ge sp rs0 m0 st -> sabort_local ge sp st rs0 m0 -> False. +Proof. + unfold sok_local, sabort_local. + intros (H1 & H2 & H3) [A1 | [A2 | (r & A3)]]; intuition eauto. +Qed. +*) + Definition hsok_local ge sp rs0 m0 (hst: hsistate_local) : Prop := (forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None) /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). @@ -87,6 +103,12 @@ Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistat /\ (hsok_local ge sp rs0 m0 hst -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) /\ (hsok_local ge sp rs0 m0 hst -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). +Lemma ssem_local_refines_hok ge sp rs0 m0 hst st rs m: + ssem_local ge sp st rs0 m0 rs m -> hsistate_local_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst. +Proof. + intros H0 (H1 & _ & _). apply H1. eapply ssem_local_sok. eauto. +Qed. + (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) Record hsistate_exit := mk_hsistate_exit @@ -97,12 +119,13 @@ Record hsistate_exit := mk_hsistate_exit and a "dynamic" part -- that depends on it *) Definition hsistate_exit_refines_stat (hext: hsistate_exit) (ext: sistate_exit): Prop := - hsi_cond hext = si_cond ext - /\ hsi_scondargs hext = si_scondargs ext - /\ hsi_ifso hext = si_ifso ext. + hsi_ifso hext = si_ifso ext. Definition hsistate_exit_refines_dyn ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := - hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext). + hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) + /\ (hsok_local ge sp rs0 m0 (hsi_elocal hext) -> + seval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem_get (hsi_elocal hext)) rs0 m0 + = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0). (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. @@ -391,6 +414,15 @@ Proof. auto. Qed. +Lemma seval_condition_refines hst st ge sp cond args rs m: + hsok_local ge sp rs m hst -> + hsistate_local_refines ge sp rs m hst st -> + seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. + Proof. + intros HOK (OKEQ & MEMEQ & RSEQ). unfold seval_condition. + rewrite <- MEMEQ; auto. rewrite hsi_smem_eval_correct. reflexivity. +Qed. + Lemma hsiexec_inst_correct_None i hst st: hsiexec_inst i hst = None -> siexec_inst i st = None. Proof. @@ -606,6 +638,12 @@ Qed. (* TODO - generalize it with a list of registers to test ? *) Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. +(* TODO: voilà les conditions que hsilocal_simub doit garantir *) +Definition hsilocal_simu_core (hst1 hst2: hsistate_local):= + forall sm, List.In sm (hsi_lsmem hst2) -> List.In sm (hsi_lsmem hst1) + /\ forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1) + /\ forall r, hsi_sreg_get hst2 r = hsi_sreg_get hst1 r. + Definition hsilocal_simu dm f (hst1 hst2: hsistate_local) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> @@ -767,14 +805,17 @@ Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := fora hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> siexit_simu dm f ctx se1 se2. - +(* Lemma seval_condition_refines hst st ge sp cond args rs m: + hsok_local ge sp rs0 m0 (hsi_elocal hext) -> hsistate_local_refines ge sp rs m hst st -> seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. Proof. -Admitted. + intros H & (_ & DEF). intuition. +Qed. +*) -Local Hint Resolve genv_match: core. +Local Hint Resolve genv_match ssem_local_refines_hok: core. Lemma hsiexit_simub_correct dm f hse1 hse2: hsiexit_simub dm f hse1 hse2 = OK tt -> @@ -785,31 +826,33 @@ Proof. intros (SREGEQ & SMEMEQ) ctx. eapply hsilocal_simub_correct in EQ2. apply revmap_check_single_correct in EQ3. - intros se1 se2 (HCOND1 & HARGS1 & HIFSO1) (HCOND2 & HARGS2 & HIFSO2) HLREF1 HLREF2. - unfold hsistate_exit_refines_dyn in *. + intros se1 se2 HIFSO1 HIFSO2 (HLREF1 & HCOND1) (HLREF2 & HCOND2). + unfold hsistate_exit_refines_stat in *. + intros is1 (SCOND & SLOCAL & SIFSO). constructor. - - unfold siexit_simu_ssem. intros is SEMEXIT. - destruct SEMEXIT as (SCOND & SLOCAL & SIFSO). - let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). + - rewrite <- HCOND1; eauto. + rewrite <- HCOND2; eauto. + rewrite SMEMEQ, EQ0, e. + eapply seval_condition_preserved; eauto. + assert (X: hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse1) -> + hsok_local (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse2)). + + admit. (* TODO: on doit savoir ça du fait que les états symboliques locaux se raffinent *) + + eauto. + - let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). destruct is2 as [icont ipc irs imem]. simpl in *. exists (mk_istate icont (si_ifso se2) irs imem). simpl. constructor; auto. + repeat (constructor; auto). - rewrite <- HARGS2. rewrite <- EQ0. - rewrite <- HCOND2. rewrite <- e. - erewrite <- seval_condition_refines; eauto. rewrite <- SMEMEQ. + rewrite <- HCOND2; eauto. + rewrite <- e, <- EQ0, <- SMEMEQ. erewrite seval_condition_preserved; eauto. - rewrite HARGS1. rewrite HCOND1. erewrite seval_condition_refines; eauto. - + unfold istate_simu in *. destruct (icontinue is) eqn:ICONT. + rewrite HCOND1; eauto. + + unfold istate_simu in *. destruct (icontinue is1) eqn:ICONT. * destruct ISIMU as (A & B & C). constructor; auto. * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. repeat (constructor; auto). simpl in *. congruence. - - rewrite <- HARGS2. rewrite <- HCOND2. erewrite <- seval_condition_refines; eauto. - rewrite <- HARGS1. rewrite <- HCOND1. erewrite <- seval_condition_refines; eauto. - rewrite SMEMEQ, EQ0, e. - eapply seval_condition_preserved; eauto. -Qed. +Admitted. Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := match lhse1 with @@ -869,10 +912,13 @@ Proof. induction 1; [unfold all_fallthrough; contradiction|]. intros X ext INEXT. eapply all_fallthrough_revcons in X. destruct X as (SEVAL & ALLFU). apply IHlist_forall2 in ALLFU. - destruct H as (_ & CONDSIMU). +Admitted. +(* BROKEN! + destruct H as (_ & CONDSIMU). inv INEXT; [|eauto]. erewrite <- CONDSIMU; eauto. Qed. +*) Lemma siexits_simu_all_fallthrough_upto dm f ctx lse1 lse2: siexits_simu dm f lse1 lse2 ctx -> forall ext1 lx1, @@ -964,6 +1010,8 @@ Proof. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). congruence. } +Admitted. +(* destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. destruct SSEME as (is2 & SSEME2 & ISIMU). unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). @@ -973,6 +1021,7 @@ Proof. + eexists; eexists; constructor; eauto. + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). Qed. +*) Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := forall st1 st2, diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index e81a107e..73d7fa2c 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1761,17 +1761,13 @@ Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sista exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) sl2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) /\ istate_simu f dm is1 is2. -Definition siexit_simu_ssem (dm: PTree.t node) (f: RTLpath.function) (se1 se2: sistate_exit) (ctx: simu_proof_context f):= - forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> - exists is2, - ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) - /\ istate_simu f dm is1 is2. - Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (ctx: simu_proof_context f) (se1 se2: sistate_exit) := - siexit_simu_ssem dm f se1 se2 ctx - /\ - seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx) - = seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx). + forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> + (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx)) + = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)) + /\ exists is2, + ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) + /\ istate_simu f dm is1 is2. Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) (ctx: simu_proof_context f) := list_forall2 (siexit_simu dm f ctx) lse1 lse2. -- cgit From 5aa7b86c2d78a005057ebd02125ef76fa62c27e7 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 10 Jul 2020 09:10:06 +0200 Subject: remove junk comment --- kvx/lib/RTLpathSE_impl.v | 10 ---------- 1 file changed, 10 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index d24efdb5..7c8b16f7 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -805,16 +805,6 @@ Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := fora hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> siexit_simu dm f ctx se1 se2. -(* -Lemma seval_condition_refines hst st ge sp cond args rs m: - hsok_local ge sp rs0 m0 (hsi_elocal hext) -> - hsistate_local_refines ge sp rs m hst st -> - seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. -Proof. - intros H & (_ & DEF). intuition. -Qed. -*) - Local Hint Resolve genv_match ssem_local_refines_hok: core. Lemma hsiexit_simub_correct dm f hse1 hse2: -- cgit From 0241c07f9826a628940c705312093d85e8819f92 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 10 Jul 2020 09:15:49 +0200 Subject: fix comment --- kvx/lib/RTLpathSE_impl.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 7c8b16f7..7f7801c0 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -826,7 +826,7 @@ Proof. eapply seval_condition_preserved; eauto. assert (X: hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse1) -> hsok_local (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse2)). - + admit. (* TODO: on doit savoir ça du fait que les états symboliques locaux se raffinent *) + + admit. (* TODO: on doit savoir ça du fait que l'état symbolique de hse2 simule celui de hse1 *) + eauto. - let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). -- cgit From dc077ee8cce7612c4ed9dacf5ceb2108131544ab Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 10 Jul 2020 16:00:41 +0200 Subject: Lemmas on hsilocal_simu ---> relies on reflexivity of istate_simu (which isn't true) --- kvx/lib/RTLpathSE_impl.v | 71 +++++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 64 insertions(+), 7 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 7f7801c0..9ec104a0 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -14,9 +14,7 @@ Local Open Scope option_monad_scope. (** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) -(** * Implementation of local symbolic internal states -TODO: use refined symbolic values instead of those from RTLpathSE_theory. -*) +(** * Implementation of local symbolic internal states *) (** name : Hash-consed Symbolic Internal state local. Later on we will use the refinement to introduce hash consing *) @@ -109,12 +107,71 @@ Proof. intros H0 (H1 & _ & _). apply H1. eapply ssem_local_sok. eauto. Qed. +Definition hsilocal_simu_core (hst1 hst2: hsistate_local) := + (forall sm, List.In sm (hsi_lsmem hst2) -> List.In sm (hsi_lsmem hst1)) + /\ (forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1)) + /\ (forall r, hsi_sreg_get hst2 r = hsi_sreg_get hst1 r) + /\ hsi_smem_get hst1 = hsi_smem_get hst2. + +Definition hsilocal_simu_coreb (hst1 hst2: hsistate_local) := OK tt. + +Theorem hsilocal_simu_coreb_correct hst1 hst2: + hsilocal_simu_coreb hst1 hst2 = OK tt -> + hsilocal_simu_core hst1 hst2. +Proof. +Admitted. + +Lemma hsilocal_simu_core_nofail ge1 ge2 sp rs0 m0 hst1 hst2: + hsilocal_simu_core hst1 hst2 -> + (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> + hsok_local ge1 sp rs0 m0 hst1 -> + hsok_local ge2 sp rs0 m0 hst2. +Proof. + intros (MEMOK & RSOK & _ & _) GFS (OKV & OKM). constructor. + - intros sv INS. apply RSOK in INS. apply OKV in INS. erewrite seval_preserved; eauto. + - intros sm INS. apply MEMOK in INS. apply OKM in INS. erewrite smem_eval_preserved; eauto. +Qed. + +Definition hsilocal_simu dm f ctx hst1 hst2 := forall st1 st2, + hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> + hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + silocal_simu dm f st1 st2 ctx. + +Remark istate_simu_reflexive f dm is: istate_simu f dm is is. +Proof. + unfold istate_simu. destruct (icontinue is). + - repeat (constructor; auto). + - eexists. constructor; [|constructor]. all: admit. +Admitted. (* Problem : this remark is not true *) + +Lemma hsilocal_simu_core_correct hst1 hst2 f ctx dm: + hsilocal_simu_core hst1 hst2 -> + hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> + hsilocal_simu dm f ctx hst1 hst2. +Proof. + intros CORE HOK1 st1 st2 HREF1 HREF2. + exploit hsilocal_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. + intros is1 (PRE & MEMEQ & RSEQ). + eexists. constructor. + - constructor; [|constructor]. + + destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. + + destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _). + destruct CORE as (_ & _ & _ & MEMEQ3). + rewrite <- MEMEQ2; auto. rewrite hsi_smem_eval_correct. rewrite <- MEMEQ3. + erewrite smem_eval_preserved; [| eapply ctx]. + rewrite <- hsi_smem_eval_correct. rewrite MEMEQ1; auto. eassumption. + + intro r. destruct HREF2 as (_ & _ & A). destruct HREF1 as (_ & _ & B). + destruct CORE as (_ & _ & C & _). rewrite <- A; auto. rewrite hsi_sreg_eval_correct. + rewrite C. erewrite seval_preserved; [| eapply ctx]. rewrite <- hsi_sreg_eval_correct. + rewrite B; auto. + - apply istate_simu_reflexive. +Qed. + (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. - (** NB: we split the refinement relation between a "static" part -- independendent of the initial context and a "dynamic" part -- that depends on it *) @@ -638,7 +695,7 @@ Qed. (* TODO - generalize it with a list of registers to test ? *) Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. -(* TODO: voilà les conditions que hsilocal_simub doit garantir *) +(* (* TODO: voilà les conditions que hsilocal_simub doit garantir *) Definition hsilocal_simu_core (hst1 hst2: hsistate_local):= forall sm, List.In sm (hsi_lsmem hst2) -> List.In sm (hsi_lsmem hst1) /\ forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1) @@ -647,11 +704,11 @@ Definition hsilocal_simu_core (hst1 hst2: hsistate_local):= Definition hsilocal_simu dm f (hst1 hst2: hsistate_local) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - silocal_simu dm f st1 st2 ctx. + silocal_simu dm f st1 st2 ctx. *) Lemma hsilocal_simub_correct dm f hst1 hst2: hsilocal_simub dm f hst1 hst2 = OK tt -> - forall ctx, hsilocal_simu dm f hst1 hst2 ctx. + forall ctx, hsilocal_simu dm f ctx hst1 hst2. Proof. Admitted. -- cgit From 2ff09832c0c3a2c50d51ec90566ad74e093ab3da Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 17:25:44 +0200 Subject: trapping loads are irreversible --- kvx/lib/PrepassSchedulingOracle.ml | 1 + kvx/lib/RTLpathScheduleraux.ml | 2 +- test/monniaux/rules.mk | 2 +- 3 files changed, 3 insertions(+), 2 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 5a48924e..c1a3804f 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -135,6 +135,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> + (if trap=TRAP then irreversible_action i); add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 11539994..12637a9d 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -112,7 +112,7 @@ let schedule_superblock sb code = let nr_instr = Array.length sb.instructions in let trailer_length = match PTree.get (sb.instructions.(nr_instr-1)) code with - | Some (Ireturn _ | Itailcall _ | Ijumptable _) -> 1 + | Some (Ireturn _ | Itailcall _ | Ijumptable _ | Icall _) -> 1 | _ -> 0 in match PrepassSchedulingOracle.schedule_sequence (Array.map (fun i -> diff --git a/test/monniaux/rules.mk b/test/monniaux/rules.mk index 95d57f80..8549e367 100644 --- a/test/monniaux/rules.mk +++ b/test/monniaux/rules.mk @@ -21,7 +21,7 @@ MEASURES?=time ALL_CFLAGS+=-Wall -D__KVX_COS__ -DMAX_MEASURES=$(MAX_MEASURES) #ALL_CFLAGS+=-g ALL_GCCFLAGS+=$(ALL_CFLAGS) -std=c99 -Wextra -Werror=implicit -ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fprepass +ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fduplicate 2 -fprepass -fallnontrap # The compilers KVX_CC?=kvx-cos-gcc -- cgit From 36d6c732567a10a893b502ab86f6b438fa5e0a8a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 17:27:38 +0200 Subject: trapping ops --- kvx/lib/PrepassSchedulingOracle.ml | 1 + 1 file changed, 1 insertion(+) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index c1a3804f..6169bf16 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -132,6 +132,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> + (if Op.is_trapping_op op then irreversible_action i); add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> -- cgit From fe7e30495c1690068e1df8aded85404299aaf329 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 18:21:22 +0200 Subject: proper ordering on calls etc. ? --- kvx/lib/PrepassSchedulingOracle.ml | 8 ++++---- test/monniaux/rules.mk | 2 +- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 6169bf16..b83078f8 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -132,21 +132,21 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> - (if Op.is_trapping_op op then irreversible_action i); + (if Op.is_trapping_op op then set_branch i); add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> - (if trap=TRAP then irreversible_action i); + (if trap=TRAP then set_branch i); add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output | Istore(chunk, addressing, addr_regs, input, _) -> - irreversible_action i; + set_branch i; add_input_regs i addr_regs; add_input_reg i input; add_output_mem i | Icall(signature, ef, inputs, output, _) -> - irreversible_action i; + set_branch i; (match ef with | Datatypes.Coq_inl r -> add_input_reg i r | Datatypes.Coq_inr symbol -> () diff --git a/test/monniaux/rules.mk b/test/monniaux/rules.mk index 8549e367..63582f96 100644 --- a/test/monniaux/rules.mk +++ b/test/monniaux/rules.mk @@ -21,7 +21,7 @@ MEASURES?=time ALL_CFLAGS+=-Wall -D__KVX_COS__ -DMAX_MEASURES=$(MAX_MEASURES) #ALL_CFLAGS+=-g ALL_GCCFLAGS+=$(ALL_CFLAGS) -std=c99 -Wextra -Werror=implicit -ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fduplicate 2 -fprepass -fallnontrap +ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fduplicate 2 -fprepass -fall-loads-nontrap # The compilers KVX_CC?=kvx-cos-gcc -- cgit From b44d87ef03279cb479101d7b02031cf78136062d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 19:25:47 +0200 Subject: oracle super restrictif --- kvx/lib/PrepassSchedulingOracle.ml | 24 ++++++++++++++++-------- 1 file changed, 16 insertions(+), 8 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index b83078f8..322dff85 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -132,11 +132,13 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> - (if Op.is_trapping_op op then set_branch i); + (* (if Op.is_trapping_op op then set_branch i); *) + set_branch i; add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> - (if trap=TRAP then set_branch i); + set_branch i; + (* (if trap=TRAP then set_branch i); *) add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output @@ -154,7 +156,8 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_input_mem i; add_input_regs i inputs; add_output_reg i (latency_of_call signature ef) output; - add_output_mem i + add_output_mem i; + failwith "Icall" | Itailcall(signature, ef, inputs) -> set_branch i; (match ef with @@ -162,24 +165,29 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = | Datatypes.Coq_inr symbol -> () ); add_input_mem i; - add_input_regs i inputs + add_input_regs i inputs; + failwith "Itailcall" | Ibuiltin(ef, builtin_inputs, builtin_output, _) -> set_branch i; add_input_mem i; List.iter (add_builtin_arg i) builtin_inputs; add_builtin_res i builtin_output; - add_output_mem i + add_output_mem i; + failwith "Ibuiltin" | Icond(cond, inputs, _, _, _) -> set_branch i; add_input_regs i inputs | Ijumptable(input, _) -> set_branch i; - add_input_reg i input + add_input_reg i input; + failwith "Ijumptable" | Ireturn(Some input) -> set_branch i; - add_input_reg i input + add_input_reg i input; + failwith "Ireturn" | Ireturn(None) -> - set_branch i + set_branch i; + failwith "Ireturn none" end seqa; !latency_constraints;; -- cgit From cf94e6e6ee08732f3914174d12cb46e67e2adab6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 19:36:45 +0200 Subject: begin relaxing --- kvx/lib/PrepassSchedulingOracle.ml | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 322dff85..7e5e2f4a 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -125,15 +125,13 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = begin fun i (insn, other_uses) -> List.iter (fun use -> - print_endline "add other use"; add_input_reg i use) (Regset.elements other_uses); match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> - (* (if Op.is_trapping_op op then set_branch i); *) - set_branch i; + (if Op.is_trapping_op op then set_branch i); add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> -- cgit From b38bf5750f0eb876d12da5d424b548abf4160484 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 21:02:17 +0200 Subject: it works but is too constrained --- kvx/lib/PrepassSchedulingOracle.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 7e5e2f4a..dce7686b 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -136,7 +136,6 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> set_branch i; - (* (if trap=TRAP then set_branch i); *) add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output @@ -174,6 +173,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = failwith "Ibuiltin" | Icond(cond, inputs, _, _, _) -> set_branch i; + add_input_mem i; add_input_regs i inputs | Ijumptable(input, _) -> set_branch i; -- cgit From 0392f0a7ef44ddbe6c5a75fc5b89a83e9add035c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 21:27:22 +0200 Subject: relaxing... --- kvx/lib/PrepassSchedulingOracle.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index dce7686b..cd8f321f 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -131,7 +131,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = match insn with | Inop _ -> () | Iop(op, inputs, output, _) -> - (if Op.is_trapping_op op then set_branch i); + (if Op.is_trapping_op op then irreversible_action i); add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> @@ -140,7 +140,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output | Istore(chunk, addressing, addr_regs, input, _) -> - set_branch i; + irreversible_action i; add_input_regs i addr_regs; add_input_reg i input; add_output_mem i -- cgit From 7abe728fea66e04db905df7fc1899904690560af Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 21:39:37 +0200 Subject: relaxing --- kvx/lib/PrepassSchedulingOracle.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index cd8f321f..4d70d0aa 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -135,7 +135,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> - set_branch i; + irreversible_action i; add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output -- cgit From ba39e941a47a87ecc67795f955cc24e5c1266428 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 10 Jul 2020 22:42:27 +0200 Subject: fix the last instruction detection code --- kvx/lib/RTLpathScheduleraux.ml | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 12637a9d..1b8ea20a 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -112,8 +112,11 @@ let schedule_superblock sb code = let nr_instr = Array.length sb.instructions in let trailer_length = match PTree.get (sb.instructions.(nr_instr-1)) code with - | Some (Ireturn _ | Itailcall _ | Ijumptable _ | Icall _) -> 1 - | _ -> 0 in + | None -> 0 + | Some ii -> + match predicted_successor ii with + | Some _ -> 0 + | None -> 1 in match PrepassSchedulingOracle.schedule_sequence (Array.map (fun i -> (match PTree.get i code with -- cgit From a8e16ac32d1cb4262fa131337c4cd9f443c09cde Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 11 Jul 2020 09:26:48 +0200 Subject: fix hsilocal_simu_core_correct --- kvx/lib/RTLpathSE_impl.v | 38 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 9ec104a0..a6f63237 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -132,11 +132,47 @@ Proof. - intros sm INS. apply MEMOK in INS. apply OKM in INS. erewrite smem_eval_preserved; eauto. Qed. +Remark istate_simulive_reflexive dm is: istate_simulive (fun _ : Regset.elt => True) dm is is. +Proof. + unfold istate_simulive. + repeat (constructor; auto). +Qed. + +Lemma hsilocal_simu_core_correct dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: + hsilocal_simu_core hst1 hst2 -> + hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> + hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> + exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) + /\ istate_simulive (fun _ => True) dm is1 is2. +Proof. + intros CORE HREF1 HREF2 is1 SEML. + exploit ssem_local_refines_hok; eauto. + intros HOK1. + exploit hsilocal_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. + destruct SEML as (PRE & MEMEQ & RSEQ). + eexists. constructor. + - constructor; [|constructor]. + + destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. + + destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _). + destruct CORE as (_ & _ & _ & MEMEQ3). + rewrite <- MEMEQ2; auto. rewrite hsi_smem_eval_correct. rewrite <- MEMEQ3. + erewrite smem_eval_preserved; [| eapply ctx]. + rewrite <- hsi_smem_eval_correct. rewrite MEMEQ1; auto. eassumption. + + intro r. destruct HREF2 as (_ & _ & A). destruct HREF1 as (_ & _ & B). + destruct CORE as (_ & _ & C & _). rewrite <- A; auto. rewrite hsi_sreg_eval_correct. + rewrite C. erewrite seval_preserved; [| eapply ctx]. rewrite <- hsi_sreg_eval_correct. + rewrite B; auto. + - apply istate_simulive_reflexive. +Qed. + + Definition hsilocal_simu dm f ctx hst1 hst2 := forall st1 st2, hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> silocal_simu dm f st1 st2 ctx. +(* BELOW: this is wrong: il faut ajouter des conditions supplémentaire vis-à-vis de hsilocal_simu_core Remark istate_simu_reflexive f dm is: istate_simu f dm is is. Proof. unfold istate_simu. destruct (icontinue is). @@ -144,6 +180,7 @@ Proof. - eexists. constructor; [|constructor]. all: admit. Admitted. (* Problem : this remark is not true *) + Lemma hsilocal_simu_core_correct hst1 hst2 f ctx dm: hsilocal_simu_core hst1 hst2 -> hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> @@ -166,6 +203,7 @@ Proof. rewrite B; auto. - apply istate_simu_reflexive. Qed. +*) (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) -- cgit From 291b7bd92b510f9dd2edabcae49d13f8c7466c25 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 11 Jul 2020 10:28:55 +0200 Subject: found another bug --- kvx/lib/PrepassSchedulingOracle.ml | 4 ++-- kvx/lib/RTLpathScheduleraux.ml | 1 + 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 4d70d0aa..9701f446 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -135,7 +135,7 @@ let get_simple_dependencies (seqa : (instruction*Regset.t) array) = add_input_regs i inputs; add_output_reg i (latency_of_op op (List.length inputs)) output | Iload(trap, chunk, addressing, addr_regs, output, _) -> - irreversible_action i; + (if trap=TRAP then irreversible_action i); add_input_mem i; add_input_regs i addr_regs; add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output @@ -414,7 +414,7 @@ let schedule_sequence (seqa : (instruction*Regset.t) array) = let problem = define_problem seqa in print_sequence stdout (Array.map fst seqa); print_problem stdout problem; - match reverse_list_scheduler problem + match (*reverse_list_scheduler*) list_scheduler problem (* scheduler_by_name !Clflags.option_fprepass_sched problem *) with | None -> Printf.printf "no solution in prepass scheduling\n"; None diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 1b8ea20a..7a43fed5 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -167,6 +167,7 @@ let change_successors i = function | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) | Ijumptable (a,[n]) -> Ijumptable (a,[s]) + | Icond _ -> failwith "Icond Wrong instruction (2)" | _ -> failwith "Wrong instruction (2)") | [s1; s2] -> ( match i with -- cgit From eea5640e55890538fa43c3d5672853a0ae015b9c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 11 Jul 2020 11:00:15 +0200 Subject: command line selection of prepass scheduler --- driver/Clflags.ml | 1 + driver/Driver.ml | 3 +++ kvx/InstructionScheduler.ml | 7 +++++++ kvx/InstructionScheduler.mli | 7 +++++-- kvx/PostpassSchedulingOracle.ml | 9 +-------- kvx/lib/PrepassSchedulingOracle.ml | 3 +-- 6 files changed, 18 insertions(+), 12 deletions(-) diff --git a/driver/Clflags.ml b/driver/Clflags.ml index fc5d9e45..829af76a 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -37,6 +37,7 @@ let option_fduplicate = ref (-1) let option_finvertcond = ref true let option_ftracelinearize = ref false let option_fprepass = ref false +let option_fprepass_sched = ref "list" let option_fpostpass = ref true let option_fpostpass_sched = ref "list" let option_fifconversion = ref true diff --git a/driver/Driver.ml b/driver/Driver.ml index 6accc22b..fef9c166 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -205,6 +205,8 @@ Processing options: -fmove-loop-invariants Perform loop-invariant code motion [off] -fredundancy Perform redundancy elimination [on] -fprepass Perform prepass scheduling (only for K1 architecture) [off] + -fprepass= Perform postpass scheduling with the specified optimization [list] + (=list: list scheduling, =ilp: ILP, =greedy: just packing bundles) -fpostpass Perform postpass scheduling (only for K1 architecture) [on] -fpostpass= Perform postpass scheduling with the specified optimization [list] (=list: list scheduling, =ilp: ILP, =greedy: just packing bundles) @@ -425,6 +427,7 @@ let cmdline_actions = @ [ Exact "-fduplicate", Integer (fun n -> option_fduplicate := n) ] @ f_opt "invertcond" option_finvertcond @ f_opt "tracelinearize" option_ftracelinearize + @ f_opt_str "prepass" option_fprepass option_fprepass_sched @ f_opt_str "postpass" option_fpostpass option_fpostpass_sched @ f_opt "inline" option_finline @ f_opt "inline-functions-called-once" option_finline_functions_called_once diff --git a/kvx/InstructionScheduler.ml b/kvx/InstructionScheduler.ml index e4dc3f97..32f394b1 100644 --- a/kvx/InstructionScheduler.ml +++ b/kvx/InstructionScheduler.ml @@ -1245,3 +1245,10 @@ let cascaded_scheduler (problem : problem) = end; Some solution;; +let scheduler_by_name name = + match name with + | "ilp" -> validated_scheduler cascaded_scheduler + | "list" -> validated_scheduler list_scheduler + | "revlist" -> validated_scheduler reverse_list_scheduler + | "greedy" -> greedy_scheduler + | s -> failwith ("unknown scheduler: " ^ s);; diff --git a/kvx/InstructionScheduler.mli b/kvx/InstructionScheduler.mli index f91c2d06..6d27b30c 100644 --- a/kvx/InstructionScheduler.mli +++ b/kvx/InstructionScheduler.mli @@ -65,10 +65,10 @@ val list_scheduler : problem -> solution option (** Schedule the problem using the order of instructions without any reordering *) val greedy_scheduler : problem -> solution option -(** Schedule a problem using a scheduler applied in the opposite direction, e.g. for list scheduling from the end instead of the start. BUGGY *) +(** Schedule a problem using a scheduler applied in the opposite direction, e.g. for list scheduling from the end instead of the start. *) val schedule_reversed : scheduler -> problem -> int array option -(** Schedule a problem from the end using a list scheduler. BUGGY *) +(** Schedule a problem from the end using a list scheduler. *) val reverse_list_scheduler : problem -> int array option (** Check that a problem is well-formed. @@ -108,3 +108,6 @@ val smt_print_problem : out_channel -> problem -> unit;; val ilp_print_problem : out_channel -> problem -> pseudo_boolean_problem_type -> pseudo_boolean_mapper;; val ilp_scheduler : pseudo_boolean_problem_type -> problem -> solution option;; + +(** Schedule a problem using a scheduler given by a string name *) +val scheduler_by_name : string -> problem -> int array option;; diff --git a/kvx/PostpassSchedulingOracle.ml b/kvx/PostpassSchedulingOracle.ml index 822c0dc0..67e3f80f 100644 --- a/kvx/PostpassSchedulingOracle.ml +++ b/kvx/PostpassSchedulingOracle.ml @@ -916,14 +916,7 @@ let print_bb oc bb = let do_schedule bb = let problem = build_problem bb - in let solution = (if !Clflags.option_fpostpass_sched = "ilp" then - validated_scheduler cascaded_scheduler - else if !Clflags.option_fpostpass_sched = "list" then - validated_scheduler list_scheduler - else if !Clflags.option_fpostpass_sched = "revlist" then - validated_scheduler reverse_list_scheduler - else if !Clflags.option_fpostpass_sched = "greedy" then - greedy_scheduler else failwith ("Invalid scheduler:" ^ !Clflags.option_fpostpass_sched)) problem + in let solution = scheduler_by_name (!Clflags.option_fpostpass_sched) problem in match solution with | None -> failwith "Could not find a valid schedule" | Some sol -> let bundles = bundlize_solution bb sol in diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml index 9701f446..78961310 100644 --- a/kvx/lib/PrepassSchedulingOracle.ml +++ b/kvx/lib/PrepassSchedulingOracle.ml @@ -414,8 +414,7 @@ let schedule_sequence (seqa : (instruction*Regset.t) array) = let problem = define_problem seqa in print_sequence stdout (Array.map fst seqa); print_problem stdout problem; - match (*reverse_list_scheduler*) list_scheduler problem - (* scheduler_by_name !Clflags.option_fprepass_sched problem *) with + match scheduler_by_name (!Clflags.option_fprepass_sched) problem with | None -> Printf.printf "no solution in prepass scheduling\n"; None | Some solution -> -- cgit From 5b756533e3fbd51b8b824e9ed87d2be3feff8f36 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 13 Jul 2020 14:29:19 +0200 Subject: Fix Icond bug --- kvx/lib/RTLpathScheduleraux.ml | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 7a43fed5..08cda134 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -167,7 +167,12 @@ let change_successors i = function | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) | Ijumptable (a,[n]) -> Ijumptable (a,[s]) - | Icond _ -> failwith "Icond Wrong instruction (2)" + | Icond (a,b,n1,n2,p) -> ( + match p with + | Some true -> Icond (a,b,s,n2,p) + | Some false -> Icond (a,b,n1,s,p) + | None -> failwith "Icond Wrong instruction (2) ; should not happen?" + ) | _ -> failwith "Wrong instruction (2)") | [s1; s2] -> ( match i with -- cgit From ba42f656e78e9ef7d3de699bd8a8e9a032de9b60 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 13 Jul 2020 18:00:40 +0200 Subject: Fix switching basic instruction with Icond --- kvx/lib/RTLpathScheduleraux.ml | 16 +++++++++++++--- 1 file changed, 13 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 08cda134..6c1d4b11 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -153,7 +153,7 @@ let schedule_superblock sb code = end *) (* stub: identity function *) -let change_successors i = function +let rec change_successors i = function | [] -> ( match i with | Itailcall _ | Ireturn _ -> i @@ -178,7 +178,7 @@ let change_successors i = function match i with | Icond (a,b,n1,n2,p) -> Icond (a,b,s1,s2,p) | Ijumptable (a, [n1; n2]) -> Ijumptable (a, [s1; s2]) - | _ -> failwith "Wrong instruction (3)") + | _ -> change_successors i [s1]) | ls -> ( match i with | Ijumptable (a, ln) -> begin @@ -221,11 +221,21 @@ let check_order code old_order new_order = begin | _ -> () end +let successors_inst_with_prediction = function +| Icond (_,_,n1,n2,p) -> ( + match p with + | Some true -> [n1; n2] + | Some false -> [n2; n1] + | None -> [n1; n2] + ) +| i -> successors_inst i + let apply_schedule code sb new_order = let tc = ref code in let old_order = sb.instructions in let last_node = Array.get old_order (Array.length old_order - 1) in - let last_successors = successors_inst @@ get_some @@ PTree.get last_node code in + let last_successors = successors_inst_with_prediction + @@ get_some @@ PTree.get last_node code in begin check_order code old_order new_order; Array.iteri (fun i n' -> -- cgit From 9d0b7d4c1728b0a07c89ddd0a77e5385a183d6ff Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 15 Jul 2020 16:40:08 +0200 Subject: More debug info --- kvx/lib/RTLpathScheduleraux.ml | 25 +++++++++++++++++-------- 1 file changed, 17 insertions(+), 8 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 6c1d4b11..1e15effe 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -221,20 +221,21 @@ let check_order code old_order new_order = begin | _ -> () end -let successors_inst_with_prediction = function +let main_successors = function +| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n] +| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] | Icond (_,_,n1,n2,p) -> ( match p with | Some true -> [n1; n2] - | Some false -> [n2; n1] - | None -> [n1; n2] - ) -| i -> successors_inst i + | Some false -> [n2; n1] + | None -> [n1; n2] ) +| Ijumptable _ | Itailcall _ | Ireturn _ -> [] let apply_schedule code sb new_order = let tc = ref code in let old_order = sb.instructions in let last_node = Array.get old_order (Array.length old_order - 1) in - let last_successors = successors_inst_with_prediction + let last_successors = main_successors @@ get_some @@ PTree.get last_node code in begin check_order code old_order new_order; @@ -274,5 +275,13 @@ let scheduler f = let pm = f.fn_path in let typing = get_ok @@ RTLtyping.type_function f.fn_RTL in let lsb = get_superblocks code entry pm typing in - let tc = do_schedule code lsb in - (((tc, entry), pm), id_ptree) + begin + (* debug_flag := true; *) + dprintf "Pathmap:\n"; dprintf "\n"; + print_path_map pm; + dprintf "Superblocks:\n"; + print_superblocks lsb code; dprintf "\n"; + (* debug_flag := false; *) + let tc = do_schedule code lsb in + (((tc, entry), pm), id_ptree) + end -- cgit From 8cdafe0117cd13b347db4ee3910c2a61a454bc73 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 15 Jul 2020 17:34:46 +0200 Subject: Fixed last instruction not having liveins --- kvx/lib/RTLpathScheduleraux.ml | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index 1e15effe..ac7238db 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -69,8 +69,14 @@ let get_superblocks code entry pm typing = let liveins = ref (PTree.empty) in let rec follow pc n = let inst = get_some @@ PTree.get pc code in - if (n == 0) then ([pc], successors_inst inst) - else + if (n == 0) then begin + (match (non_predicted_successors inst) with + | [pcout] -> + let live = (get_some @@ PTree.get pcout pm).input_regs in + liveins := PTree.set pc live !liveins + | _ -> ()); + ([pc], successors_inst inst) + end else let nexts_from_exit = match (non_predicted_successors inst) with | [pcout] -> let live = (get_some @@ PTree.get pcout pm).input_regs in begin -- cgit From 509015c29fe083b1e85b326b3dfa71d82636c9e2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 15 Jul 2020 20:27:38 +0200 Subject: flags --- test/monniaux/rules.mk | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/rules.mk b/test/monniaux/rules.mk index 63582f96..a5754574 100644 --- a/test/monniaux/rules.mk +++ b/test/monniaux/rules.mk @@ -21,7 +21,7 @@ MEASURES?=time ALL_CFLAGS+=-Wall -D__KVX_COS__ -DMAX_MEASURES=$(MAX_MEASURES) #ALL_CFLAGS+=-g ALL_GCCFLAGS+=$(ALL_CFLAGS) -std=c99 -Wextra -Werror=implicit -ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fduplicate 2 -fprepass -fall-loads-nontrap +ALL_CCOMPFLAGS+=$(ALL_CFLAGS) -fduplicate 2 -fprepass -fprepass= list -fall-loads-nontrap # The compilers KVX_CC?=kvx-cos-gcc -- cgit From dd36a11eea8a9867f05cc25ad6ab2df55197162b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 16 Jul 2020 16:05:56 +0200 Subject: Progress --- kvx/lib/RTLpathSE_impl.v | 258 +++++++++++++++++++++++++++++++++++---------- kvx/lib/RTLpathSE_theory.v | 4 +- 2 files changed, 207 insertions(+), 55 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index a6f63237..20d7fe3d 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -96,13 +96,13 @@ Definition hsok_local ge sp rs0 m0 (hst: hsistate_local) : Prop := /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). (* refinement link between a (st: sistate_local) and (hst: hsistate_local) *) -Definition hsistate_local_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistate_local) := +Definition hsilocal_refines ge sp rs0 m0 (hst: hsistate_local) (st: sistate_local) := (sok_local ge sp rs0 m0 st <-> hsok_local ge sp rs0 m0 hst) /\ (hsok_local ge sp rs0 m0 hst -> hsi_smem_eval ge sp hst rs0 m0 = seval_smem ge sp st.(si_smem) rs0 m0) /\ (hsok_local ge sp rs0 m0 hst -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0). Lemma ssem_local_refines_hok ge sp rs0 m0 hst st rs m: - ssem_local ge sp st rs0 m0 rs m -> hsistate_local_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst. + ssem_local ge sp st rs0 m0 rs m -> hsilocal_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst. Proof. intros H0 (H1 & _ & _). apply H1. eapply ssem_local_sok. eauto. Qed. @@ -138,10 +138,10 @@ Proof. repeat (constructor; auto). Qed. -Lemma hsilocal_simu_core_correct dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: +Lemma hsilocal_simu_core_correct' dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: hsilocal_simu_core hst1 hst2 -> - hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + hsilocal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> + hsilocal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) /\ istate_simulive (fun _ => True) dm is1 is2. @@ -167,10 +167,22 @@ Proof. Qed. -Definition hsilocal_simu dm f ctx hst1 hst2 := forall st1 st2, +Theorem hsilocal_simu_core_correct hst1 hst2 ge1 ge2 sp rs0 m0 rs m st1 st2: + hsilocal_simu_core hst1 hst2 -> + hsilocal_refines ge1 sp rs0 m0 hst1 st1 -> + hsilocal_refines ge2 sp rs0 m0 hst2 st2 -> + (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> + ssem_local ge1 sp st1 rs0 m0 rs m -> + ssem_local ge2 sp st2 rs0 m0 rs m. +Proof. +Admitted. + + +(* Definition hsilocal_simu dm f ctx hst1 hst2 := forall st1 st2, hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - silocal_simu dm f st1 st2 ctx. + silocal_simu dm f st1 st2 ctx. *) + (* BELOW: this is wrong: il faut ajouter des conditions supplémentaire vis-à-vis de hsilocal_simu_core Remark istate_simu_reflexive f dm is: istate_simu f dm is is. @@ -210,28 +222,155 @@ Qed. Record hsistate_exit := mk_hsistate_exit { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. +Definition hsiexit_simu_core dm f (hse1 hse2: hsistate_exit) := + (exists path, (fn_path f) ! (hsi_ifso hse1) = Some path) + /\ dm ! (hsi_ifso hse2) = Some (hsi_ifso hse1) + /\ hsi_cond hse1 = hsi_cond hse2 + /\ hsi_scondargs hse1 = hsi_scondargs hse2 (* FIXME - should there be something about okvals ? *) + /\ hsilocal_simu_core (hsi_elocal hse1) (hsi_elocal hse2) + /\ hsi_ifso hse1 = hsi_ifso hse2. + (** NB: we split the refinement relation between a "static" part -- independendent of the initial context and a "dynamic" part -- that depends on it *) -Definition hsistate_exit_refines_stat (hext: hsistate_exit) (ext: sistate_exit): Prop := - hsi_ifso hext = si_ifso ext. +Definition hsiexit_refines_stat (hext: hsistate_exit) (ext: sistate_exit): Prop := + hsi_cond hext = si_cond ext + /\ hsi_ifso hext = si_ifso ext. + +Definition hsok_exit ge sp rs m hse := hsok_local ge sp rs m (hsi_elocal hse). -Definition hsistate_exit_refines_dyn ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := - hsistate_local_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) - /\ (hsok_local ge sp rs0 m0 (hsi_elocal hext) -> +Definition hsiexit_refines_dyn ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := + hsilocal_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) + /\ (hsok_exit ge sp rs0 m0 hext -> seval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem_get (hsi_elocal hext)) rs0 m0 = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0). +Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, + hsiexit_refines_stat hse1 se1 -> + hsiexit_refines_stat hse2 se2 -> + hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> + hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> + siexit_simu dm f ctx se1 se2. + +Lemma hsiexit_simu_core_nofail dm f hse1 hse2 ge1 ge2 sp rs m: + hsiexit_simu_core dm f hse1 hse2 -> + (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> + hsok_exit ge1 sp rs m hse1 -> + hsok_exit ge2 sp rs m hse2. +Proof. +Admitted. + +Theorem hsiexit_simu_core_correct dm f hse1 hse2 ctx: + hsiexit_simu_core dm f hse1 hse2 -> + hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> + hsiexit_simu dm f hse1 hse2 ctx. +Proof. + intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2 is1 ICONT SSEME. + exploit hsiexit_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. + assert (SEVALC: + seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond st1) (si_scondargs st1) (si_smem (si_elocal st1)) + (the_rs0 ctx) (the_m0 ctx) = + seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond st2) (si_scondargs st2) (si_smem (si_elocal st2)) + (the_rs0 ctx) (the_m0 ctx)). + { destruct HDYN1 as (_ & SCOND1). rewrite <- SCOND1 by assumption. clear SCOND1. + destruct HDYN2 as (_ & SCOND2). rewrite <- SCOND2 by assumption. clear SCOND2. + destruct SIMUC as (_ & _ & CONDEQ & ARGSEQ & LSIMU & IFEQ). destruct LSIMU as (_ & _ & _ & MEMEQ). + rewrite CONDEQ. rewrite ARGSEQ. rewrite MEMEQ. erewrite <- seval_condition_preserved; eauto. + eapply ctx. } + constructor; auto. + exists (mk_istate false (si_ifso st2) (irs is1) (imem is1)). simpl. constructor. + - constructor; [|constructor]. + + rewrite <- SEVALC. destruct SSEME as (SCOND & _ & _). assumption. + + destruct SSEME as (_ & SLOC & _). destruct SIMUC as (_ & _ & _ & _ & LSIMU & _). + destruct HDYN1 as (LREF1 & _). destruct HDYN2 as (LREF2 & _). + eapply hsilocal_simu_core_correct; eauto. apply ctx. + + reflexivity. + - unfold istate_simu. rewrite ICONT. simpl. + destruct SIMUC as ((path & PATH) & REVEQ & _ & _ & _ & _). + assert (PCEQ: hsi_ifso hse1 = ipc is1). { destruct SSEME as (_ & _ & PCEQ). destruct HREF1 as (_ & IFSO). congruence. } + exists path. constructor; [|constructor]. + + congruence. + + constructor; [|constructor]; simpl; auto. + constructor; auto. + + destruct HREF2 as (_ & IFSO). congruence. +Qed. + +Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := + | hsok_exits_nil: hsok_exits ge sp rs m nil + | hsok_exits_cons: forall he lhe, + hsok_exits ge sp rs m lhe -> + hsok_exit ge sp rs m he -> + hsok_exits ge sp rs m (he::lhe). + (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. +Definition hsistate_simu_core dm f (hse1 hse2: hsistate) := + hsi_pc hse1 = hsi_pc hse2 + /\ list_forall2 (hsiexit_simu_core dm f) (hsi_exits hse1) (hsi_exits hse2) + /\ hsilocal_simu_core (hsi_local hse1) (hsi_local hse2). + Definition hsistate_refines_stat (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st - /\ list_forall2 hsistate_exit_refines_stat (hsi_exits hst) (si_exits st). + /\ list_forall2 hsiexit_refines_stat (hsi_exits hst) (si_exits st). Definition hsistate_refines_dyn ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop := - list_forall2 (hsistate_exit_refines_dyn ge sp rs0 m0) (hsi_exits hst) (si_exits st) - /\ hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). + list_forall2 (hsiexit_refines_dyn ge sp rs0 m0) (hsi_exits hst) (si_exits st) + /\ hsilocal_refines ge sp rs0 m0 (hsi_local hst) (si_local st). + +Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, + hsistate_refines_stat hst1 st1 -> + hsistate_refines_stat hst2 st2 -> + hsistate_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> + hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + sistate_simu dm f st1 st2 ctx. + +Definition hsok ge sp rs m hst := + hsok_exits ge sp rs m (hsi_exits hst) + /\ hsok_local ge sp rs m (hsi_local hst). + +Theorem hsistate_simu_core_correct dm f hst1 hst2 ctx: + hsistate_simu_core dm f hst1 hst2 -> + hsok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> + hsistate_simu dm f hst1 hst2 ctx. +Proof. + intros SIMUC HOK st1 st2 HREF1 HREF2 DREF1 DREF2 is1 SEMI. + unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. + - destruct SEMI as (SSEML & PCEQ & ALLFU). + +(* + unfold hsistate_simub. intro. explore. unfold hsistate_simu. + intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) (HREF1 & HLREF1) (HREF2 & HLREF2) is1 SEMI. + exploit hsiexits_simub_correct; eauto. intro ESIMU. + unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. + - destruct SEMI as (SSEML & PCEQ & ALLFU). + exploit hsilocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). + destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. + exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. + + unfold ssem_internal. simpl. + unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. + destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. + rewrite <- CONTEQ. rewrite ICONT. + constructor; [|constructor]; [eassumption | auto | eapply siexits_simu_all_fallthrough; eauto]. + + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. + destruct ISIMU as (A & B & C). simpl in *. constructor; auto. + - destruct SEMI as (ext & lx & SSEME & ALLFU). + exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). + assert (EXTSIMU: siexit_simu dm f ctx ext ext2). { + eapply list_forall2_nth_error; eauto. + - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. + - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. + assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). + congruence. } *) +Admitted. + +Definition hsistate_simu_coreb (hst1 hst2: hsistate) := OK tt. + +Theorem hsistate_simu_coreb_correct dm f hst1 hst2: + hsistate_simu_coreb hst1 hst2 = OK tt -> + hsistate_simu_core dm f hst1 hst2. +Proof. +Admitted. Lemma hsistate_refines_stat_pceq st hst: hsistate_refines_stat hst st -> @@ -242,13 +381,15 @@ Qed. Lemma hsistate_refines_dyn_local_refines ge sp rs0 m0 hst st: hsistate_refines_dyn ge sp rs0 m0 hst st -> - hsistate_local_refines ge sp rs0 m0 (hsi_local hst) (si_local st). + hsilocal_refines ge sp rs0 m0 (hsi_local hst) (si_local st). Proof. unfold hsistate_refines_dyn; intuition. Qed. Local Hint Resolve hsistate_refines_dyn_local_refines: core. + + (** * Symbolic execution of one internal step TODO: to refine symbolic values/symbolic memories with hash-consed symbolic values *) @@ -275,9 +416,9 @@ Proof. Qed. Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st hsm sm: - hsistate_local_refines ge sp rs0 m0 hst st -> + hsilocal_refines ge sp rs0 m0 hst st -> (hsok_local ge sp rs0 m0 hst -> seval_smem ge sp hsm rs0 m0 = seval_smem ge sp sm rs0 m0) -> - hsistate_local_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm). + hsilocal_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm). Proof. intros LOCREF. intros SMEMEQ. destruct LOCREF as (OKEQ & SMEMEQ' & REGEQ). constructor; [| constructor ]. @@ -308,7 +449,7 @@ Qed. Lemma hsist_set_local_correct_dyn ge sp rs0 m0 hst st pc hnxt nxt: hsistate_refines_dyn ge sp rs0 m0 hst st -> - hsistate_local_refines ge sp rs0 m0 hnxt nxt -> + hsilocal_refines ge sp rs0 m0 hnxt nxt -> hsistate_refines_dyn ge sp rs0 m0 (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt). Proof. unfold hsistate_refines_dyn; simpl; intuition. @@ -378,7 +519,8 @@ Definition simplify (rsv: root_sval) lsv sm: sval := end. Lemma simplify_correct (rsv: root_sval) lsv sm (ge: RTL.genv) (sp:val) (rs0: regset) (m0: mem) v: - seval_sval ge sp (rsv lsv sm) rs0 m0 = Some v -> seval_sval ge sp (simplify rsv lsv sm) rs0 m0 = Some v. + seval_sval ge sp (rsv lsv sm) rs0 m0 = Some v -> + seval_sval ge sp (simplify rsv lsv sm) rs0 m0 = Some v. Proof. destruct rsv; simpl; auto. - (* Rop *) @@ -438,11 +580,11 @@ Definition ok_args ge sp rs0 m0 hst lsv sm := (seval_list_sval ge sp (list_sval_inj lsv) rs0 m0 <> None /\ seval_smem ge sp sm rs0 m0 <> None). Lemma hslocal_set_sreg_correct ge sp rs0 m0 hst st r (rsv:root_sval) lsv sm sv': - hsistate_local_refines ge sp rs0 m0 hst st -> + hsilocal_refines ge sp rs0 m0 hst st -> (forall ge sp rs0 m0, ok_args ge sp rs0 m0 hst lsv sm -> (hsok_local ge sp rs0 m0 hst -> seval_sval ge sp sv' rs0 m0 = seval_sval ge sp (rsv lsv sm) rs0 m0) ) -> - hsistate_local_refines ge sp rs0 m0 (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). + hsilocal_refines ge sp rs0 m0 (hslocal_set_sreg hst r rsv lsv sm) (slocal_set_sreg st r sv'). Admitted. (** ** Execution of one instruction *) @@ -474,11 +616,12 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := | _ => None (* TODO jumptable ? *) end. -Local Hint Resolve hsist_set_local_correct_stat hsist_set_local_correct_dyn hslocal_set_mem_correct: core. +Local Hint Resolve hsist_set_local_correct_stat + hsist_set_local_correct_dyn hslocal_set_mem_correct: core. Lemma seval_sval_refines ge sp rs0 m0 hst st p: hsok_local ge sp rs0 m0 hst -> - hsistate_local_refines ge sp rs0 m0 hst st -> + hsilocal_refines ge sp rs0 m0 hst st -> seval_sval ge sp (hsi_sreg_get hst p) rs0 m0 = seval_sval ge sp (si_sreg st p) rs0 m0. Proof. intros OKL HREF. destruct HREF as (_ & _ & RSEQ). @@ -487,7 +630,7 @@ Qed. Lemma seval_list_sval_refines ge sp rs0 m0 hst st l: hsok_local ge sp rs0 m0 hst -> - hsistate_local_refines ge sp rs0 m0 hst st -> + hsilocal_refines ge sp rs0 m0 hst st -> seval_list_sval ge sp (list_sval_inj (map (hsi_sreg_get hst) l)) rs0 m0 = seval_list_sval ge sp (list_sval_inj (map (si_sreg st) l)) rs0 m0. Proof. @@ -501,7 +644,7 @@ Qed. Lemma seval_smem_refines ge sp rs0 m0 hst st : hsok_local ge sp rs0 m0 hst -> - hsistate_local_refines ge sp rs0 m0 hst st -> + hsilocal_refines ge sp rs0 m0 hst st -> seval_smem ge sp (hsi_smem_get hst) rs0 m0 = seval_smem ge sp (si_smem st) rs0 m0. Proof. intros OKL HLREF. destruct HLREF as (_ & MSEQ & _). @@ -511,8 +654,9 @@ Qed. Lemma seval_condition_refines hst st ge sp cond args rs m: hsok_local ge sp rs m hst -> - hsistate_local_refines ge sp rs m hst st -> - seval_condition ge sp cond args (hsi_smem_get hst) rs m = seval_condition ge sp cond args (si_smem st) rs m. + hsilocal_refines ge sp rs m hst st -> + seval_condition ge sp cond args (hsi_smem_get hst) rs m + = seval_condition ge sp cond args (si_smem st) rs m. Proof. intros HOK (OKEQ & MEMEQ & RSEQ). unfold seval_condition. rewrite <- MEMEQ; auto. rewrite hsi_smem_eval_correct. reflexivity. @@ -527,7 +671,8 @@ Qed. Lemma hsiexec_inst_correct_stat i hst hst' st: hsiexec_inst i hst = Some hst' -> - exists st', siexec_inst i st = Some st' /\ (hsistate_refines_stat hst st -> hsistate_refines_stat hst' st'). + exists st', siexec_inst i st = Some st' + /\ (hsistate_refines_stat hst st -> hsistate_refines_stat hst' st'). Proof. destruct i; simpl; intros STEPI; inversion_clear STEPI; discriminate || eexists; split; eauto. (* TODO *) @@ -538,7 +683,8 @@ Lemma hsiexec_inst_correct_dyn ge sp rs0 m0 i hst st hst' st': siexec_inst i st = Some st' -> hsistate_refines_dyn ge sp rs0 m0 hst st -> hsistate_refines_dyn ge sp rs0 m0 hst' st'. Proof. - destruct i; simpl; intros STEP1 STEP2; inversion_clear STEP1; inversion_clear STEP2; discriminate || (intro REF; eauto). + destruct i; simpl; intros STEP1 STEP2; inversion_clear STEP1; + inversion_clear STEP2; discriminate || (intro REF; eauto). - (* Iop *) eapply hsist_set_local_correct_dyn; eauto. eapply hslocal_set_sreg_correct; eauto. @@ -636,13 +782,13 @@ Proof. Admitted. *) Lemma sfind_function_conserves hsl sl pge ge sp s rs0 m0: - hsistate_local_refines ge sp rs0 m0 hsl sl -> + hsilocal_refines ge sp rs0 m0 hsl sl -> sfind_function pge ge sp (sum_left_map (hsi_sreg_get hsl) s) rs0 m0 = sfind_function pge ge sp (sum_left_map (si_sreg sl) s) rs0 m0. Admitted. Lemma hsexec_final_correct hsl sl i: - (forall ge sp rs0 m0, hsistate_local_refines ge sp rs0 m0 hsl sl) -> + (forall ge sp rs0 m0, hsilocal_refines ge sp rs0 m0 hsl sl) -> hsexec_final i hsl = sexec_final i sl. Proof. (* destruct i; simpl; intros HLREF; try (apply hfinal_refines_snone). @@ -665,10 +811,11 @@ Proof. Admitted. -Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. +Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; + hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. Remark init_hsistate_local_correct ge sp rs0 m0: - hsistate_local_refines ge sp rs0 m0 init_hsistate_local init_sistate_local. + hsilocal_refines ge sp rs0 m0 init_hsistate_local init_sistate_local. Proof. constructor; constructor; simpl. - intro. destruct H as (_ & SMEM & SVAL). constructor; [contradiction|]. @@ -704,7 +851,8 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). -Local Hint Resolve init_hsistate_correct_stat init_hsistate_correct_dyn hsexec_final_correct hsiexec_inst_correct_dyn hsiexec_path_correct_dyn hfinal_refines_snone: core. +Local Hint Resolve init_hsistate_correct_stat init_hsistate_correct_dyn hsexec_final_correct + hsiexec_inst_correct_dyn hsiexec_path_correct_dyn hfinal_refines_snone: core. Lemma hsexec_correct f pc hst: hsexec f pc = Some hst -> @@ -739,14 +887,16 @@ Definition hsilocal_simu_core (hst1 hst2: hsistate_local):= /\ forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1) /\ forall r, hsi_sreg_get hst2 r = hsi_sreg_get hst1 r. +*) + Definition hsilocal_simu dm f (hst1 hst2: hsistate_local) (ctx: simu_proof_context f): Prop := forall st1 st2, - hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - silocal_simu dm f st1 st2 ctx. *) + hsilocal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> + hsilocal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> + silocal_simu dm f st1 st2 ctx. Lemma hsilocal_simub_correct dm f hst1 hst2: hsilocal_simub dm f hst1 hst2 = OK tt -> - forall ctx, hsilocal_simu dm f ctx hst1 hst2. + forall ctx, hsilocal_simu dm f hst1 hst2 ctx. Proof. Admitted. @@ -893,12 +1043,12 @@ Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hs else Error (msg "siexit_simub: conditions do not match") . -Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, - hsistate_exit_refines_stat hse1 se1 -> - hsistate_exit_refines_stat hse2 se2 -> - hsistate_exit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> - hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> - siexit_simu dm f ctx se1 se2. +(* Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, + hsiexit_refines_stat hse1 se1 -> + hsiexit_refines_stat hse2 se2 -> + hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> + hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> + siexit_simu dm f ctx se1 se2. *) Local Hint Resolve genv_match ssem_local_refines_hok: core. @@ -906,13 +1056,13 @@ Lemma hsiexit_simub_correct dm f hse1 hse2: hsiexit_simub dm f hse1 hse2 = OK tt -> forall ctx, hsiexit_simu dm f hse1 hse2 ctx. Proof. - unfold hsiexit_simub. intro. explore. +(* unfold hsiexit_simub. intro. explore. apply list_sval_simub_correct in EQ0. exploit hsilocal_simub_eq; eauto. intros (SREGEQ & SMEMEQ) ctx. eapply hsilocal_simub_correct in EQ2. apply revmap_check_single_correct in EQ3. intros se1 se2 HIFSO1 HIFSO2 (HLREF1 & HCOND1) (HLREF2 & HCOND2). - unfold hsistate_exit_refines_stat in *. + unfold hsiexit_refines_stat in *. intros is1 (SCOND & SLOCAL & SIFSO). constructor. - rewrite <- HCOND1; eauto. @@ -936,7 +1086,7 @@ Proof. + unfold istate_simu in *. destruct (icontinue is1) eqn:ICONT. * destruct ISIMU as (A & B & C). constructor; auto. * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. - repeat (constructor; auto). simpl in *. congruence. + repeat (constructor; auto). simpl in *. congruence. *) Admitted. Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := @@ -957,10 +1107,10 @@ Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: l end. Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit) (ctx: simu_proof_context f): Prop := forall se1 se2, - list_forall2 hsistate_exit_refines_stat hse1 se1 -> - list_forall2 hsistate_exit_refines_stat hse2 se2 -> - list_forall2 (hsistate_exit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse1 se1 -> - list_forall2 (hsistate_exit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse2 se2 -> + list_forall2 hsiexit_refines_stat hse1 se1 -> + list_forall2 hsiexit_refines_stat hse2 se2 -> + list_forall2 (hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse1 se1 -> + list_forall2 (hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse2 se2 -> siexits_simu dm f se1 se2 ctx. Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, @@ -982,12 +1132,12 @@ Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: h do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); OK tt. -Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, +(* Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, hsistate_refines_stat hst1 st1 -> hsistate_refines_stat hst2 st2 -> hsistate_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - sistate_simu dm f st1 st2 ctx. + sistate_simu dm f st1 st2 ctx. *) Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, siexits_simu dm f lse1 lse2 ctx -> diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 73d7fa2c..1299758a 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1762,7 +1762,9 @@ Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sista /\ istate_simu f dm is1 is2. Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (ctx: simu_proof_context f) (se1 se2: sistate_exit) := - forall is1, ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> + forall is1, + icontinue is1 = false -> + ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx)) = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)) /\ exists is2, -- cgit From c24c4b0ba6cd425965f9d073cca3572fcbaf333f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 20 Jul 2020 18:27:05 +0200 Subject: More robust code for changing order of instructions --- kvx/lib/RTLpathScheduleraux.ml | 161 ++++++++++++++++++++++++++++++----------- 1 file changed, 118 insertions(+), 43 deletions(-) diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml index ac7238db..88f777a5 100644 --- a/kvx/lib/RTLpathScheduleraux.ml +++ b/kvx/lib/RTLpathScheduleraux.ml @@ -159,6 +159,123 @@ let schedule_superblock sb code = end *) (* stub: identity function *) +(** + * Perform basic checks on the new order : + * - must have the same length as the old order + * - non basic instructions (call, tailcall, return, jumptable, non predicted CB) must not move + *) +let check_order code old_order new_order = begin + assert ((Array.length old_order) == (Array.length new_order)); + let length = Array.length new_order in + if length > 0 then + let last_inst = Array.get old_order (length - 1) in + let instr = get_some @@ PTree.get last_inst code in + match predicted_successor instr with + | None -> + if (last_inst != Array.get new_order (length - 1)) then + failwith "The last instruction of the superblock is not basic, but was moved" + | _ -> () +end + +type sinst = + (* Each middle instruction has a direct successor *) + (* A Smid can be the last instruction of a superblock, but a Send cannot be moved *) + | Smid of RTL.instruction * node + | Send of RTL.instruction + +let rinst_to_sinst inst = + match inst with + | Inop n -> Smid(inst, n) + | Iop (_,_,_,n) -> Smid(inst, n) + | Iload (_,_,_,_,_,n) -> Smid(inst, n) + | Istore (_,_,_,_,n) -> Smid(inst, n) + | Icond (_,_,n1,n2,p) -> ( + match p with + | Some true -> Smid(inst, n1) + | Some false -> Smid(inst, n2) + | None -> Send(inst) + ) + | Icall _ | Ibuiltin _ | Ijumptable _ | Itailcall _ | Ireturn _ -> Send(inst) + +let change_predicted_successor s = function + | Smid(i, n) -> Smid(i, s) + | Send _ -> failwith "Called change_predicted_successor on Send. Are you trying to move a non-basic instruction in the middle of the block?" + +(* Forwards the successor changes into an RTL instruction *) +let sinst_to_rinst = function + | Smid(inst, s) -> ( + match inst with + | Inop n -> Inop s + | Iop (a,b,c,n) -> Iop (a,b,c,s) + | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) + | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s) + | Icond (a,b,n1,n2,p) -> ( + match p with + | Some true -> Icond(a, b, s, n2, p) + | Some false -> Icond(a, b, n1, s, p) + | None -> failwith "Non predicted Icond as a middle instruction!" + ) + | _ -> failwith "That instruction shouldn't be a middle instruction" + ) + | Send i -> i + +let apply_schedule code sb new_order = + let tc = ref code in + let old_order = sb.instructions in + begin + check_order code old_order new_order; + Array.iteri (fun i n' -> + let inst' = get_some @@ PTree.get n' code in + let iend = Array.length old_order - 1 in + let new_inst = + if (i == iend) then + let final_inst_node = Array.get old_order iend in + let sinst' = rinst_to_sinst inst' in + match sinst' with + (* The below assert fails if a Send is in the middle of the original superblock *) + | Send i -> (assert (final_inst_node == n'); i) + | Smid _ -> + let final_inst = get_some @@ PTree.get final_inst_node code in + match rinst_to_sinst final_inst with + | Smid (_, s') -> sinst_to_rinst @@ change_predicted_successor s' sinst' + | Send _ -> assert(false) (* should have failed earlier *) + else + sinst_to_rinst + (* this will fail if the moved instruction is a Send *) + @@ change_predicted_successor (Array.get old_order (i+1)) + @@ rinst_to_sinst inst' + in tc := PTree.set (Array.get old_order i) new_inst !tc + ) new_order; + !tc + end + + (* +let main_successors = function +| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n] +| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] +| Icond (_,_,n1,n2,p) -> ( + match p with + | Some true -> [n1; n2] + | Some false -> [n2; n1] + | None -> [n1; n2] ) +| Ijumptable _ | Itailcall _ | Ireturn _ -> [] + +let change_predicted_successor i s = match i with + | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5) - Tailcalls and returns should not be moved in the middle of a superblock" + | Ijumptable _ -> failwith "Wrong instruction (6) - Jumptables should not be moved in the middle of a superblock" + | Inop n -> Inop s + | Iop (a,b,c,n) -> Iop (a,b,c,s) + | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) + | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s) + | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) + | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) + | Icond (a,b,n1,n2,p) -> ( + match p with + | Some true -> Icond (a,b,s,n2,p) + | Some false -> Icond (a,b,n1,s,p) + | None -> failwith "Predicted a successor for an Icond with p=None - unpredicted CB should not be moved in the middle of the superblock" + ) + let rec change_successors i = function | [] -> ( match i with @@ -193,49 +310,6 @@ let rec change_successors i = function end | _ -> failwith "Wrong instruction (4)") -let change_predicted_successor i s = match i with - | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5) - Tailcalls and returns should not be moved in the middle of a superblock" - | Ijumptable _ -> failwith "Wrong instruction (6) - Jumptables should not be moved in the middle of a superblock" - | Inop n -> Inop s - | Iop (a,b,c,n) -> Iop (a,b,c,s) - | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s) - | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s) - | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s) - | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s) - | Icond (a,b,n1,n2,p) -> ( - match p with - | Some true -> Icond (a,b,s,n2,p) - | Some false -> Icond (a,b,n1,s,p) - | None -> failwith "Predicted a successor for an Icond with p=None - unpredicted CB should not be moved in the middle of the superblock" - ) - -(** - * Perform basic checks on the new order : - * - must have the same length as the old order - * - non basic instructions (call, tailcall, return, jumptable, non predicted CB) must not move - *) -let check_order code old_order new_order = begin - assert ((Array.length old_order) == (Array.length new_order)); - let length = Array.length new_order in - if length > 0 then - let last_inst = Array.get old_order (length - 1) in - let instr = get_some @@ PTree.get last_inst code in - match predicted_successor instr with - | None -> - if (last_inst != Array.get new_order (length - 1)) then - failwith "The last instruction of the superblock is not basic, but was moved" - | _ -> () -end - -let main_successors = function -| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n] -| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n] -| Icond (_,_,n1,n2,p) -> ( - match p with - | Some true -> [n1; n2] - | Some false -> [n2; n1] - | None -> [n1; n2] ) -| Ijumptable _ | Itailcall _ | Ireturn _ -> [] let apply_schedule code sb new_order = let tc = ref code in @@ -256,6 +330,7 @@ let apply_schedule code sb new_order = ) new_order; !tc end +*) let rec do_schedule code = function | [] -> code -- cgit From 392bfbaa7374ec625e5ec660c65515744f0dffd0 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 21 Jul 2020 12:06:25 +0200 Subject: hsistate_simu_core_correct --- kvx/lib/RTLpathSE_impl.v | 345 +++++++++++++++++++++++---------------------- kvx/lib/RTLpathSE_theory.v | 8 +- 2 files changed, 183 insertions(+), 170 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 20d7fe3d..8c5b84a1 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -138,7 +138,7 @@ Proof. repeat (constructor; auto). Qed. -Lemma hsilocal_simu_core_correct' dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: +(* Lemma hsilocal_simu_core_correct' dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: hsilocal_simu_core hst1 hst2 -> hsilocal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> hsilocal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> @@ -164,8 +164,7 @@ Proof. rewrite C. erewrite seval_preserved; [| eapply ctx]. rewrite <- hsi_sreg_eval_correct. rewrite B; auto. - apply istate_simulive_reflexive. -Qed. - +Qed. *) Theorem hsilocal_simu_core_correct hst1 hst2 ge1 ge2 sp rs0 m0 rs m st1 st2: hsilocal_simu_core hst1 hst2 -> @@ -175,47 +174,22 @@ Theorem hsilocal_simu_core_correct hst1 hst2 ge1 ge2 sp rs0 m0 rs m st1 st2: ssem_local ge1 sp st1 rs0 m0 rs m -> ssem_local ge2 sp st2 rs0 m0 rs m. Proof. -Admitted. - - -(* Definition hsilocal_simu dm f ctx hst1 hst2 := forall st1 st2, - hsistate_local_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsistate_local_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - silocal_simu dm f st1 st2 ctx. *) - - -(* BELOW: this is wrong: il faut ajouter des conditions supplémentaire vis-à-vis de hsilocal_simu_core -Remark istate_simu_reflexive f dm is: istate_simu f dm is is. -Proof. - unfold istate_simu. destruct (icontinue is). - - repeat (constructor; auto). - - eexists. constructor; [|constructor]. all: admit. -Admitted. (* Problem : this remark is not true *) - - -Lemma hsilocal_simu_core_correct hst1 hst2 f ctx dm: - hsilocal_simu_core hst1 hst2 -> - hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> - hsilocal_simu dm f ctx hst1 hst2. -Proof. - intros CORE HOK1 st1 st2 HREF1 HREF2. - exploit hsilocal_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. - intros is1 (PRE & MEMEQ & RSEQ). - eexists. constructor. - - constructor; [|constructor]. - + destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. - + destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _). - destruct CORE as (_ & _ & _ & MEMEQ3). - rewrite <- MEMEQ2; auto. rewrite hsi_smem_eval_correct. rewrite <- MEMEQ3. - erewrite smem_eval_preserved; [| eapply ctx]. - rewrite <- hsi_smem_eval_correct. rewrite MEMEQ1; auto. eassumption. - + intro r. destruct HREF2 as (_ & _ & A). destruct HREF1 as (_ & _ & B). - destruct CORE as (_ & _ & C & _). rewrite <- A; auto. rewrite hsi_sreg_eval_correct. - rewrite C. erewrite seval_preserved; [| eapply ctx]. rewrite <- hsi_sreg_eval_correct. - rewrite B; auto. - - apply istate_simu_reflexive. + intros CORE HREF1 HREF2 GFS SEML. + exploit ssem_local_refines_hok; eauto. intro HOK1. + exploit hsilocal_simu_core_nofail; eauto. intro HOK2. + destruct SEML as (PRE & MEMEQ & RSEQ). + constructor; [|constructor]. + + destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. + + destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _). + destruct CORE as (_ & _ & _ & MEMEQ3). + rewrite <- MEMEQ2; auto. rewrite hsi_smem_eval_correct. rewrite <- MEMEQ3. + erewrite smem_eval_preserved; [| eapply GFS]. + rewrite <- hsi_smem_eval_correct. rewrite MEMEQ1; auto. + + intro r. destruct HREF2 as (_ & _ & A). destruct HREF1 as (_ & _ & B). + destruct CORE as (_ & _ & C & _). rewrite <- A; auto. rewrite hsi_sreg_eval_correct. + rewrite C. erewrite seval_preserved; [| eapply GFS]. rewrite <- hsi_sreg_eval_correct. + rewrite B; auto. Qed. -*) (* Syntax and semantics of symbolic exit states *) (* TODO: add hash-consing *) @@ -227,8 +201,7 @@ Definition hsiexit_simu_core dm f (hse1 hse2: hsistate_exit) := /\ dm ! (hsi_ifso hse2) = Some (hsi_ifso hse1) /\ hsi_cond hse1 = hsi_cond hse2 /\ hsi_scondargs hse1 = hsi_scondargs hse2 (* FIXME - should there be something about okvals ? *) - /\ hsilocal_simu_core (hsi_elocal hse1) (hsi_elocal hse2) - /\ hsi_ifso hse1 = hsi_ifso hse2. + /\ hsilocal_simu_core (hsi_elocal hse1) (hsi_elocal hse2). (** NB: we split the refinement relation between a "static" part -- independendent of the initial context and a "dynamic" part -- that depends on it @@ -245,7 +218,7 @@ Definition hsiexit_refines_dyn ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_ seval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem_get (hsi_elocal hext)) rs0 m0 = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0). -Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, +Definition hsiexit_simu dm f (ctx: simu_proof_context f) hse1 hse2: Prop := forall se1 se2, hsiexit_refines_stat hse1 se1 -> hsiexit_refines_stat hse2 se2 -> hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> @@ -258,14 +231,17 @@ Lemma hsiexit_simu_core_nofail dm f hse1 hse2 ge1 ge2 sp rs m: hsok_exit ge1 sp rs m hse1 -> hsok_exit ge2 sp rs m hse2. Proof. -Admitted. + intros CORE GFS HOK1. + destruct CORE as (_ & _ & _ & _ & CORE). + eapply hsilocal_simu_core_nofail; eauto. +Qed. Theorem hsiexit_simu_core_correct dm f hse1 hse2 ctx: hsiexit_simu_core dm f hse1 hse2 -> hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> - hsiexit_simu dm f hse1 hse2 ctx. + hsiexit_simu dm f ctx hse1 hse2. Proof. - intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2 is1 ICONT SSEME. + intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2 is1 (* ICONT *) SSEME. exploit hsiexit_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. assert (SEVALC: seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond st1) (si_scondargs st1) (si_smem (si_elocal st1)) @@ -274,25 +250,27 @@ Proof. (the_rs0 ctx) (the_m0 ctx)). { destruct HDYN1 as (_ & SCOND1). rewrite <- SCOND1 by assumption. clear SCOND1. destruct HDYN2 as (_ & SCOND2). rewrite <- SCOND2 by assumption. clear SCOND2. - destruct SIMUC as (_ & _ & CONDEQ & ARGSEQ & LSIMU & IFEQ). destruct LSIMU as (_ & _ & _ & MEMEQ). + destruct SIMUC as (_ & _ & CONDEQ & ARGSEQ & LSIMU). destruct LSIMU as (_ & _ & _ & MEMEQ). rewrite CONDEQ. rewrite ARGSEQ. rewrite MEMEQ. erewrite <- seval_condition_preserved; eauto. eapply ctx. } constructor; auto. - exists (mk_istate false (si_ifso st2) (irs is1) (imem is1)). simpl. constructor. + exists (mk_istate (icontinue is1) (si_ifso st2) (irs is1) (imem is1)). simpl. constructor. - constructor; [|constructor]. + rewrite <- SEVALC. destruct SSEME as (SCOND & _ & _). assumption. - + destruct SSEME as (_ & SLOC & _). destruct SIMUC as (_ & _ & _ & _ & LSIMU & _). + + destruct SSEME as (_ & SLOC & _). destruct SIMUC as (_ & _ & _ & _ & LSIMU). destruct HDYN1 as (LREF1 & _). destruct HDYN2 as (LREF2 & _). eapply hsilocal_simu_core_correct; eauto. apply ctx. + reflexivity. - - unfold istate_simu. rewrite ICONT. simpl. - destruct SIMUC as ((path & PATH) & REVEQ & _ & _ & _ & _). - assert (PCEQ: hsi_ifso hse1 = ipc is1). { destruct SSEME as (_ & _ & PCEQ). destruct HREF1 as (_ & IFSO). congruence. } - exists path. constructor; [|constructor]. - + congruence. - + constructor; [|constructor]; simpl; auto. + - unfold istate_simu. destruct (icontinue is1) eqn:ICONT. + * constructor; [|constructor]; simpl; auto. constructor; auto. - + destruct HREF2 as (_ & IFSO). congruence. + * simpl. destruct SIMUC as ((path & PATH) & REVEQ & _ & _ & _ & _). + assert (PCEQ: hsi_ifso hse1 = ipc is1). { destruct SSEME as (_ & _ & PCEQ). destruct HREF1 as (_ & IFSO). congruence. } + exists path. constructor; [|constructor]. + + congruence. + + constructor; [|constructor]; simpl; auto. + constructor; auto. + + destruct HREF2 as (_ & IFSO). congruence. Qed. Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := @@ -302,6 +280,28 @@ Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := hsok_exit ge sp rs m he -> hsok_exits ge sp rs m (he::lhe). +Definition hsiexits_simu dm f (ctx: simu_proof_context f) lhse1 lhse2: Prop := + list_forall2 (hsiexit_simu dm f ctx) lhse1 lhse2. + +Definition hsiexits_simu_core dm f lhse1 lhse2: Prop := + list_forall2 (hsiexit_simu_core dm f) lhse1 lhse2. + +Theorem hsiexits_simu_core_correct dm f lhse1 lhse2 ctx: + hsiexits_simu_core dm f lhse1 lhse2 -> + hsok_exits (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse1 -> + hsiexits_simu dm f ctx lhse1 lhse2. +Proof. + induction 1; [constructor|]. + intros HOKS. inv HOKS. constructor; [|apply IHlist_forall2; assumption]. + apply hsiexit_simu_core_correct; assumption. +Qed. + +Definition hsiexits_refines_stat lhse lse := + list_forall2 hsiexit_refines_stat lhse lse. + +Definition hsiexits_refines_dyn ge sp rs0 m0 lhse se := + list_forall2 (hsiexit_refines_dyn ge sp rs0 m0) lhse se. + (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. @@ -312,10 +312,10 @@ Definition hsistate_simu_core dm f (hse1 hse2: hsistate) := Definition hsistate_refines_stat (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st - /\ list_forall2 hsiexit_refines_stat (hsi_exits hst) (si_exits st). + /\ hsiexits_refines_stat (hsi_exits hst) (si_exits st). Definition hsistate_refines_dyn ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop := - list_forall2 (hsiexit_refines_dyn ge sp rs0 m0) (hsi_exits hst) (si_exits st) + hsiexits_refines_dyn ge sp rs0 m0 (hsi_exits hst) (si_exits st) /\ hsilocal_refines ge sp rs0 m0 (hsi_local hst) (si_local st). Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, @@ -329,40 +329,128 @@ Definition hsok ge sp rs m hst := hsok_exits ge sp rs m (hsi_exits hst) /\ hsok_local ge sp rs m (hsi_local hst). +Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, + siexits_simu dm f lse1 lse2 ctx -> + all_fallthrough (the_ge1 ctx) (the_sp ctx) lse1 (the_rs0 ctx) (the_m0 ctx) -> + all_fallthrough (the_ge2 ctx) (the_sp ctx) lse2 (the_rs0 ctx) (the_m0 ctx). +Proof. + induction 1; [unfold all_fallthrough; contradiction|]. + intros X ext INEXT. eapply all_fallthrough_revcons in X. destruct X as (SEVAL & ALLFU). + apply IHlist_forall2 in ALLFU. +Admitted. +(* BROKEN! + destruct H as (_ & CONDSIMU). + inv INEXT; [|eauto]. + erewrite <- CONDSIMU; eauto. +Qed. +*) + +Lemma hsiexits_simu_siexits dm f ctx lhse1 lhse2 lse1 lse2: + hsiexits_simu dm f ctx lhse1 lhse2 -> + hsiexits_refines_stat lhse1 lse1 -> + hsiexits_refines_stat lhse2 lse2 -> + hsiexits_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse1 lse1 -> + hsiexits_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse2 lse2 -> + siexits_simu dm f lse1 lse2 ctx. +Proof. +Admitted. + +Lemma siexits_simu_all_fallthrough_upto dm f ctx lse1 lse2: + siexits_simu dm f lse1 lse2 ctx -> forall ext1 lx1, + all_fallthrough_upto_exit (the_ge1 ctx) (the_sp ctx) ext1 lx1 lse1 (the_rs0 ctx) (the_m0 ctx) -> + exists ext2 lx2, + all_fallthrough_upto_exit (the_ge2 ctx) (the_sp ctx) ext2 lx2 lse2 (the_rs0 ctx) (the_m0 ctx) + /\ length lx1 = length lx2. +Proof. + induction 1. + - intros. destruct H as (ITAIL & ALLFU). eapply is_tail_false in ITAIL. contradiction. + - intros ext1 lx1 ALLFUE. + destruct ALLFUE as (ITAIL & ALLFU). inv ITAIL. + + eexists; eexists. + constructor; [| eapply list_forall2_length; eauto]. + constructor; [econstructor | eapply siexits_simu_all_fallthrough; eauto]. + + exploit IHlist_forall2; [constructor; eauto|]. + intros (ext2 & lx2 & ALLFUE2 & LENEQ). + eexists; eexists. constructor; eauto. + eapply all_fallthrough_upto_exit_cons; eauto. +Qed. + +Lemma list_forall2_nth_error {A} (l1 l2: list A) P: + list_forall2 P l1 l2 -> + forall x1 x2 n, + nth_error l1 n = Some x1 -> + nth_error l2 n = Some x2 -> + P x1 x2. +Proof. + induction 1. + - intros. rewrite nth_error_nil in H. discriminate. + - intros x1 x2 n. destruct n as [|n]; simpl. + + intros. inv H1. inv H2. assumption. + + apply IHlist_forall2. +Qed. + +Lemma is_tail_length {A} (l1 l2: list A): + is_tail l1 l2 -> + (length l1 <= length l2)%nat. +Proof. + induction l2. + - intro. destruct l1; auto. apply is_tail_false in H. contradiction. + - intros ITAIL. inv ITAIL; auto. + apply IHl2 in H1. clear IHl2. simpl. omega. +Qed. + +Lemma is_tail_nth_error {A} (l1 l2: list A) x: + is_tail (x::l1) l2 -> + nth_error l2 ((length l2) - length l1 - 1) = Some x. +Proof. + induction l2. + - intro ITAIL. apply is_tail_false in ITAIL. contradiction. + - intros ITAIL. assert (length (a::l2) = S (length l2)) by auto. rewrite H. clear H. + assert (forall n n', ((S n) - n' - 1)%nat = (n - n')%nat) by (intros; omega). rewrite H. clear H. + inv ITAIL. + + assert (forall n, (n - n)%nat = 0%nat) by (intro; omega). rewrite H. + simpl. reflexivity. + + exploit IHl2; eauto. intros. clear IHl2. + assert (forall n n', (n > n')%nat -> (n - n')%nat = S (n - n' - 1)%nat) by (intros; omega). + exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2. + assert ((length l2 > length l1)%nat) by omega. clear H2. + rewrite H0; auto. +Qed. + Theorem hsistate_simu_core_correct dm f hst1 hst2 ctx: hsistate_simu_core dm f hst1 hst2 -> hsok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> hsistate_simu dm f hst1 hst2 ctx. Proof. intros SIMUC HOK st1 st2 HREF1 HREF2 DREF1 DREF2 is1 SEMI. + destruct HREF1 as (PCREF1 & EREF1). destruct HREF2 as (PCREF2 & EREF2). + destruct DREF1 as (DEREF1 & LREF1). destruct DREF2 as (DEREF2 & LREF2). + destruct SIMUC as (PCSIMU & ESIMU & LSIMU). destruct HOK as (EOK1 & LOK1). + exploit hsiexits_simu_core_correct; eauto. intro HESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). - -(* - unfold hsistate_simub. intro. explore. unfold hsistate_simu. - intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) (HREF1 & HLREF1) (HREF2 & HLREF2) is1 SEMI. - exploit hsiexits_simub_correct; eauto. intro ESIMU. - unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - - destruct SEMI as (SSEML & PCEQ & ALLFU). - exploit hsilocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). - destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. - exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. - + unfold ssem_internal. simpl. - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. - destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. - rewrite <- CONTEQ. rewrite ICONT. - constructor; [|constructor]; [eassumption | auto | eapply siexits_simu_all_fallthrough; eauto]. - + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. - destruct ISIMU as (A & B & C). simpl in *. constructor; auto. + exploit hsilocal_simu_core_correct; eauto; [apply ctx|]. intro SSEML2. + exists (mk_istate (icontinue is1) (si_pc st2) (irs is1) (imem is1)). constructor. + + unfold ssem_internal. simpl. rewrite ICONT. constructor; [assumption | constructor; [reflexivity |]]. + eapply siexits_simu_all_fallthrough; eauto. eapply hsiexits_simu_siexits; eauto. + + unfold istate_simu. rewrite ICONT. constructor; [simpl; assumption | constructor; [| reflexivity]]. + constructor. - destruct SEMI as (ext & lx & SSEME & ALLFU). + assert (SESIMU: siexits_simu dm f (si_exits st1) (si_exits st2) ctx) by (eapply hsiexits_simu_siexits; eauto). exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). assert (EXTSIMU: siexit_simu dm f ctx ext ext2). { eapply list_forall2_nth_error; eauto. - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). - congruence. } *) -Admitted. + congruence. } + apply EXTSIMU in SSEME. destruct SSEME as (CONDEVAL & (is2 & SSEME2 & ISIMU)). + exists (mk_istate (icontinue is1) (ipc is2) (irs is2) (imem is2)). constructor. + + unfold ssem_internal. simpl. rewrite ICONT. exists ext2, lx2. constructor; assumption. + + unfold istate_simu in *. rewrite ICONT in *. destruct ISIMU as (path & PATHEQ & ISIMULIVE & DMEQ). + destruct ISIMULIVE as (CONTEQ & REGEQ & MEMEQ). + exists path. repeat (constructor; auto). +Qed. Definition hsistate_simu_coreb (hst1 hst2: hsistate) := OK tt. @@ -1054,7 +1142,7 @@ Local Hint Resolve genv_match ssem_local_refines_hok: core. Lemma hsiexit_simub_correct dm f hse1 hse2: hsiexit_simub dm f hse1 hse2 = OK tt -> - forall ctx, hsiexit_simu dm f hse1 hse2 ctx. + forall ctx, hsiexit_simu dm f ctx hse1 hse2. Proof. (* unfold hsiexit_simub. intro. explore. apply list_sval_simub_correct in EQ0. exploit hsilocal_simub_eq; eauto. @@ -1106,26 +1194,19 @@ Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: l end end. -Definition hsiexits_simu dm f (hse1 hse2: list hsistate_exit) (ctx: simu_proof_context f): Prop := forall se1 se2, - list_forall2 hsiexit_refines_stat hse1 se1 -> - list_forall2 hsiexit_refines_stat hse2 se2 -> - list_forall2 (hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse1 se1 -> - list_forall2 (hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx)) hse2 se2 -> - siexits_simu dm f se1 se2 ctx. - Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, hsiexits_simub dm f lhse1 lhse2 = OK tt -> - hsiexits_simu dm f lhse1 lhse2 ctx. + hsiexits_simu dm f ctx lhse1 lhse2. Proof. - induction lhse1. +(* induction lhse1. - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. inv HEREFS1. inv HEREFS2. constructor. - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. fold hsiexits_simub in EQ1. eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. - apply EQ1; assumption. -Qed. + apply EQ1; assumption. *) +Admitted. Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := do _ <- hsilocal_simub dm f (hsi_local hst1) (hsi_local hst2); @@ -1139,90 +1220,20 @@ Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: h hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> sistate_simu dm f st1 st2 ctx. *) -Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, - siexits_simu dm f lse1 lse2 ctx -> - all_fallthrough (the_ge1 ctx) (the_sp ctx) lse1 (the_rs0 ctx) (the_m0 ctx) -> - all_fallthrough (the_ge2 ctx) (the_sp ctx) lse2 (the_rs0 ctx) (the_m0 ctx). -Proof. - induction 1; [unfold all_fallthrough; contradiction|]. - intros X ext INEXT. eapply all_fallthrough_revcons in X. destruct X as (SEVAL & ALLFU). - apply IHlist_forall2 in ALLFU. -Admitted. -(* BROKEN! - destruct H as (_ & CONDSIMU). - inv INEXT; [|eauto]. - erewrite <- CONDSIMU; eauto. -Qed. -*) -Lemma siexits_simu_all_fallthrough_upto dm f ctx lse1 lse2: - siexits_simu dm f lse1 lse2 ctx -> forall ext1 lx1, - all_fallthrough_upto_exit (the_ge1 ctx) (the_sp ctx) ext1 lx1 lse1 (the_rs0 ctx) (the_m0 ctx) -> - exists ext2 lx2, - all_fallthrough_upto_exit (the_ge2 ctx) (the_sp ctx) ext2 lx2 lse2 (the_rs0 ctx) (the_m0 ctx) - /\ length lx1 = length lx2. -Proof. - induction 1. - - intros. destruct H as (ITAIL & ALLFU). eapply is_tail_false in ITAIL. contradiction. - - intros ext1 lx1 ALLFUE. - destruct ALLFUE as (ITAIL & ALLFU). inv ITAIL. - + eexists; eexists. - constructor; [| eapply list_forall2_length; eauto]. - constructor; [econstructor | eapply siexits_simu_all_fallthrough; eauto]. - + exploit IHlist_forall2; [constructor; eauto|]. - intros (ext2 & lx2 & ALLFUE2 & LENEQ). - eexists; eexists. constructor; eauto. - eapply all_fallthrough_upto_exit_cons; eauto. -Qed. -Lemma list_forall2_nth_error {A} (l1 l2: list A) P: - list_forall2 P l1 l2 -> - forall x1 x2 n, - nth_error l1 n = Some x1 -> - nth_error l2 n = Some x2 -> - P x1 x2. -Proof. - induction 1. - - intros. rewrite nth_error_nil in H. discriminate. - - intros x1 x2 n. destruct n as [|n]; simpl. - + intros. inv H1. inv H2. assumption. - + apply IHlist_forall2. -Qed. -Lemma is_tail_length {A} (l1 l2: list A): - is_tail l1 l2 -> - (length l1 <= length l2)%nat. -Proof. - induction l2. - - intro. destruct l1; auto. apply is_tail_false in H. contradiction. - - intros ITAIL. inv ITAIL; auto. - apply IHl2 in H1. clear IHl2. simpl. omega. -Qed. -Lemma is_tail_nth_error {A} (l1 l2: list A) x: - is_tail (x::l1) l2 -> - nth_error l2 ((length l2) - length l1 - 1) = Some x. -Proof. - induction l2. - - intro ITAIL. apply is_tail_false in ITAIL. contradiction. - - intros ITAIL. assert (length (a::l2) = S (length l2)) by auto. rewrite H. clear H. - assert (forall n n', ((S n) - n' - 1)%nat = (n - n')%nat) by (intros; omega). rewrite H. clear H. - inv ITAIL. - + assert (forall n, (n - n)%nat = 0%nat) by (intro; omega). rewrite H. - simpl. reflexivity. - + exploit IHl2; eauto. intros. clear IHl2. - assert (forall n n', (n > n')%nat -> (n - n')%nat = S (n - n' - 1)%nat) by (intros; omega). - exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2. - assert ((length l2 > length l1)%nat) by omega. clear H2. - rewrite H0; auto. -Qed. + + + (* Very adapted from sistate_simub_correct ---> should delete sistate_simub_correct after *) Lemma hsistate_simub_correct dm f hst1 hst2: hsistate_simub dm f hst1 hst2 = OK tt -> forall ctx, hsistate_simu dm f hst1 hst2 ctx. Proof. - unfold hsistate_simub. intro. explore. unfold hsistate_simu. +(* unfold hsistate_simub. intro. explore. unfold hsistate_simu. intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) (HREF1 & HLREF1) (HREF2 & HLREF2) is1 SEMI. exploit hsiexits_simub_correct; eauto. intro ESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. @@ -1244,7 +1255,7 @@ Proof. - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). - congruence. } + congruence. } *) Admitted. (* destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 1299758a..76a9fde8 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1763,10 +1763,12 @@ Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sista Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (ctx: simu_proof_context f) (se1 se2: sistate_exit) := forall is1, - icontinue is1 = false -> +(* icontinue is1 = false -> *) ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> - (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx)) - = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)) + (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) + (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx)) + = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) + (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)) /\ exists is2, ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) /\ istate_simu f dm is1 is2. -- cgit From b402e1d2d3fc09a3686b3cde3337672777baa41e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 21 Jul 2020 18:28:42 +0200 Subject: Progress --- kvx/lib/RTLpathSE_impl.v | 164 ++++++++++++++++++++++++++++----------------- kvx/lib/RTLpathSE_theory.v | 17 +++-- 2 files changed, 109 insertions(+), 72 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 8c5b84a1..965f8578 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -138,34 +138,6 @@ Proof. repeat (constructor; auto). Qed. -(* Lemma hsilocal_simu_core_correct' dm f hst1 hst2 (ctx: simu_proof_context f) st1 st2: - hsilocal_simu_core hst1 hst2 -> - hsilocal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsilocal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> - exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) - /\ istate_simulive (fun _ => True) dm is1 is2. -Proof. - intros CORE HREF1 HREF2 is1 SEML. - exploit ssem_local_refines_hok; eauto. - intros HOK1. - exploit hsilocal_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. - destruct SEML as (PRE & MEMEQ & RSEQ). - eexists. constructor. - - constructor; [|constructor]. - + destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. - + destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _). - destruct CORE as (_ & _ & _ & MEMEQ3). - rewrite <- MEMEQ2; auto. rewrite hsi_smem_eval_correct. rewrite <- MEMEQ3. - erewrite smem_eval_preserved; [| eapply ctx]. - rewrite <- hsi_smem_eval_correct. rewrite MEMEQ1; auto. eassumption. - + intro r. destruct HREF2 as (_ & _ & A). destruct HREF1 as (_ & _ & B). - destruct CORE as (_ & _ & C & _). rewrite <- A; auto. rewrite hsi_sreg_eval_correct. - rewrite C. erewrite seval_preserved; [| eapply ctx]. rewrite <- hsi_sreg_eval_correct. - rewrite B; auto. - - apply istate_simulive_reflexive. -Qed. *) - Theorem hsilocal_simu_core_correct hst1 hst2 ge1 ge2 sp rs0 m0 rs m st1 st2: hsilocal_simu_core hst1 hst2 -> hsilocal_refines ge1 sp rs0 m0 hst1 st1 -> @@ -203,6 +175,14 @@ Definition hsiexit_simu_core dm f (hse1 hse2: hsistate_exit) := /\ hsi_scondargs hse1 = hsi_scondargs hse2 (* FIXME - should there be something about okvals ? *) /\ hsilocal_simu_core (hsi_elocal hse1) (hsi_elocal hse2). +Definition hsiexit_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := OK tt (* TODO *). + +Theorem hsiexit_simu_coreb_correct hse1 hse2 dm f: + hsiexit_simu_coreb dm f hse1 hse2 = OK tt -> + hsiexit_simu_core dm f hse1 hse2. +Proof. +Admitted. + (** NB: we split the refinement relation between a "static" part -- independendent of the initial context and a "dynamic" part -- that depends on it *) @@ -241,7 +221,7 @@ Theorem hsiexit_simu_core_correct dm f hse1 hse2 ctx: hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> hsiexit_simu dm f ctx hse1 hse2. Proof. - intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2 is1 (* ICONT *) SSEME. + intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2. exploit hsiexit_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. assert (SEVALC: seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond st1) (si_scondargs st1) (si_smem (si_elocal st1)) @@ -253,7 +233,8 @@ Proof. destruct SIMUC as (_ & _ & CONDEQ & ARGSEQ & LSIMU). destruct LSIMU as (_ & _ & _ & MEMEQ). rewrite CONDEQ. rewrite ARGSEQ. rewrite MEMEQ. erewrite <- seval_condition_preserved; eauto. eapply ctx. } - constructor; auto. + constructor; [assumption|]. + intros is1 SSEME. exists (mk_istate (icontinue is1) (si_ifso st2) (irs is1) (imem is1)). simpl. constructor. - constructor; [|constructor]. + rewrite <- SEVALC. destruct SSEME as (SCOND & _ & _). assumption. @@ -273,6 +254,17 @@ Proof. + destruct HREF2 as (_ & IFSO). congruence. Qed. +Remark hsiexit_simu_siexit dm f ctx hse1 hse2 se1 se2: + hsiexit_simu dm f ctx hse1 hse2 -> + hsiexit_refines_stat hse1 se1 -> + hsiexit_refines_stat hse2 se2 -> + hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> + hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> + siexit_simu dm f ctx se1 se2. +Proof. + auto. +Qed. + Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := | hsok_exits_nil: hsok_exits ge sp rs m nil | hsok_exits_cons: forall he lhe, @@ -286,6 +278,15 @@ Definition hsiexits_simu dm f (ctx: simu_proof_context f) lhse1 lhse2: Prop := Definition hsiexits_simu_core dm f lhse1 lhse2: Prop := list_forall2 (hsiexit_simu_core dm f) lhse1 lhse2. +(* TODO *) +Definition hsiexits_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := OK tt. + +Theorem hsiexits_simu_coreb_correct dm f lhse1 lhse2: + hsiexits_simu_coreb dm f lhse1 lhse2 = OK tt -> + hsiexits_simu_core dm f lhse1 lhse2. +Proof. +Admitted. + Theorem hsiexits_simu_core_correct dm f lhse1 lhse2 ctx: hsiexits_simu_core dm f lhse1 lhse2 -> hsok_exits (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse1 -> @@ -306,10 +307,18 @@ Definition hsiexits_refines_dyn ge sp rs0 m0 lhse se := Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. Definition hsistate_simu_core dm f (hse1 hse2: hsistate) := - hsi_pc hse1 = hsi_pc hse2 + dm ! (hsi_pc hse2) = Some (hsi_pc hse1) /\ list_forall2 (hsiexit_simu_core dm f) (hsi_exits hse1) (hsi_exits hse2) /\ hsilocal_simu_core (hsi_local hse1) (hsi_local hse2). +Definition hsistate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate) := OK tt. (* TODO *) + +Theorem hsistate_simu_coreb_correct dm f hse1 hse2: + hsistate_simu_coreb dm f hse1 hse2 = OK tt -> + hsistate_simu_core dm f hse1 hse2. +Proof. +Admitted. + Definition hsistate_refines_stat (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st /\ hsiexits_refines_stat (hsi_exits hst) (si_exits st). @@ -337,23 +346,26 @@ Proof. induction 1; [unfold all_fallthrough; contradiction|]. intros X ext INEXT. eapply all_fallthrough_revcons in X. destruct X as (SEVAL & ALLFU). apply IHlist_forall2 in ALLFU. -Admitted. -(* BROKEN! - destruct H as (_ & CONDSIMU). + destruct H as (CONDSIMU & _). inv INEXT; [|eauto]. erewrite <- CONDSIMU; eauto. Qed. -*) -Lemma hsiexits_simu_siexits dm f ctx lhse1 lhse2 lse1 lse2: +Lemma hsiexits_simu_siexits dm f ctx lhse1 lhse2: hsiexits_simu dm f ctx lhse1 lhse2 -> + forall lse1 lse2, hsiexits_refines_stat lhse1 lse1 -> hsiexits_refines_stat lhse2 lse2 -> hsiexits_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse1 lse1 -> hsiexits_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse2 lse2 -> siexits_simu dm f lse1 lse2 ctx. Proof. -Admitted. + induction 1. + - intros. inv H. inv H0. constructor. + - intros lse1 lse2 SREF1 SREF2 DREF1 DREF2. inv SREF1. inv SREF2. inv DREF1. inv DREF2. + constructor; [| eapply IHlist_forall2; eauto]. + eapply hsiexit_simu_siexit; eauto. +Qed. Lemma siexits_simu_all_fallthrough_upto dm f ctx lse1 lse2: siexits_simu dm f lse1 lse2 ctx -> forall ext1 lx1, @@ -444,7 +456,8 @@ Proof. - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). congruence. } - apply EXTSIMU in SSEME. destruct SSEME as (CONDEVAL & (is2 & SSEME2 & ISIMU)). + destruct EXTSIMU as (CONDEVAL & EXTSIMU). + apply EXTSIMU in SSEME. clear EXTSIMU. destruct SSEME as (is2 & SSEME2 & ISIMU). exists (mk_istate (icontinue is1) (ipc is2) (irs is2) (imem is2)). constructor. + unfold ssem_internal. simpl. rewrite ICONT. exists ext2, lx2. constructor; assumption. + unfold istate_simu in *. rewrite ICONT in *. destruct ISIMU as (path & PATHEQ & ISIMULIVE & DMEQ). @@ -452,12 +465,57 @@ Proof. exists path. repeat (constructor; auto). Qed. -Definition hsistate_simu_coreb (hst1 hst2: hsistate) := OK tt. +(* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', + ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) + +(* FIXME - too strong let's change it later.. *) +Definition hfinal_refines (hfv fv: sfval) := hfv = fv. -Theorem hsistate_simu_coreb_correct dm f hst1 hst2: - hsistate_simu_coreb hst1 hst2 = OK tt -> - hsistate_simu_core dm f hst1 hst2. +Remark hfinal_refines_snone: hfinal_refines Snone Snone. Proof. + reflexivity. +Qed. + +Record hsstate := { hinternal:> hsistate; hfinal: sfval }. + +Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := + hsistate_refines_stat (hinternal hst) (internal st) + /\ (forall ge sp rs0 m0, hsistate_refines_dyn ge sp rs0 m0 (hinternal hst) (internal st)) + /\ hfinal_refines (hfinal hst) (final st). + +Definition hsstate_simu_core (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := + hsistate_simu_core dm f (hinternal hst1) (hinternal hst2) + /\ hfinal hst1 = hfinal hst2. + +Definition hok ge sp rs m hst := hsok ge sp rs m (hinternal hst). + +Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := + forall st1 st2, + hsstate_refines hst1 st1 -> + hsstate_refines hst2 st2 -> sstate_simu dm f st1 st2 ctx. + +Remark sfval_simu_reflexive dm f ctx pc1 pc2 hf: + dm ! pc2 = Some pc1 -> + sfval_simu dm f pc1 pc2 ctx hf hf. +Proof. + intro DMEQ. destruct hf. + - constructor; auto. + - constructor; auto. +Abort. (* theorem not true !! *) + +Theorem hsstate_simu_core_correct dm f ctx hst1 hst2: + hsstate_simu_core dm f hst1 hst2 -> + hok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> + hsstate_simu dm f hst1 hst2 ctx. +Proof. + intros (SCORE & FSIMU) HOK1. intros st1 st2 HREF1 HREF2. + destruct HREF1 as (SREF1 & DREF1 & FREF1). destruct HREF2 as (SREF2 & DREF2 & FREF2). + assert (PCEQ: dm ! (hsi_pc hst2) = Some (hsi_pc hst1)) by apply SCORE. + eapply hsistate_simu_core_correct in SCORE; [|eassumption]. + constructor; [apply SCORE; auto|]. +(* rewrite <- FREF1. rewrite <- FREF2. rewrite FSIMU. (* FIXME - not that simple *) + apply sfval_simu_reflexive. + destruct SREF1 as (PC1 & _). destruct SREF2 as (PC2 & _). congruence. *) Admitted. Lemma hsistate_refines_stat_pceq st hst: @@ -819,22 +877,6 @@ Lemma hsiexec_path_correct_dyn ge sp rs0 m0 ps f hst hst' st st': Proof. Admitted. -(* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', - ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) - -(* FIXME - too strong let's change it later.. *) -Definition hfinal_refines (hfv fv: sfval) := hfv = fv. - -Remark hfinal_refines_snone: hfinal_refines Snone Snone. -Proof. -Admitted. - -Record hsstate := { hinternal:> hsistate; hfinal: sfval }. - -Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := - hsistate_refines_stat (hinternal hst) (internal st) - /\ (forall ge sp rs0 m0, hsistate_refines_dyn ge sp rs0 m0 (hinternal hst) (internal st)) - /\ hfinal_refines (hfinal hst) (final st). Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := match i with @@ -1269,10 +1311,6 @@ Admitted. Qed. *) -Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := - forall st1 st2, - hsstate_refines hst1 st1 -> - hsstate_refines hst2 st2 -> sstate_simu f dm st1 st2 ctx. Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := @@ -1540,7 +1578,7 @@ Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hs Lemma hsstate_simub_correct dm f hst1 hst2 ctx: hsstate_simub dm f hst1 hst2 = OK tt -> - hsstate_simu f dm hst1 hst2 ctx. + hsstate_simu dm f hst1 hst2 ctx. Proof. unfold hsstate_simub. intro. explore. eapply sfval_simub_correct in EQ1. eapply hsistate_simub_correct in EQ. diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 76a9fde8..81351f6a 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -1757,21 +1757,20 @@ Here are intermediate definitions. *) Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) (sl1 sl2: sistate_local) (ctx: simu_proof_context f): Prop := - forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) sl1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> + forall is1, ssem_local (the_ge1 ctx) (the_sp ctx) sl1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) -> exists is2, ssem_local (the_ge2 ctx) (the_sp ctx) sl2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) /\ istate_simu f dm is1 is2. Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) (ctx: simu_proof_context f) (se1 se2: sistate_exit) := - forall is1, -(* icontinue is1 = false -> *) - ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> - (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) + (seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond se1) (si_scondargs se1) (si_smem (si_elocal se1)) (the_rs0 ctx) (the_m0 ctx)) - = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) + = (seval_condition (the_ge2 ctx) (the_sp ctx) (si_cond se2) (si_scondargs se2) (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)) - /\ exists is2, - ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) - /\ istate_simu f dm is1 is2. + /\ forall is1, + ssem_exit (the_ge1 ctx) (the_sp ctx) se1 (the_rs0 ctx) (the_m0 ctx) (irs is1) (imem is1) (ipc is1) -> + exists is2, + ssem_exit (the_ge2 ctx) (the_sp ctx) se2 (the_rs0 ctx) (the_m0 ctx) (irs is2) (imem is2) (ipc is2) + /\ istate_simu f dm is1 is2. Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) (lse1 lse2: list sistate_exit) (ctx: simu_proof_context f) := list_forall2 (siexit_simu dm f ctx) lse1 lse2. -- cgit From 70618966a9ba8b3d0b5c5ec970ec71f1629fbedf Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 22 Jul 2020 09:09:45 +0200 Subject: start a junk implementation of the pre-pass verifier --- configure | 2 +- kvx/lib/RTLpathSE_impl_junk.v | 589 ++++++++++++++++++++++++++++++++++++++++++ kvx/lib/RTLpathScheduler.v | 2 +- 3 files changed, 591 insertions(+), 2 deletions(-) create mode 100644 kvx/lib/RTLpathSE_impl_junk.v diff --git a/configure b/configure index 959d0590..c700557d 100755 --- a/configure +++ b/configure @@ -845,7 +845,7 @@ EXECUTE=kvx-cluster --syscall=libstd_scalls.so -- CFLAGS= -D __KVX_COS__ SIMU=kvx-cluster -- BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\ - RTLpath.v RTLpathproof.v RTLpathLivegen.v RTLpathLivegenproof.v RTLpathSE_theory.v RTLpathSE_impl.v RTLpathScheduler.v RTLpathSchedulerproof.v\\ + RTLpath.v RTLpathproof.v RTLpathLivegen.v RTLpathLivegenproof.v RTLpathSE_theory.v RTLpathSE_impl_junk.v RTLpathScheduler.v RTLpathSchedulerproof.v\\ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v Asmblockprops.v\\ ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\ diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v new file mode 100644 index 00000000..8ae21a2d --- /dev/null +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -0,0 +1,589 @@ +(** Implementation and refinement of the symbolic execution + +* a JUNK VERSION WITHOUT ANY FORMAL PROOF !!! + + *) + +Require Import Coqlib Maps Floats. +Require Import AST Integers Values Events Memory Globalenvs Smallstep. +Require Import Op Registers. +Require Import RTL RTLpath. +Require Import Errors Duplicate. +Require Import RTLpathSE_theory. +Require Import Axioms. + +Local Open Scope error_monad_scope. +Local Open Scope option_monad_scope. + +Require Export Impure.ImpHCons. +Export Notations. +Import HConsing. + + +(* +(** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) + +(** * Implementation of local symbolic internal states +TODO: use refined symbolic values instead of those from RTLpathSE_theory. +*) + +(** name : Hash-consed Symbolic Internal state local. Later on we will use the + refinement to introduce hash consing *) +Record hsistate_local := + { + (** [hsi_lsmem] represents the list of smem symbolic evaluations. + The first one of the list is the current smem *) + hsi_lsmem:> list smem; + (** For the values in registers: + 1) we store a list of sval evaluations + 2) we encode the symbolic regset by a PTree *) + hsi_ok_lsval: list sval; + hsi_sreg:> PTree.t sval + }. + +Definition hsi_sreg_get (hst: PTree.t sval) r: sval := + match PTree.get r hst with + | None => Sinput r + | Some sv => sv + end. + + +Definition hsi_smem_get (hst: list smem): smem := + match hst with + | nil => Sinit + | sm::_ => sm + end. + +(* Syntax and semantics of symbolic exit states *) +(* TODO: add hash-consing *) +Record hsistate_exit := mk_hsistate_exit + { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. + + +(** * Syntax and Semantics of symbolic internal state *) +Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. + +(** ** Assignment of memory *) +Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := + {| hsi_lsmem := sm::hsi_lsmem hst; + hsi_ok_lsval := hsi_ok_lsval hst; + hsi_sreg:= hsi_sreg hst + |}. + +(** ** Assignment of local state *) + +Definition hsist_set_local (hst: hsistate) (pc: node) (hnxt: hsistate_local): hsistate := + {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= hnxt |}. + +(** ** 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) (lsv: list sval) (sm: smem): sval := + match rsv with + | Rop op => Sop op (list_sval_inj lsv) sm + | Rload trap chunk addr => Sload sm trap chunk addr (list_sval_inj lsv) + end. +Coercion root_apply: root_sval >-> Funclass. + +Local Open Scope lazy_bool_scope. + +(* NB: return [false] if the rsv cannot fail *) +Definition may_trap (rsv: root_sval) (lsv: list sval) (sm: smem): bool := + match rsv with + | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lsv) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) + | Rload TRAP _ _ => true + | _ => false + end. + +(* simplify a symbolic value before assignment to a register *) +Definition simplify (rsv: root_sval) lsv sm: sval := + match rsv with + | Rload TRAP chunk addr => Sload sm NOTRAP chunk addr (list_sval_inj lsv) + | Rop op => + match is_move_operation op lsv with + | Some arg => arg (* optimization of Omove *) + | None => + if op_depends_on_memory op then + rsv lsv sm + else + Sop op (list_sval_inj lsv) Sinit (* magically remove the dependency on sm ! *) + end + | _ => rsv lsv sm + end. + +Definition red_PTree_set (r:reg) (sv: sval) (hst: PTree.t sval): PTree.t sval := + match sv with + | Sinput r' => + if Pos.eq_dec r r' + then PTree.remove r' hst + else PTree.set r sv hst + | _ => PTree.set r sv hst + end. + +Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := + {| hsi_lsmem := hsi_lsmem hst; + hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; + hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. + +(** ** Execution of one instruction *) + +Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := + match i with + | Inop pc' => + Some (hsist_set_local hst pc' hst.(hsi_local)) + | Iop op args dst pc' => + let prev := hst.(hsi_local) in + let vargs := List.map (hsi_sreg_get prev) args in + let next := hslocal_set_sreg prev dst (Rop op) vargs (hsi_smem_get prev) in + Some (hsist_set_local hst pc' next) + | Iload trap chunk addr args dst pc' => + let prev := hst.(hsi_local) in + let vargs := List.map (hsi_sreg_get prev) args in + let next := hslocal_set_sreg prev dst (Rload trap chunk addr) vargs (hsi_smem_get prev) in + Some (hsist_set_local hst pc' next) + | Istore chunk addr args src pc' => + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in + Some (hsist_set_local hst pc' next) + | Icond cond args ifso ifnot _ => + let prev := hst.(hsi_local) in + let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in + Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |} + | _ => None (* TODO jumptable ? *) + end. + +Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := + match path with + | O => Some hst + | S p => + SOME i <- (fn_code f)!(hst.(hsi_pc)) IN + SOME hst1 <- hsiexec_inst i hst IN + hsiexec_path p f hst1 + end. + +(* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', + ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) + +(* FIXME - too strong let's change it later.. *) +Definition hfinal_refines (hfv fv: sfval) := hfv = fv. + +Record hsstate := { hinternal:> hsistate; hfinal: sfval }. + +Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := + match i with + | Icall sig ros args res pc => + let svos := sum_left_map (hsi_sreg_get prev) ros in + let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + Scall sig svos sargs res pc + | Itailcall sig ros args => + let svos := sum_left_map (hsi_sreg_get prev) ros in + let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + Stailcall sig svos sargs + | Ibuiltin ef args res pc => + let sargs := List.map (builtin_arg_inj (hsi_sreg_get prev)) args in + Sbuiltin ef sargs res pc + | Ireturn or => + let sor := SOME r <- or IN Some (hsi_sreg_get prev r) in + Sreturn sor + | Ijumptable reg tbl => + let sv := hsi_sreg_get prev reg in + Sjumptable sv tbl + | _ => Snone + end. + +Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. + +Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. + +Definition hsexec (f: function) (pc:node): option hsstate := + SOME path <- (fn_path f)!pc IN + SOME hst <- hsiexec_path path.(psize) f (init_hsistate pc) IN + SOME i <- (fn_code f)!(hst.(hsi_pc)) IN + Some (match hsiexec_inst i hst with + | Some hst' => {| hinternal := hst'; hfinal := Snone |} + | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} + end). + +(** * The simulation test of concrete symbolic execution *) + +(* TODO - generalize it with a list of registers to test ? *) +Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. + +Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := + match dm ! tn with + | None => Error (msg "revmap_check_single: no mapping for tn") + | Some res => if (Pos.eq_dec n res) then OK tt + else Error (msg "revmap_check_single: n and res do not match") + end. + +(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) +Fixpoint sval_simub (sv1 sv2: sval) := + match sv1 with + | Sinput r => + match sv2 with + | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") + | _ => Error (msg "sval_simub: Sinput expected") + end + | Sop op lsv sm => + match sv2 with + | Sop op' lsv' sm' => + if (eq_operation op op') then + do _ <- list_sval_simub lsv lsv'; + smem_simub sm sm' + else Error (msg "sval_simub: different operations in Sop") + | _ => Error (msg "sval_simub: Sop expected") + end + | Sload sm trap chk addr lsv => + match sv2 with + | Sload sm' trap' chk' addr' lsv' => + if (trapping_mode_eq trap trap') then + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + list_sval_simub lsv lsv' + else Error (msg "sval_simub: addressings do not match") + else Error (msg "sval_simub: chunks do not match") + else Error (msg "sval_simub: trapping modes do not match") + (* FIXME - should be refined once we introduce non trapping loads *) + | _ => Error (msg "sval_simub: Sload expected") + end + end +with list_sval_simub (lsv1 lsv2: list_sval) := + match lsv1 with + | Snil => + match lsv2 with + | Snil => OK tt + | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") + end + | Scons sv1 lsv1 => + match lsv2 with + | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") + | Scons sv2 lsv2 => + do _ <- sval_simub sv1 sv2; + list_sval_simub lsv1 lsv2 + end + end +with smem_simub (sm1 sm2: smem) := + match sm1 with + | Sinit => + match sm2 with + | Sinit => OK tt + | _ => Error (msg "smem_simub: Sinit expected") + end + | Sstore sm chk addr lsv sv => + match sm2 with + | Sstore sm' chk' addr' lsv' sv' => + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + do _ <- list_sval_simub lsv lsv'; + sval_simub sv sv' + else Error (msg "smem_simub: addressings do not match") + else Error (msg "smem_simub: chunks not match") + | _ => Error (msg "smem_simub: Sstore expected") + end + end. + +Lemma sval_simub_correct sv1: forall sv2, + sval_simub sv1 sv2 = OK tt -> sv1 = sv2. +Proof. + induction sv1 using sval_mut with + (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) + (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. + (* Sinput *) + - destruct sv2; try discriminate. + destruct (Pos.eq_dec r r0); [congruence|discriminate]. + (* Sop *) + - destruct sv2; try discriminate. + destruct (eq_operation _ _); [|discriminate]. subst. + intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sload *) + - destruct sv2; try discriminate. + destruct (trapping_mode_eq _ _); [|discriminate]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. + congruence. + (* Snil *) + - destruct lsv2; [|discriminate]. congruence. + (* Scons *) + - destruct lsv2; [discriminate|]. intro. explore. + apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sinit *) + - destruct sm2; [|discriminate]. congruence. + (* Sstore *) + - destruct sm2; [discriminate|]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. + assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. + congruence. +Qed. + +Lemma list_sval_simub_correct lsv1: forall lsv2, + list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. +Proof. + induction lsv1; simpl; auto. + - destruct lsv2; try discriminate. reflexivity. + - destruct lsv2; try discriminate. intro. explore. + apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. + congruence. +Qed. + +Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := + if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then + do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); + do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); + revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) + else Error (msg "siexit_simub: conditions do not match") +. + +Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := + match lhse1 with + | nil => + match lhse2 with + | nil => OK tt + | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") + end + | hse1 :: lhse1 => + match lhse2 with + | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") + | hse2 :: lhse2 => + do _ <- hsiexit_simub dm f hse1 hse2; + do _ <- hsiexits_simub dm f lhse1 lhse2; + OK tt + end + end. + +Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := + do _ <- hsilocal_simub dm f (hsi_local hst1) (hsi_local hst2); + do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); + OK tt. + +Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := + match ln with + | nil => + match ln' with + | nil => OK tt + | _ => Error (msg "revmap_check_list: lists have different lengths") + end + | n::ln => + match ln' with + | nil => Error (msg "revmap_check_list: lists have different lengths") + | n'::ln' => do _ <- revmap_check_single dm n n'; revmap_check_list dm ln ln' + end + end. + +Definition svos_simub (svos1 svos2: sval + ident) := + match svos1 with + | inl sv1 => + match svos2 with + | inl sv2 => sval_simub sv1 sv2 + | _ => Error (msg "svos_simub: expected sval") + end + | inr id1 => + match svos2 with + | inr id2 => + if (ident_eq id1 id2) then OK tt + else Error (msg "svos_simub: ids do not match") + | _ => Error (msg "svos_simub: expected id") + end + end. + +Fixpoint builtin_arg_simub (bs bs': builtin_arg sval) := + match bs with + | BA sv => + match bs' with + | BA sv' => sval_simub sv sv' + | _ => Error (msg "builtin_arg_simub: BA expected") + end + | BA_int n => + match bs' with + | BA_int n' => if (Integers.int_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_int expected") + end + | BA_long n => + match bs' with + | BA_long n' => if (int64_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") + | _ => Error (msg "buitin_arg_simub: BA_long expected") + end + | BA_float f => + match bs' with + | BA_float f' => if (float_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_float expected") + end + | BA_single f => + match bs' with + | BA_single f' => if (float32_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") + | _ => Error (msg "builtin_arg_simub: BA_single expected") + end + | BA_loadstack chk ptr => + match bs' with + | BA_loadstack chk' ptr' => + if (chunk_eq chk chk') then + if (ptrofs_eq ptr ptr') then OK tt + else Error (msg "builtin_arg_simub: ptr do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadstack expected") + end + | BA_addrstack ptr => + match bs' with + | BA_addrstack ptr' => if (ptrofs_eq ptr ptr') then OK tt else Error (msg "builtin_arg_simub: ptr do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrstack expected") + end + | BA_loadglobal chk id ofs => + match bs' with + | BA_loadglobal chk' id' ofs' => + if (chunk_eq chk chk') then + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + else Error (msg "builtin_arg_simub: chunks do not match") + | _ => Error (msg "builtin_arg_simub: BA_loadglobal expected") + end + | BA_addrglobal id ofs => + match bs' with + | BA_addrglobal id' ofs' => + if (ident_eq id id') then + if (ptrofs_eq ofs ofs') then OK tt + else Error (msg "builtin_arg_simub: offsets do not match") + else Error (msg "builtin_arg_simub: identifiers do not match") + | _ => Error (msg "builtin_arg_simub: BA_addrglobal expected") + end + | BA_splitlong lo hi => + match bs' with + | BA_splitlong lo' hi' => do _ <- builtin_arg_simub lo lo'; builtin_arg_simub hi hi' + | _ => Error (msg "builtin_arg_simub: BA_splitlong expected") + end + | BA_addptr b1 b2 => + match bs' with + | BA_addptr b1' b2' => do _ <- builtin_arg_simub b1 b1'; builtin_arg_simub b2 b2' + | _ => Error (msg "builtin_arg_simub: BA_addptr expected") + end + end. + +Fixpoint list_builtin_arg_simub lbs1 lbs2 := + match lbs1 with + | nil => + match lbs2 with + | nil => OK tt + | _ => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs2 is bigger)") + end + | bs1::lbs1 => + match lbs2 with + | nil => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs1 is bigger)") + | bs2::lbs2 => + do _ <- builtin_arg_simub bs1 bs2; + list_builtin_arg_simub lbs1 lbs2 + end + end. + +(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) +Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := + match fv1 with + | Snone => + match fv2 with + | Snone => revmap_check_single dm pc1 pc2 + | _ => Error (msg "sfval_simub: Snone expected") + end + | Scall sig svos lsv res pc1 => + match fv2 with + | Scall sig2 svos2 lsv2 res2 pc2 => + do _ <- revmap_check_single dm pc1 pc2; + if (signature_eq sig sig2) then + if (Pos.eq_dec res res2) then + do _ <- svos_simub svos svos2; + list_sval_simub lsv lsv2 + else Error (msg "sfval_simub: Scall res do not match") + else Error (msg "sfval_simub: Scall different signatures") + | _ => Error (msg "sfval_simub: Scall expected") + end + | Stailcall sig svos lsv => + match fv2 with + | Stailcall sig' svos' lsv' => + if (signature_eq sig sig') then + do _ <- svos_simub svos svos'; + list_sval_simub lsv lsv' + else Error (msg "sfval_simub: signatures do not match") + | _ => Error (msg "sfval_simub: Stailcall expected") + end + | Sbuiltin ef lbs br pc => + match fv2 with + | Sbuiltin ef' lbs' br' pc' => + if (external_function_eq ef ef') then + if (builtin_res_eq_pos br br') then + do _ <- revmap_check_single dm pc pc'; + list_builtin_arg_simub lbs lbs' + else Error (msg "sfval_simub: builtin res do not match") + else Error (msg "sfval_simub: external functions do not match") + | _ => Error (msg "sfval_simub: Sbuiltin expected") + end + | Sjumptable sv ln => + match fv2 with + | Sjumptable sv' ln' => + do _ <- revmap_check_list dm ln ln'; sval_simub sv sv' + | _ => Error (msg "sfval_simub: Sjumptable expected") + end + | Sreturn osv => + match fv2 with + | Sreturn osv' => + match osv with + | Some sv => + match osv' with + | Some sv' => sval_simub sv sv' + | None => Error (msg "sfval_simub sv' expected") + end + | None => + match osv' with + | Some sv' => Error (msg "None expected") + | None => OK tt + end + end + | _ => Error (msg "sfval_simub: Sreturn expected") + end + end. + +Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := + do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); + do u2 <- sfval_simub dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2); + OK tt. + +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := + let (pc2, pc1) := m in + match (hsexec f pc1) with + | None => Error (msg "simu_check_single: hsexec f pc1 failed") + | Some hst1 => + match (hsexec tf pc2) with + | None => Error (msg "simu_check_single: hsexec tf pc2 failed") + | Some hst2 => hsstate_simub dm f hst1 hst2 + end + end. + +Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := + match lm with + | nil => OK tt + | m :: lm => do u1 <- simu_check_single dm f tf m; + do u2 <- simu_check_rec dm f tf lm; + OK tt + end. + +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := + simu_check_rec dm f tf (PTree.elements dm). +*) + + +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := OK tt. + +Lemma simu_check_correct dm f tf: + simu_check dm f tf = OK tt -> + forall pc1 pc2, dm ! pc2 = Some pc1 -> + sexec_simu dm f tf pc1 pc2. +Admitted. \ No newline at end of file diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v index 5a1c7f1c..b70c1685 100644 --- a/kvx/lib/RTLpathScheduler.v +++ b/kvx/lib/RTLpathScheduler.v @@ -6,7 +6,7 @@ This module is inspired from [Duplicate] and [Duplicateproof] Require Import AST Linking Values Maps Globalenvs Smallstep Registers. Require Import Coqlib Maps Events Errors Op. -Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory RTLpathSE_impl. +Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory RTLpathSE_impl_junk. Notation "'ASSERT' A 'WITH' MSG 'IN' B" := (if A then B else Error (msg MSG)) -- cgit From f37341ec18aad35ab87c407345a4b82ffd2adacb Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Jul 2020 10:45:46 +0200 Subject: hfinal_simu_core --- kvx/lib/RTLpathSE_impl.v | 71 +++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 70 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 965f8578..bfa66d8b 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -468,7 +468,7 @@ Qed. (* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) -(* FIXME - too strong let's change it later.. *) +(* FIXME - might be too strong, let's change it later.. *) Definition hfinal_refines (hfv fv: sfval) := hfv = fv. Remark hfinal_refines_snone: hfinal_refines Snone Snone. @@ -476,6 +476,75 @@ Proof. reflexivity. Qed. +Definition hfinal_simu_core (dm: PTree.t node) (f: RTLpath.function) (hf1 hf2: sfval): Prop := + match hf1 with + | Scall sig1 svos1 lsv1 res1 pc1 => + match hf2 with + | Scall sig2 svos2 lsv2 res2 pc2 => + dm ! pc2 = Some pc1 /\ sig1 = sig2 /\ svos1 = svos2 /\ lsv1 = lsv2 /\ res1 = res2 + | _ => False + end + | Sbuiltin ef1 lbs1 br1 pc1 => + match hf2 with + | Sbuiltin ef2 lbs2 br2 pc2 => + dm ! pc2 = Some pc1 /\ ef1 = ef2 /\ lbs1 = lbs2 /\ br1 = br2 + | _ => False + end + | Sjumptable sv1 lpc1 => + match hf2 with + | Sjumptable sv2 lpc2 => + ptree_get_list dm lpc2 = Some lpc1 /\ sv1 = sv2 + | _ => False + end + (* Snone, Stailcall, Sreturn *) + | _ => hf1 = hf2 + end. + +Definition hfinal_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hf1 hf2: sfval) := OK tt. (* TODO *) + +Theorem hfinal_simu_coreb_correct dm f hf1 hf2: + hfinal_simu_coreb dm f hf1 hf2 = OK tt -> + hfinal_simu_core dm f hf1 hf2. +Proof. +Admitted. + +Lemma svident_simu_refl f ctx s: + svident_simu f ctx s s. +Proof. + destruct s; constructor; [| reflexivity]. + erewrite <- seval_preserved; [| eapply ctx]. constructor. +Qed. + +Theorem hfinal_simu_core_correct dm f ctx opc1 opc2 hf1 hf2 f1 f2: + hfinal_simu_core dm f hf1 hf2 -> + hfinal_refines hf1 f1 -> + hfinal_refines hf2 f2 -> + dm ! opc2 = Some opc1 -> + sfval_simu dm f opc1 opc2 ctx f1 f2. +Proof. + intros CORE FREF1 FREF2 OPCEQ. + rewrite <- FREF1. rewrite <- FREF2. clear FREF1. clear FREF2. (* FIXME - to change when the refinement is more complex *) + destruct hf1. + (* Snone *) + - simpl in CORE. rewrite <- CORE. constructor. assumption. + (* Scall *) + - simpl in CORE. destruct hf2; try contradiction. destruct CORE as (PCEQ & ? & ? & ? & ?). subst. + constructor; [assumption | apply svident_simu_refl|]. + erewrite <- list_sval_eval_preserved; [| eapply ctx]. constructor. + (* Stailcall *) + - simpl in CORE. rewrite <- CORE. constructor; [apply svident_simu_refl|]. + erewrite <- list_sval_eval_preserved; [| eapply ctx]. constructor. + (* Sbuiltin *) + - simpl in CORE. destruct hf2; try contradiction. destruct CORE as (PCEQ & ? & ? & ?). subst. + constructor; [assumption | reflexivity]. + (* Sjumptable *) + - simpl in CORE. destruct hf2; try contradiction. destruct CORE as (PCEQ & ?). subst. + constructor; [assumption|]. + erewrite <- seval_preserved; [| eapply ctx]. constructor. + (* Sreturn *) + - simpl in CORE. subst. constructor. +Qed. + Record hsstate := { hinternal:> hsistate; hfinal: sfval }. Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := -- cgit From af410660b53a6e91a0be2457d6dc5adb1d332d16 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Jul 2020 11:27:01 +0200 Subject: Branching simu_coreb --- kvx/lib/RTLpathSE_impl.v | 269 ++++++++++------------------------------------- 1 file changed, 58 insertions(+), 211 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index bfa66d8b..f25143fd 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -205,6 +205,14 @@ Definition hsiexit_simu dm f (ctx: simu_proof_context f) hse1 hse2: Prop := fora hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> siexit_simu dm f ctx se1 se2. +(* Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := + if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then + do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); + do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); + revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) + else Error (msg "siexit_simub: conditions do not match") +. *) + Lemma hsiexit_simu_core_nofail dm f hse1 hse2 ge1 ge2 sp rs m: hsiexit_simu_core dm f hse1 hse2 -> (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> @@ -278,6 +286,38 @@ Definition hsiexits_simu dm f (ctx: simu_proof_context f) lhse1 lhse2: Prop := Definition hsiexits_simu_core dm f lhse1 lhse2: Prop := list_forall2 (hsiexit_simu_core dm f) lhse1 lhse2. +(* Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := + match lhse1 with + | nil => + match lhse2 with + | nil => OK tt + | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") + end + | hse1 :: lhse1 => + match lhse2 with + | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") + | hse2 :: lhse2 => + do _ <- hsiexit_simub dm f hse1 hse2; + do _ <- hsiexits_simub dm f lhse1 lhse2; + OK tt + end + end. *) + +(* Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, + hsiexits_simub dm f lhse1 lhse2 = OK tt -> + hsiexits_simu dm f ctx lhse1 lhse2. +Proof. +(* induction lhse1. + - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. + inv HEREFS1. inv HEREFS2. constructor. + - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. + fold hsiexits_simub in EQ1. + eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. + intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. + apply EQ1; assumption. *) +Admitted. + *) + (* TODO *) Definition hsiexits_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := OK tt. @@ -554,7 +594,15 @@ Definition hsstate_refines (hst: hsstate) (st:sstate): Prop := Definition hsstate_simu_core (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := hsistate_simu_core dm f (hinternal hst1) (hinternal hst2) - /\ hfinal hst1 = hfinal hst2. + /\ hfinal_simu_core dm f (hfinal hst1) (hfinal hst2). + +Definition hsstate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) + +Theorem hsstate_simu_coreb_correct dm f hst1 hst2: + hsstate_simu_coreb dm f hst1 hst2 = OK tt -> + hsstate_simu_core dm f hst1 hst2. +Proof. +Admitted. Definition hok ge sp rs m hst := hsok ge sp rs m (hinternal hst). @@ -563,15 +611,6 @@ Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := hsstate_refines hst1 st1 -> hsstate_refines hst2 st2 -> sstate_simu dm f st1 st2 ctx. -Remark sfval_simu_reflexive dm f ctx pc1 pc2 hf: - dm ! pc2 = Some pc1 -> - sfval_simu dm f pc1 pc2 ctx hf hf. -Proof. - intro DMEQ. destruct hf. - - constructor; auto. - - constructor; auto. -Abort. (* theorem not true !! *) - Theorem hsstate_simu_core_correct dm f ctx hst1 hst2: hsstate_simu_core dm f hst1 hst2 -> hok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> @@ -581,11 +620,11 @@ Proof. destruct HREF1 as (SREF1 & DREF1 & FREF1). destruct HREF2 as (SREF2 & DREF2 & FREF2). assert (PCEQ: dm ! (hsi_pc hst2) = Some (hsi_pc hst1)) by apply SCORE. eapply hsistate_simu_core_correct in SCORE; [|eassumption]. + eapply hfinal_simu_core_correct in FSIMU; eauto. constructor; [apply SCORE; auto|]. -(* rewrite <- FREF1. rewrite <- FREF2. rewrite FSIMU. (* FIXME - not that simple *) - apply sfval_simu_reflexive. - destruct SREF1 as (PC1 & _). destruct SREF2 as (PC2 & _). congruence. *) -Admitted. + destruct SREF1 as (PC1 & _). destruct SREF2 as (PC2 & _). rewrite <- PC1. rewrite <- PC2. + eapply FSIMU. +Qed. Lemma hsistate_refines_stat_pceq st hst: hsistate_refines_stat hst st -> @@ -1077,35 +1116,6 @@ Qed. (** * The simulation test of concrete symbolic execution *) -(* TODO - generalize it with a list of registers to test ? *) -Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. - -(* (* TODO: voilà les conditions que hsilocal_simub doit garantir *) -Definition hsilocal_simu_core (hst1 hst2: hsistate_local):= - forall sm, List.In sm (hsi_lsmem hst2) -> List.In sm (hsi_lsmem hst1) - /\ forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1) - /\ forall r, hsi_sreg_get hst2 r = hsi_sreg_get hst1 r. - -*) - -Definition hsilocal_simu dm f (hst1 hst2: hsistate_local) (ctx: simu_proof_context f): Prop := forall st1 st2, - hsilocal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsilocal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - silocal_simu dm f st1 st2 ctx. - -Lemma hsilocal_simub_correct dm f hst1 hst2: - hsilocal_simub dm f hst1 hst2 = OK tt -> - forall ctx, hsilocal_simu dm f hst1 hst2 ctx. -Proof. -Admitted. - -Lemma hsilocal_simub_eq dm f hsl1 hsl2: - hsilocal_simub dm f hsl1 hsl2 = OK tt -> - hsi_sreg_get (hsi_sreg hsl1) = hsi_sreg_get (hsi_sreg hsl2) (* FIXME: a bit too strong ? *) - /\ hsi_smem_get (hsi_lsmem hsl1) = hsi_smem_get (hsi_lsmem hsl2). -Proof. -Admitted. - Definition revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := match dm ! tn with | None => Error (msg "revmap_check_single: no mapping for tn") @@ -1234,154 +1244,8 @@ Proof. congruence. Qed. -Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := - if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then - do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); - do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); - revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) - else Error (msg "siexit_simub: conditions do not match") -. - -(* Definition hsiexit_simu dm f hse1 hse2 (ctx: simu_proof_context f): Prop := forall se1 se2, - hsiexit_refines_stat hse1 se1 -> - hsiexit_refines_stat hse2 se2 -> - hsiexit_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 se1 -> - hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> - siexit_simu dm f ctx se1 se2. *) - Local Hint Resolve genv_match ssem_local_refines_hok: core. -Lemma hsiexit_simub_correct dm f hse1 hse2: - hsiexit_simub dm f hse1 hse2 = OK tt -> - forall ctx, hsiexit_simu dm f ctx hse1 hse2. -Proof. -(* unfold hsiexit_simub. intro. explore. - apply list_sval_simub_correct in EQ0. exploit hsilocal_simub_eq; eauto. - intros (SREGEQ & SMEMEQ) ctx. - eapply hsilocal_simub_correct in EQ2. - apply revmap_check_single_correct in EQ3. - intros se1 se2 HIFSO1 HIFSO2 (HLREF1 & HCOND1) (HLREF2 & HCOND2). - unfold hsiexit_refines_stat in *. - intros is1 (SCOND & SLOCAL & SIFSO). - constructor. - - rewrite <- HCOND1; eauto. - rewrite <- HCOND2; eauto. - rewrite SMEMEQ, EQ0, e. - eapply seval_condition_preserved; eauto. - assert (X: hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse1) -> - hsok_local (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse2)). - + admit. (* TODO: on doit savoir ça du fait que l'état symbolique de hse2 simule celui de hse1 *) - + eauto. - - let H0 := fresh "SLOCAL" in assert (H0 := SLOCAL). - eapply EQ2 in SLOCAL0; eauto. destruct SLOCAL0 as (is2 & SLOCAL' & ISIMU). - destruct is2 as [icont ipc irs imem]. simpl in *. - exists (mk_istate icont (si_ifso se2) irs imem). simpl. - constructor; auto. - + repeat (constructor; auto). - rewrite <- HCOND2; eauto. - rewrite <- e, <- EQ0, <- SMEMEQ. - erewrite seval_condition_preserved; eauto. - rewrite HCOND1; eauto. - + unfold istate_simu in *. destruct (icontinue is1) eqn:ICONT. - * destruct ISIMU as (A & B & C). constructor; auto. - * destruct ISIMU as (path & PATHEQ & ISIMU & REVEQ). exists path. - repeat (constructor; auto). simpl in *. congruence. *) -Admitted. - -Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := - match lhse1 with - | nil => - match lhse2 with - | nil => OK tt - | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") - end - | hse1 :: lhse1 => - match lhse2 with - | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") - | hse2 :: lhse2 => - do _ <- hsiexit_simub dm f hse1 hse2; - do _ <- hsiexits_simub dm f lhse1 lhse2; - OK tt - end - end. - -Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, - hsiexits_simub dm f lhse1 lhse2 = OK tt -> - hsiexits_simu dm f ctx lhse1 lhse2. -Proof. -(* induction lhse1. - - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. - inv HEREFS1. inv HEREFS2. constructor. - - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. - fold hsiexits_simub in EQ1. - eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. - intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. - apply EQ1; assumption. *) -Admitted. - -Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := - do _ <- hsilocal_simub dm f (hsi_local hst1) (hsi_local hst2); - do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); - OK tt. - -(* Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): Prop := forall st1 st2, - hsistate_refines_stat hst1 st1 -> - hsistate_refines_stat hst2 st2 -> - hsistate_refines_dyn (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 st1 -> - hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> - sistate_simu dm f st1 st2 ctx. *) - - - - - - - - - -(* Very adapted from sistate_simub_correct ---> should delete sistate_simub_correct after *) -Lemma hsistate_simub_correct dm f hst1 hst2: - hsistate_simub dm f hst1 hst2 = OK tt -> - forall ctx, hsistate_simu dm f hst1 hst2 ctx. -Proof. -(* unfold hsistate_simub. intro. explore. unfold hsistate_simu. - intros ctx st1 st2 (PCEQ1 & HEREFD1) (PCEQ2 & HEREFD2) (HREF1 & HLREF1) (HREF2 & HLREF2) is1 SEMI. - exploit hsiexits_simub_correct; eauto. intro ESIMU. - unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - - destruct SEMI as (SSEML & PCEQ & ALLFU). - exploit hsilocal_simub_correct; eauto. intros (is2 & SSEML' & ISIMU). - destruct is2 as [icont2 ipc2 irs2 imem2]. simpl in *. - exists (mk_istate icont2 (si_pc st2) irs2 imem2). constructor; auto. - + unfold ssem_internal. simpl. - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. - destruct ISIMU as (CONTEQ & RSEQ & MEMEQ). simpl in *. - rewrite <- CONTEQ. rewrite ICONT. - constructor; [|constructor]; [eassumption | auto | eapply siexits_simu_all_fallthrough; eauto]. - + unfold istate_simu in *; simpl in *. rewrite ICONT in ISIMU. rewrite ICONT. - destruct ISIMU as (A & B & C). simpl in *. constructor; auto. - - destruct SEMI as (ext & lx & SSEME & ALLFU). - exploit siexits_simu_all_fallthrough_upto; eauto. intros (ext2 & lx2 & ALLFU2 & LENEQ). - assert (EXTSIMU: siexit_simu dm f ctx ext ext2). { - eapply list_forall2_nth_error; eauto. - - destruct ALLFU as (ITAIL & _). eapply is_tail_nth_error; eauto. - - destruct ALLFU2 as (ITAIL & _). eapply is_tail_nth_error in ITAIL. - assert (LENEQ': length (si_exits st1) = length (si_exits st2)) by (eapply list_forall2_length; eauto). - congruence. } *) -Admitted. -(* - destruct EXTSIMU as (SEMSIMU & _). eapply SEMSIMU in SSEME; eauto. - destruct SSEME as (is2 & SSEME2 & ISIMU). - unfold istate_simu in ISIMU. rewrite ICONT in ISIMU. destruct ISIMU as (path & PATHEQ & SIMULIVE & REVBIND). - destruct SIMULIVE as (CONTEQ & REGLIVE & MEMEQ). - unfold ssem_internal. exists is2. - rewrite <- CONTEQ. rewrite ICONT. constructor; auto. - + eexists; eexists; constructor; eauto. - + unfold istate_simu. rewrite ICONT. eexists. repeat (constructor; eauto). -Qed. -*) - - - Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := match ln with | nil => @@ -1640,25 +1504,6 @@ Proof. + constructor; auto. Admitted. -Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := - do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); - do u2 <- sfval_simub dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2); - OK tt. - -Lemma hsstate_simub_correct dm f hst1 hst2 ctx: - hsstate_simub dm f hst1 hst2 = OK tt -> - hsstate_simu dm f hst1 hst2 ctx. -Proof. - unfold hsstate_simub. intro. explore. - eapply sfval_simub_correct in EQ1. eapply hsistate_simub_correct in EQ. - intros st1 st2 (SREF1 & LREF1 & FREF1) (SREF2 & LREF2 & FREF2). - constructor; [eauto |]. - destruct SREF1 as (PCEQ1 & _ ). destruct SREF2 as (PCEQ2 & _ ). - rewrite <- PCEQ1. rewrite <- PCEQ2. rewrite <- FREF1. rewrite <- FREF2. - eauto. - (* FIXME - that proof will have to be changed once hfinal_refines is more complex *) -Qed. - Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in match (hsexec f pc1) with @@ -1666,7 +1511,7 @@ Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpa | Some hst1 => match (hsexec tf pc2) with | None => Error (msg "simu_check_single: hsexec tf pc2 failed") - | Some hst2 => hsstate_simub dm f hst1 hst2 + | Some hst2 => hsstate_simu_coreb dm f hst1 hst2 end end. @@ -1685,8 +1530,10 @@ Proof. intros (st0 & SEXEC1 & REF1). rewrite SEXEC1 in SEXEC; inversion SEXEC; subst. eexists; split; eauto. - intros; eapply hsstate_simub_correct; eauto. -Qed. + intros ctx. eapply hsstate_simu_coreb_correct in H. + eapply hsstate_simu_core_correct; eauto. + (* TODO - Needs a sok in _theory.v, and a sok -> hok lemma *) +Admitted. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with -- cgit From 712a17a33559baaa1dc6c6cab487393ec2c45407 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 22 Jul 2020 12:49:36 +0200 Subject: smart constructors with hash-consing --- kvx/lib/RTLpathSE_impl_junk.v | 145 +++++++++++++++++++++++++++++++++++++++++- 1 file changed, 144 insertions(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index 8ae21a2d..dd8e80d7 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -19,9 +19,146 @@ Require Export Impure.ImpHCons. Export Notations. Import HConsing. +Local Open Scope impure. + +Import ListNotations. +Local Open Scope list_scope. + +(** * Refinement of symbolic values/symbolic memories with hash-consing data *) + +Inductive hsval := + | HSinput (r: reg) (hid:hashcode) + | HSop (op:operation) (hlsv: hlist_sval) (hsm: hsmem) (hid:hashcode) + | HSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr:addressing) (hlsv:hlist_sval) (hid:hashcode) +with hlist_sval := + | HSnil (hid:hashcode) + | HScons (hsv: hsval) (hlsv: hlist_sval) (hid:hashcode) +(* symbolic memory *) +with hsmem := + | HSinit (hid:hashcode) + | HSstore (hsm: hsmem) (chunk:memory_chunk) (addr:addressing) (hlsv:hlist_sval) (srce: hsval) (hid:hashcode). + +Scheme hsval_mut := Induction for hsval Sort Prop +with hlist_sval_mut := Induction for hlist_sval Sort Prop +with hsmem_mut := Induction for hsmem Sort Prop. + +Definition hsval_get_hid (hsv: hsval): hashcode := + match hsv with + | HSinput _ hid => hid + | HSop _ _ _ hid => hid + | HSload _ _ _ _ _ hid => hid + end. + +Definition hlist_sval_get_hid (hlsv: hlist_sval): hashcode := + match hlsv with + | HSnil hid => hid + | HScons _ _ hid => hid + end. + +Definition hsmem_get_hid (hsm: hsmem ): hashcode := + match hsm with + | HSinit hid => hid + | HSstore _ _ _ _ _ hid => hid + end. + +Definition hsval_set_hid (hsv: hsval) (hid: hashcode): hsval := + match hsv with + | HSinput r _ => HSinput r hid + | HSop o hlsv hsm _ => HSop o hlsv hsm hid + | HSload hsm trap chunk addr hlsv _ => HSload hsm trap chunk addr hlsv hid + end. + +Definition hlist_sval_set_hid (hlsv: hlist_sval) (hid: hashcode): hlist_sval := + match hlsv with + | HSnil _ => HSnil hid + | HScons hsv hlsv _ => HScons hsv hlsv hid + end. + +Definition hsmem_get_sid (hsm: hsmem ) (hid: hashcode): hsmem := + match hsm with + | HSinit _ => HSinit hid + | HSstore hsm chunk addr hlsv srce _ => HSstore hsm chunk addr hlsv srce hid + end. + + +Section CanonBuilding. + +Variable hC_hsval: hashinfo hsval -> ?? hsval. +(*Hypothesis hC_hsval_correct: TODO *) + +Variable hC_hlist_sval: hashinfo hlist_sval -> ?? hlist_sval. +(*Hypothesis hC_hlist_sval_correct: TODO *) + +Variable hC_hsmem: hashinfo hsmem -> ?? hsmem. +(*Hypothesis hC_hsval_correct: TODO *) + +(* First, we wrap constructors for hashed values !*) + +Definition hSinput_hcodes (r: reg) := + DO hc <~ hash 1;; + DO hv <~ hash r;; + RET [hc;hv]. +Extraction Inline hSinput_hcodes. + +Definition hSinput (r:reg): ?? hsval := + DO hv <~ hSinput_hcodes r;; + hC_hsval {| hdata:=HSinput r unknown_hid; hcodes :=hv; |}. + + +Definition hSop_hcodes (op:operation) (hlsv: hlist_sval) (hsm: hsmem) := + DO hc <~ hash 2;; + DO hv <~ hash op;; + RET [hc;hv;hlist_sval_get_hid hlsv; hsmem_get_hid hsm]. +Extraction Inline hSop_hcodes. + +Definition hSop (op:operation) (hlsv: hlist_sval) (hsm: hsmem): ?? hsval := + DO hv <~ hSop_hcodes op hlsv hsm;; + hC_hsval {| hdata:=HSop op hlsv hsm unknown_hid; hcodes :=hv |}. + + +Definition hSload_hcodes (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr:addressing) (hlsv:hlist_sval):= + DO hc <~ hash 3;; + DO hv1 <~ hash trap;; + DO hv2 <~ hash chunk;; + DO hv3 <~ hash addr;; + RET [hc;hsmem_get_hid hsm;hv1;hv2;hv3;hlist_sval_get_hid hlsv]. +Extraction Inline hSload_hcodes. + +Definition hSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr:addressing) (hlsv:hlist_sval): ?? hsval := + DO hv <~ hSload_hcodes hsm trap chunk addr hlsv;; + hC_hsval {| hdata:=HSload hsm trap chunk addr hlsv unknown_hid; hcodes :=hv |}. + + +Definition hSTnil (_: unit): ?? hlist_sval := + hC_hlist_sval {| hdata:=HSnil unknown_hid; hcodes := nil |}. + +Definition hSTcons (hsv: hsval) (hlsv: hlist_sval): ?? hlist_sval := + hC_hlist_sval {| hdata:=HScons hsv hlsv unknown_hid; hcodes := [hsval_get_hid hsv; hlist_sval_get_hid hlsv] |}. + +Definition hSinit (_: unit): ?? hsmem := + hC_hsmem {| hdata:=HSinit unknown_hid; hcodes := nil |}. + +Definition hSstore_hcodes (hsm: hsmem) (chunk: memory_chunk) (addr:addressing) (hlsv:hlist_sval) (srce: hsval):= + DO hv1 <~ hash chunk;; + DO hv2 <~ hash addr;; + RET [hsmem_get_hid hsm;hv1;hv2;hlist_sval_get_hid hlsv;hsval_get_hid srce]. +Extraction Inline hSstore_hcodes. + +Definition hSstore (hsm: hsmem) (chunk:memory_chunk) (addr:addressing) (hlsv:hlist_sval) (srce: hsval): ?? hsmem := + DO hv <~ hSstore_hcodes hsm chunk addr hlsv srce;; + hC_hsmem {| hdata:=HSstore hsm chunk addr hlsv srce unknown_hid; hcodes := hv |}. + + + (* -(** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) + +Fixpoint list_sval_inj (l: list sval): list_sval := + match l with + | nil => Snil + | v::l => Scons v (list_sval_inj l) + end. + (** * Implementation of local symbolic internal states TODO: use refined symbolic values instead of those from RTLpathSE_theory. @@ -211,6 +348,12 @@ Definition hsexec (f: function) (pc:node): option hsstate := | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} end). + +*) + +End CanonBuilding. + +(* (** * The simulation test of concrete symbolic execution *) (* TODO - generalize it with a list of registers to test ? *) -- cgit From 71d5fc0f4ace599d74618880fab0618ce8846e6b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Jul 2020 15:14:53 +0200 Subject: Solving the hok problem --- kvx/lib/RTLpathSE_impl.v | 45 ++++++++++++++------------------------------- 1 file changed, 14 insertions(+), 31 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index f25143fd..e55cd461 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -194,9 +194,8 @@ Definition hsok_exit ge sp rs m hse := hsok_local ge sp rs m (hsi_elocal hse). Definition hsiexit_refines_dyn ge sp rs0 m0 (hext: hsistate_exit) (ext: sistate_exit): Prop := hsilocal_refines ge sp rs0 m0 (hsi_elocal hext) (si_elocal ext) - /\ (hsok_exit ge sp rs0 m0 hext -> - seval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem_get (hsi_elocal hext)) rs0 m0 - = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0). + /\ seval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem_get (hsi_elocal hext)) rs0 m0 + = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0. Definition hsiexit_simu dm f (ctx: simu_proof_context f) hse1 hse2: Prop := forall se1 se2, hsiexit_refines_stat hse1 se1 -> @@ -226,11 +225,9 @@ Qed. Theorem hsiexit_simu_core_correct dm f hse1 hse2 ctx: hsiexit_simu_core dm f hse1 hse2 -> - hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1 -> hsiexit_simu dm f ctx hse1 hse2. Proof. - intros SIMUC HOK1 st1 st2 HREF1 HREF2 HDYN1 HDYN2. - exploit hsiexit_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. + intros SIMUC (* HOK1 *) st1 st2 HREF1 HREF2 HDYN1 HDYN2. assert (SEVALC: seval_condition (the_ge1 ctx) (the_sp ctx) (si_cond st1) (si_scondargs st1) (si_smem (si_elocal st1)) (the_rs0 ctx) (the_m0 ctx) = @@ -241,8 +238,11 @@ Proof. destruct SIMUC as (_ & _ & CONDEQ & ARGSEQ & LSIMU). destruct LSIMU as (_ & _ & _ & MEMEQ). rewrite CONDEQ. rewrite ARGSEQ. rewrite MEMEQ. erewrite <- seval_condition_preserved; eauto. eapply ctx. } - constructor; [assumption|]. - intros is1 SSEME. + constructor; [assumption|]. intros is1 SSEME. + assert (HOK1: hsok_exit (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse1). { + unfold hsok_exit. destruct SSEME as (_ & SSEML & _). apply ssem_local_sok in SSEML. + destruct HDYN1 as (LREF & _). destruct LREF as (OKEQ & _ & _). rewrite <- OKEQ. assumption. } + exploit hsiexit_simu_core_nofail. 2: eapply ctx. all: eauto. intro HOK2. exists (mk_istate (icontinue is1) (si_ifso st2) (irs is1) (imem is1)). simpl. constructor. - constructor; [|constructor]. + rewrite <- SEVALC. destruct SSEME as (SCOND & _ & _). assumption. @@ -273,13 +273,6 @@ Proof. auto. Qed. -Inductive hsok_exits ge sp rs m: list hsistate_exit -> Prop := - | hsok_exits_nil: hsok_exits ge sp rs m nil - | hsok_exits_cons: forall he lhe, - hsok_exits ge sp rs m lhe -> - hsok_exit ge sp rs m he -> - hsok_exits ge sp rs m (he::lhe). - Definition hsiexits_simu dm f (ctx: simu_proof_context f) lhse1 lhse2: Prop := list_forall2 (hsiexit_simu dm f ctx) lhse1 lhse2. @@ -329,11 +322,10 @@ Admitted. Theorem hsiexits_simu_core_correct dm f lhse1 lhse2 ctx: hsiexits_simu_core dm f lhse1 lhse2 -> - hsok_exits (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) lhse1 -> hsiexits_simu dm f ctx lhse1 lhse2. Proof. induction 1; [constructor|]. - intros HOKS. inv HOKS. constructor; [|apply IHlist_forall2; assumption]. + constructor; [|apply IHlist_forall2; assumption]. apply hsiexit_simu_core_correct; assumption. Qed. @@ -374,10 +366,6 @@ Definition hsistate_simu dm f (hst1 hst2: hsistate) (ctx: simu_proof_context f): hsistate_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst2 st2 -> sistate_simu dm f st1 st2 ctx. -Definition hsok ge sp rs m hst := - hsok_exits ge sp rs m (hsi_exits hst) - /\ hsok_local ge sp rs m (hsi_local hst). - Lemma siexits_simu_all_fallthrough dm f ctx: forall lse1 lse2, siexits_simu dm f lse1 lse2 ctx -> all_fallthrough (the_ge1 ctx) (the_sp ctx) lse1 (the_rs0 ctx) (the_m0 ctx) -> @@ -471,13 +459,12 @@ Qed. Theorem hsistate_simu_core_correct dm f hst1 hst2 ctx: hsistate_simu_core dm f hst1 hst2 -> - hsok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> hsistate_simu dm f hst1 hst2 ctx. Proof. - intros SIMUC HOK st1 st2 HREF1 HREF2 DREF1 DREF2 is1 SEMI. + intros SIMUC st1 st2 HREF1 HREF2 DREF1 DREF2 is1 SEMI. destruct HREF1 as (PCREF1 & EREF1). destruct HREF2 as (PCREF2 & EREF2). destruct DREF1 as (DEREF1 & LREF1). destruct DREF2 as (DEREF2 & LREF2). - destruct SIMUC as (PCSIMU & ESIMU & LSIMU). destruct HOK as (EOK1 & LOK1). + destruct SIMUC as (PCSIMU & ESIMU & LSIMU). exploit hsiexits_simu_core_correct; eauto. intro HESIMU. unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT. - destruct SEMI as (SSEML & PCEQ & ALLFU). @@ -604,8 +591,6 @@ Theorem hsstate_simu_coreb_correct dm f hst1 hst2: Proof. Admitted. -Definition hok ge sp rs m hst := hsok ge sp rs m (hinternal hst). - Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := forall st1 st2, hsstate_refines hst1 st1 -> @@ -613,13 +598,12 @@ Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := Theorem hsstate_simu_core_correct dm f ctx hst1 hst2: hsstate_simu_core dm f hst1 hst2 -> - hok (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hst1 -> hsstate_simu dm f hst1 hst2 ctx. Proof. - intros (SCORE & FSIMU) HOK1. intros st1 st2 HREF1 HREF2. + intros (SCORE & FSIMU). intros st1 st2 HREF1 HREF2. destruct HREF1 as (SREF1 & DREF1 & FREF1). destruct HREF2 as (SREF2 & DREF2 & FREF2). assert (PCEQ: dm ! (hsi_pc hst2) = Some (hsi_pc hst1)) by apply SCORE. - eapply hsistate_simu_core_correct in SCORE; [|eassumption]. + eapply hsistate_simu_core_correct in SCORE. eapply hfinal_simu_core_correct in FSIMU; eauto. constructor; [apply SCORE; auto|]. destruct SREF1 as (PC1 & _). destruct SREF2 as (PC2 & _). rewrite <- PC1. rewrite <- PC2. @@ -1532,8 +1516,7 @@ Proof. eexists; split; eauto. intros ctx. eapply hsstate_simu_coreb_correct in H. eapply hsstate_simu_core_correct; eauto. - (* TODO - Needs a sok in _theory.v, and a sok -> hok lemma *) -Admitted. +Qed. Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := match lm with -- cgit From efca4a363642abc530756242eab7850ddabf8c33 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 22 Jul 2020 16:33:36 +0200 Subject: Verificators of inclusion --- kvx/lib/RTLpathSE_impl.v | 492 +++++++++++++++++++++++++++-------------------- 1 file changed, 286 insertions(+), 206 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index e55cd461..8adc92f9 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -14,7 +14,7 @@ Local Open Scope option_monad_scope. (** * TODO: refine symbolic values/symbolic memories with hash-consed symbolic values *) -(** * Implementation of local symbolic internal states *) +(** * Implementation of local symbolic internal states - definitions and core simulation properties *) (** name : Hash-consed Symbolic Internal state local. Later on we will use the refinement to introduce hash consing *) @@ -82,15 +82,6 @@ Proof. intuition congruence. Qed. -(* -Lemma sok_local_no_abort ge sp rs0 m0 st: - sok_local ge sp rs0 m0 st -> sabort_local ge sp st rs0 m0 -> False. -Proof. - unfold sok_local, sabort_local. - intros (H1 & H2 & H3) [A1 | [A2 | (r & A3)]]; intuition eauto. -Qed. -*) - Definition hsok_local ge sp rs0 m0 (hst: hsistate_local) : Prop := (forall sv, List.In sv (hsi_ok_lsval hst) -> seval_sval ge sp sv rs0 m0 <> None) /\ (forall sm, List.In sm (hsi_lsmem hst) -> seval_smem ge sp sm rs0 m0 <> None). @@ -107,20 +98,14 @@ Proof. intros H0 (H1 & _ & _). apply H1. eapply ssem_local_sok. eauto. Qed. +Definition is_subset {A: Type} (lv2 lv1: list A) := forall v, In v lv2 -> In v lv1. + Definition hsilocal_simu_core (hst1 hst2: hsistate_local) := - (forall sm, List.In sm (hsi_lsmem hst2) -> List.In sm (hsi_lsmem hst1)) - /\ (forall sv, List.In sv (hsi_ok_lsval hst2) -> List.In sv (hsi_ok_lsval hst1)) + is_subset (hsi_lsmem hst2) (hsi_lsmem hst1) + /\ is_subset (hsi_ok_lsval hst2) (hsi_ok_lsval hst1) /\ (forall r, hsi_sreg_get hst2 r = hsi_sreg_get hst1 r) /\ hsi_smem_get hst1 = hsi_smem_get hst2. -Definition hsilocal_simu_coreb (hst1 hst2: hsistate_local) := OK tt. - -Theorem hsilocal_simu_coreb_correct hst1 hst2: - hsilocal_simu_coreb hst1 hst2 = OK tt -> - hsilocal_simu_core hst1 hst2. -Proof. -Admitted. - Lemma hsilocal_simu_core_nofail ge1 ge2 sp rs0 m0 hst1 hst2: hsilocal_simu_core hst1 hst2 -> (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> @@ -204,14 +189,6 @@ Definition hsiexit_simu dm f (ctx: simu_proof_context f) hse1 hse2: Prop := fora hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 -> siexit_simu dm f ctx se1 se2. -(* Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := - if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then - do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); - do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); - revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) - else Error (msg "siexit_simub: conditions do not match") -. *) - Lemma hsiexit_simu_core_nofail dm f hse1 hse2 ge1 ge2 sp rs m: hsiexit_simu_core dm f hse1 hse2 -> (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) -> @@ -279,47 +256,6 @@ Definition hsiexits_simu dm f (ctx: simu_proof_context f) lhse1 lhse2: Prop := Definition hsiexits_simu_core dm f lhse1 lhse2: Prop := list_forall2 (hsiexit_simu_core dm f) lhse1 lhse2. -(* Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := - match lhse1 with - | nil => - match lhse2 with - | nil => OK tt - | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") - end - | hse1 :: lhse1 => - match lhse2 with - | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") - | hse2 :: lhse2 => - do _ <- hsiexit_simub dm f hse1 hse2; - do _ <- hsiexits_simub dm f lhse1 lhse2; - OK tt - end - end. *) - -(* Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, - hsiexits_simub dm f lhse1 lhse2 = OK tt -> - hsiexits_simu dm f ctx lhse1 lhse2. -Proof. -(* induction lhse1. - - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. - inv HEREFS1. inv HEREFS2. constructor. - - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. - fold hsiexits_simub in EQ1. - eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. - intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. - apply EQ1; assumption. *) -Admitted. - *) - -(* TODO *) -Definition hsiexits_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := OK tt. - -Theorem hsiexits_simu_coreb_correct dm f lhse1 lhse2: - hsiexits_simu_coreb dm f lhse1 lhse2 = OK tt -> - hsiexits_simu_core dm f lhse1 lhse2. -Proof. -Admitted. - Theorem hsiexits_simu_core_correct dm f lhse1 lhse2 ctx: hsiexits_simu_core dm f lhse1 lhse2 -> hsiexits_simu dm f ctx lhse1 lhse2. @@ -343,14 +279,6 @@ Definition hsistate_simu_core dm f (hse1 hse2: hsistate) := /\ list_forall2 (hsiexit_simu_core dm f) (hsi_exits hse1) (hsi_exits hse2) /\ hsilocal_simu_core (hsi_local hse1) (hsi_local hse2). -Definition hsistate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate) := OK tt. (* TODO *) - -Theorem hsistate_simu_coreb_correct dm f hse1 hse2: - hsistate_simu_coreb dm f hse1 hse2 = OK tt -> - hsistate_simu_core dm f hse1 hse2. -Proof. -Admitted. - Definition hsistate_refines_stat (hst: hsistate) (st:sistate): Prop := hsi_pc hst = si_pc st /\ hsiexits_refines_stat (hsi_exits hst) (si_exits st). @@ -527,14 +455,6 @@ Definition hfinal_simu_core (dm: PTree.t node) (f: RTLpath.function) (hf1 hf2: s | _ => hf1 = hf2 end. -Definition hfinal_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hf1 hf2: sfval) := OK tt. (* TODO *) - -Theorem hfinal_simu_coreb_correct dm f hf1 hf2: - hfinal_simu_coreb dm f hf1 hf2 = OK tt -> - hfinal_simu_core dm f hf1 hf2. -Proof. -Admitted. - Lemma svident_simu_refl f ctx s: svident_simu f ctx s s. Proof. @@ -583,14 +503,6 @@ Definition hsstate_simu_core (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2 hsistate_simu_core dm f (hinternal hst1) (hinternal hst2) /\ hfinal_simu_core dm f (hfinal hst1) (hfinal hst2). -Definition hsstate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) - -Theorem hsstate_simu_coreb_correct dm f hst1 hst2: - hsstate_simu_coreb dm f hst1 hst2 = OK tt -> - hsstate_simu_core dm f hst1 hst2. -Proof. -Admitted. - Definition hsstate_simu dm f (hst1 hst2: hsstate) ctx: Prop := forall st1 st2, hsstate_refines hst1 st1 -> @@ -610,6 +522,285 @@ Proof. eapply FSIMU. Qed. +(** * Verificators of *_simu_core properties *) + +(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) +Fixpoint sval_simub (sv1 sv2: sval) := + match sv1 with + | Sinput r => + match sv2 with + | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") + | _ => Error (msg "sval_simub: Sinput expected") + end + | Sop op lsv sm => + match sv2 with + | Sop op' lsv' sm' => + if (eq_operation op op') then + do _ <- list_sval_simub lsv lsv'; + smem_simub sm sm' + else Error (msg "sval_simub: different operations in Sop") + | _ => Error (msg "sval_simub: Sop expected") + end + | Sload sm trap chk addr lsv => + match sv2 with + | Sload sm' trap' chk' addr' lsv' => + if (trapping_mode_eq trap trap') then + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + list_sval_simub lsv lsv' + else Error (msg "sval_simub: addressings do not match") + else Error (msg "sval_simub: chunks do not match") + else Error (msg "sval_simub: trapping modes do not match") + (* FIXME - should be refined once we introduce non trapping loads *) + | _ => Error (msg "sval_simub: Sload expected") + end + end +with list_sval_simub (lsv1 lsv2: list_sval) := + match lsv1 with + | Snil => + match lsv2 with + | Snil => OK tt + | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") + end + | Scons sv1 lsv1 => + match lsv2 with + | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") + | Scons sv2 lsv2 => + do _ <- sval_simub sv1 sv2; + list_sval_simub lsv1 lsv2 + end + end +with smem_simub (sm1 sm2: smem) := + match sm1 with + | Sinit => + match sm2 with + | Sinit => OK tt + | _ => Error (msg "smem_simub: Sinit expected") + end + | Sstore sm chk addr lsv sv => + match sm2 with + | Sstore sm' chk' addr' lsv' sv' => + if (chunk_eq chk chk') then + if (eq_addressing addr addr') then + do _ <- smem_simub sm sm'; + do _ <- list_sval_simub lsv lsv'; + sval_simub sv sv' + else Error (msg "smem_simub: addressings do not match") + else Error (msg "smem_simub: chunks not match") + | _ => Error (msg "smem_simub: Sstore expected") + end + end. + +Lemma sval_simub_correct sv1: forall sv2, + sval_simub sv1 sv2 = OK tt -> sv1 = sv2. +Proof. + induction sv1 using sval_mut with + (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) + (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. + (* Sinput *) + - destruct sv2; try discriminate. + destruct (Pos.eq_dec r r0); [congruence|discriminate]. + (* Sop *) + - destruct sv2; try discriminate. + destruct (eq_operation _ _); [|discriminate]. subst. + intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sload *) + - destruct sv2; try discriminate. + destruct (trapping_mode_eq _ _); [|discriminate]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. + congruence. + (* Snil *) + - destruct lsv2; [|discriminate]. congruence. + (* Scons *) + - destruct lsv2; [discriminate|]. intro. explore. + apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. + (* Sinit *) + - destruct sm2; [|discriminate]. congruence. + (* Sstore *) + - destruct sm2; [discriminate|]. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. + intro. explore. + assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. + congruence. +Qed. + +Lemma list_sval_simub_correct lsv1: forall lsv2, + list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. +Proof. + induction lsv1; simpl; auto. + - destruct lsv2; try discriminate. reflexivity. + - destruct lsv2; try discriminate. intro. explore. + apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. + congruence. +Qed. + +Lemma smem_simub_correct sm1: forall sm2, + smem_simub sm1 sm2 = OK tt -> sm1 = sm2. +Proof. + induction sm1; simpl; auto. + - destruct sm2; try discriminate. reflexivity. + - destruct sm2; try discriminate. + destruct (chunk_eq _ _); [|discriminate]. + destruct (eq_addressing _ _); [|discriminate]. intro. explore. + apply sval_simub_correct in EQ2. apply list_sval_simub_correct in EQ1. + apply IHsm1 in EQ. congruence. +Qed. + +Definition is_structural {A: Type} (cmp: A -> A -> bool) := + forall x y, cmp x y = true -> x = y. + +Fixpoint is_part_of {A: Type} (cmp: A -> A -> bool) (elt: A) (lv: list A): bool := + match lv with + | nil => false + | v::lv => if (cmp elt v) then true else is_part_of cmp elt lv + end. + +Lemma is_part_of_correct {A: Type} cmp lv (e: A): + is_structural cmp -> + is_part_of cmp e lv = true -> + In e lv. +Proof. + induction lv. + - intros. simpl in H0. discriminate. + - intros. simpl in H0. destruct (cmp e a) eqn:CMP. + + apply H in CMP. subst. constructor; auto. + + right. apply IHlv; assumption. +Qed. + +(* Checks if lv2 is a subset of lv1 *) +Fixpoint is_subsetb {A: Type} (cmp: A -> A -> bool) (lv2 lv1: list A): bool := + match lv2 with + | nil => true + | v2 :: lv2 => if (is_part_of cmp v2 lv1) then is_subsetb cmp lv2 lv1 + else false + end. + +Lemma is_subset_cons {A: Type} (x: A) lv lx: + In x lv /\ is_subset lx lv -> is_subset (x::lx) lv. +Proof. + intros (ISIN & ISSUB). unfold is_subset. + intros. inv H. + - assumption. + - apply ISSUB. assumption. +Qed. + +Lemma is_subset_correct {A: Type} cmp (lv2 lv1: list A): + is_structural cmp -> + is_subsetb cmp lv2 lv1 = true -> + is_subset lv2 lv1. +Proof. + induction lv2. + - simpl. intros. intro. intro. apply in_nil in H1. contradiction. + - intros. simpl in H0. apply is_subset_cons. + explore. apply is_part_of_correct in EQ; [|assumption]. + apply IHlv2 in H0; [|assumption]. constructor; assumption. +Qed. + +Definition simub_bool {A: Type} (simub: A -> A -> res unit) (sv1 sv2: A) := + match simub sv1 sv2 with + | OK tt => true + | _ => false + end. + +Lemma simub_bool_correct {A: Type} simub (sv1 sv2: A): + (forall x y, simub x y = OK tt -> x = y) -> + simub_bool simub sv1 sv2 = true -> sv1 = sv2. +Proof. + intros. unfold simub_bool in H0. destruct (simub sv1 sv2) eqn:SIMU; explore. + - apply H. assumption. + - discriminate. +Qed. + +Definition hsilocal_simu_coreb hst1 hst2 := + if (is_subsetb (simub_bool smem_simub) (hsi_lsmem hst2) (hsi_lsmem hst1)) then + if (is_subsetb (simub_bool sval_simub) (hsi_ok_lsval hst2) (hsi_ok_lsval hst1)) then + (* TODO - compare on the whole ptree *) OK tt + else Error (msg "hsi_ok_lsval sets aren't included") + else Error (msg "hsi_lsmem sets aren't included"). + +Theorem hsilocal_simu_coreb_correct hst1 hst2: + hsilocal_simu_coreb hst1 hst2 = OK tt -> + hsilocal_simu_core hst1 hst2. +Proof. +Admitted. + +(* Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := + if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then + do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); + do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); + revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) + else Error (msg "siexit_simub: conditions do not match") +. *) + +(* Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := + match lhse1 with + | nil => + match lhse2 with + | nil => OK tt + | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") + end + | hse1 :: lhse1 => + match lhse2 with + | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") + | hse2 :: lhse2 => + do _ <- hsiexit_simub dm f hse1 hse2; + do _ <- hsiexits_simub dm f lhse1 lhse2; + OK tt + end + end. *) + +(* Lemma hsiexits_simub_correct dm f ctx lhse1: forall lhse2, + hsiexits_simub dm f lhse1 lhse2 = OK tt -> + hsiexits_simu dm f ctx lhse1 lhse2. +Proof. +(* induction lhse1. + - simpl. intros. destruct lhse2; try discriminate. intros se1 se2 HEREFS1 HEREFS2 _ _. + inv HEREFS1. inv HEREFS2. constructor. + - (* simpl. *) unfold hsiexits_simub. intros. destruct lhse2; try discriminate. explore. + fold hsiexits_simub in EQ1. + eapply hsiexit_simub_correct in EQ. apply IHlhse1 in EQ1. + intros se1 se2 HEREFS1 HEREFS2 HEREFD1 HEREFD2. inv HEREFS1. inv HEREFS2. inv HEREFD1. inv HEREFD2. constructor; auto. + apply EQ1; assumption. *) +Admitted. + *) + +(* TODO *) +Definition hsiexits_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := OK tt. + +Theorem hsiexits_simu_coreb_correct dm f lhse1 lhse2: + hsiexits_simu_coreb dm f lhse1 lhse2 = OK tt -> + hsiexits_simu_core dm f lhse1 lhse2. +Proof. +Admitted. + +Definition hsistate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate) := OK tt. (* TODO *) + +Theorem hsistate_simu_coreb_correct dm f hse1 hse2: + hsistate_simu_coreb dm f hse1 hse2 = OK tt -> + hsistate_simu_core dm f hse1 hse2. +Proof. +Admitted. + +Definition hsstate_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := OK tt. (* TODO *) + +Theorem hsstate_simu_coreb_correct dm f hst1 hst2: + hsstate_simu_coreb dm f hst1 hst2 = OK tt -> + hsstate_simu_core dm f hst1 hst2. +Proof. +Admitted. + +Definition hfinal_simu_coreb (dm: PTree.t node) (f: RTLpath.function) (hf1 hf2: sfval) := OK tt. (* TODO *) + +Theorem hfinal_simu_coreb_correct dm f hf1 hf2: + hfinal_simu_coreb dm f hf1 hf2 = OK tt -> + hfinal_simu_core dm f hf1 hf2. +Proof. +Admitted. + Lemma hsistate_refines_stat_pceq st hst: hsistate_refines_stat hst st -> (hsi_pc hst) = (si_pc st). @@ -624,6 +815,8 @@ Proof. unfold hsistate_refines_dyn; intuition. Qed. + + Local Hint Resolve hsistate_refines_dyn_local_refines: core. @@ -1114,119 +1307,6 @@ Proof. unfold revmap_check_single. explore; try discriminate. congruence. Qed. -(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) -Fixpoint sval_simub (sv1 sv2: sval) := - match sv1 with - | Sinput r => - match sv2 with - | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") - | _ => Error (msg "sval_simub: Sinput expected") - end - | Sop op lsv sm => - match sv2 with - | Sop op' lsv' sm' => - if (eq_operation op op') then - do _ <- list_sval_simub lsv lsv'; - smem_simub sm sm' - else Error (msg "sval_simub: different operations in Sop") - | _ => Error (msg "sval_simub: Sop expected") - end - | Sload sm trap chk addr lsv => - match sv2 with - | Sload sm' trap' chk' addr' lsv' => - if (trapping_mode_eq trap trap') then - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - list_sval_simub lsv lsv' - else Error (msg "sval_simub: addressings do not match") - else Error (msg "sval_simub: chunks do not match") - else Error (msg "sval_simub: trapping modes do not match") - (* FIXME - should be refined once we introduce non trapping loads *) - | _ => Error (msg "sval_simub: Sload expected") - end - end -with list_sval_simub (lsv1 lsv2: list_sval) := - match lsv1 with - | Snil => - match lsv2 with - | Snil => OK tt - | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") - end - | Scons sv1 lsv1 => - match lsv2 with - | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") - | Scons sv2 lsv2 => - do _ <- sval_simub sv1 sv2; - list_sval_simub lsv1 lsv2 - end - end -with smem_simub (sm1 sm2: smem) := - match sm1 with - | Sinit => - match sm2 with - | Sinit => OK tt - | _ => Error (msg "smem_simub: Sinit expected") - end - | Sstore sm chk addr lsv sv => - match sm2 with - | Sstore sm' chk' addr' lsv' sv' => - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - do _ <- list_sval_simub lsv lsv'; - sval_simub sv sv' - else Error (msg "smem_simub: addressings do not match") - else Error (msg "smem_simub: chunks not match") - | _ => Error (msg "smem_simub: Sstore expected") - end - end. - -Lemma sval_simub_correct sv1: forall sv2, - sval_simub sv1 sv2 = OK tt -> sv1 = sv2. -Proof. - induction sv1 using sval_mut with - (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) - (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. - (* Sinput *) - - destruct sv2; try discriminate. - destruct (Pos.eq_dec r r0); [congruence|discriminate]. - (* Sop *) - - destruct sv2; try discriminate. - destruct (eq_operation _ _); [|discriminate]. subst. - intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sload *) - - destruct sv2; try discriminate. - destruct (trapping_mode_eq _ _); [|discriminate]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. - congruence. - (* Snil *) - - destruct lsv2; [|discriminate]. congruence. - (* Scons *) - - destruct lsv2; [discriminate|]. intro. explore. - apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sinit *) - - destruct sm2; [|discriminate]. congruence. - (* Sstore *) - - destruct sm2; [discriminate|]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. - assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. - congruence. -Qed. - -Lemma list_sval_simub_correct lsv1: forall lsv2, - list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. -Proof. - induction lsv1; simpl; auto. - - destruct lsv2; try discriminate. reflexivity. - - destruct lsv2; try discriminate. intro. explore. - apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. - congruence. -Qed. Local Hint Resolve genv_match ssem_local_refines_hok: core. -- cgit From 01cb61ae13cefdcaddf0a7f53db5c3b5affa2992 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 22 Jul 2020 19:05:34 +0200 Subject: symbolic execution of internal nodes with hash-consing --- kvx/lib/RTLpathSE_impl_junk.v | 151 ++++++++++++++++++++---------------------- 1 file changed, 70 insertions(+), 81 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index dd8e80d7..264706d5 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -129,10 +129,10 @@ Definition hSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr hC_hsval {| hdata:=HSload hsm trap chunk addr hlsv unknown_hid; hcodes :=hv |}. -Definition hSTnil (_: unit): ?? hlist_sval := +Definition hSnil (_: unit): ?? hlist_sval := hC_hlist_sval {| hdata:=HSnil unknown_hid; hcodes := nil |}. -Definition hSTcons (hsv: hsval) (hlsv: hlist_sval): ?? hlist_sval := +Definition hScons (hsv: hsval) (hlsv: hlist_sval): ?? hlist_sval := hC_hlist_sval {| hdata:=HScons hsv hlsv unknown_hid; hcodes := [hsval_get_hid hsv; hlist_sval_get_hid hlsv] |}. Definition hSinit (_: unit): ?? hsmem := @@ -149,63 +149,57 @@ Definition hSstore (hsm: hsmem) (chunk:memory_chunk) (addr:addressing) (hlsv:hli hC_hsmem {| hdata:=HSstore hsm chunk addr hlsv srce unknown_hid; hcodes := hv |}. - - -(* - -Fixpoint list_sval_inj (l: list sval): list_sval := - match l with - | nil => Snil - | v::l => Scons v (list_sval_inj l) - end. - - (** * Implementation of local symbolic internal states -TODO: use refined symbolic values instead of those from RTLpathSE_theory. *) (** name : Hash-consed Symbolic Internal state local. Later on we will use the refinement to introduce hash consing *) Record hsistate_local := { - (** [hsi_lsmem] represents the list of smem symbolic evaluations. - The first one of the list is the current smem *) - hsi_lsmem:> list smem; + (** [hsi_smem] represents the current smem symbolic evaluations. + (we can recover the previous one from smem) *) + hsi_smem:> hsmem; (** For the values in registers: 1) we store a list of sval evaluations 2) we encode the symbolic regset by a PTree *) - hsi_ok_lsval: list sval; - hsi_sreg:> PTree.t sval + hsi_ok_lsval: list hsval; + hsi_sreg:> PTree.t hsval }. -Definition hsi_sreg_get (hst: PTree.t sval) r: sval := +Definition hsi_sreg_get (hst: PTree.t hsval) r: ?? hsval := match PTree.get r hst with - | None => Sinput r - | Some sv => sv + | None => hSinput r + | Some sv => RET sv end. +Fixpoint hlist_args (hst: PTree.t hsval) (l: list reg): ?? hlist_sval := + match l with + | nil => hSnil() + | r::l => + DO v <~ hsi_sreg_get hst r;; + DO hlsv <~ hlist_args hst l;; + hScons v hlsv + end. -Definition hsi_smem_get (hst: list smem): smem := - match hst with - | nil => Sinit - | sm::_ => sm - end. (* Syntax and semantics of symbolic exit states *) -(* TODO: add hash-consing *) Record hsistate_exit := mk_hsistate_exit - { hsi_cond: condition; hsi_scondargs: list_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. + { hsi_cond: condition; hsi_scondargs: hlist_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. (** * Syntax and Semantics of symbolic internal state *) Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. (** ** Assignment of memory *) -Definition hslocal_set_smem (hst:hsistate_local) (sm:smem) := - {| hsi_lsmem := sm::hsi_lsmem hst; - hsi_ok_lsval := hsi_ok_lsval hst; - hsi_sreg:= hsi_sreg hst - |}. +Definition hslocal_store (hst:hsistate_local) chunk addr args src: ?? hsistate_local := + let pt := hst.(hsi_sreg) in + DO hargs <~ hlist_args pt args;; + DO hsrc <~ hsi_sreg_get pt src;; + DO hm <~ hSstore hst chunk addr hargs hsrc;; + RET {| hsi_smem := hm; + hsi_ok_lsval := hsi_ok_lsval hst; + hsi_sreg:= hsi_sreg hst + |}. (** ** Assignment of local state *) @@ -220,17 +214,17 @@ Inductive root_sval: Type := | Rload (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) . -Definition root_apply (rsv: root_sval) (lsv: list sval) (sm: smem): sval := +Definition root_apply (rsv: root_sval) (lsv: list reg) (hst: hsistate_local) : ?? hsval := + DO hlsv <~ hlist_args hst lsv;; match rsv with - | Rop op => Sop op (list_sval_inj lsv) sm - | Rload trap chunk addr => Sload sm trap chunk addr (list_sval_inj lsv) + | Rop op => hSop op hlsv hst + | Rload trap chunk addr => hSload hst trap chunk addr hlsv end. -Coercion root_apply: root_sval >-> Funclass. Local Open Scope lazy_bool_scope. (* NB: return [false] if the rsv cannot fail *) -Definition may_trap (rsv: root_sval) (lsv: list sval) (sm: smem): bool := +Definition may_trap (rsv: root_sval) (lsv: list reg): bool := match rsv with | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lsv) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *) | Rload TRAP _ _ => true @@ -238,79 +232,74 @@ Definition may_trap (rsv: root_sval) (lsv: list sval) (sm: smem): bool := end. (* simplify a symbolic value before assignment to a register *) -Definition simplify (rsv: root_sval) lsv sm: sval := +Definition simplify (rsv: root_sval) (lsv: list reg) (hst: hsistate_local): ?? hsval := match rsv with - | Rload TRAP chunk addr => Sload sm NOTRAP chunk addr (list_sval_inj lsv) | Rop op => match is_move_operation op lsv with - | Some arg => arg (* optimization of Omove *) - | None => - if op_depends_on_memory op then - rsv lsv sm - else - Sop op (list_sval_inj lsv) Sinit (* magically remove the dependency on sm ! *) + | Some arg => hsi_sreg_get hst arg (* optimization of Omove *) + | None => + DO hsi <~ hSinit ();; + DO hlsv <~ hlist_args hst lsv;; + hSop op hlsv hsi (* magically remove the dependency on sm ! *) end - | _ => rsv lsv sm + | Rload _ chunk addr => + DO hlsv <~ hlist_args hst lsv;; + hSload hst NOTRAP chunk addr hlsv end. -Definition red_PTree_set (r:reg) (sv: sval) (hst: PTree.t sval): PTree.t sval := +Definition red_PTree_set (r:reg) (sv: hsval) (hst: PTree.t hsval): PTree.t hsval := match sv with - | Sinput r' => + | HSinput r' _ => if Pos.eq_dec r r' then PTree.remove r' hst else PTree.set r sv hst | _ => PTree.set r sv hst end. -Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv sm := - {| hsi_lsmem := hsi_lsmem hst; - hsi_ok_lsval := if may_trap rsv lsv sm then (rsv lsv sm)::(hsi_ok_lsval hst) else hsi_ok_lsval hst; - hsi_sreg := red_PTree_set r (simplify rsv lsv sm) (hsi_sreg hst) |}. +Definition hslocal_set_sreg (hst:hsistate_local) (r:reg) (rsv:root_sval) lsv: ?? hsistate_local := + DO hsiok <~ + (if may_trap rsv lsv + then DO hv <~ root_apply rsv lsv hst;; RET (hv::(hsi_ok_lsval hst)) + else RET (hsi_ok_lsval hst));; + DO simp <~ simplify rsv lsv hst;; + RET {| hsi_smem := hst; + hsi_ok_lsval := hsiok; + hsi_sreg := red_PTree_set r simp (hsi_sreg hst) |}. (** ** Execution of one instruction *) -Definition hsiexec_inst (i: instruction) (hst: hsistate): option hsistate := +Definition hsiexec_inst (i: instruction) (hst: hsistate): ?? (option hsistate) := match i with | Inop pc' => - Some (hsist_set_local hst pc' hst.(hsi_local)) + RET (Some (hsist_set_local hst pc' hst.(hsi_local))) | Iop op args dst pc' => - let prev := hst.(hsi_local) in - let vargs := List.map (hsi_sreg_get prev) args in - let next := hslocal_set_sreg prev dst (Rop op) vargs (hsi_smem_get prev) in - Some (hsist_set_local hst pc' next) + DO next <~ hslocal_set_sreg hst.(hsi_local) dst (Rop op) args;; + RET (Some (hsist_set_local hst pc' next)) | Iload trap chunk addr args dst pc' => - let prev := hst.(hsi_local) in - let vargs := List.map (hsi_sreg_get prev) args in - let next := hslocal_set_sreg prev dst (Rload trap chunk addr) vargs (hsi_smem_get prev) in - Some (hsist_set_local hst pc' next) + DO next <~ hslocal_set_sreg hst.(hsi_local) dst (Rload trap chunk addr) args;; + RET (Some (hsist_set_local hst pc' next)) | Istore chunk addr args src pc' => - let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - let next := hslocal_set_smem prev (Sstore (hsi_smem_get prev) chunk addr vargs (hsi_sreg_get prev src)) in - Some (hsist_set_local hst pc' next) + DO next <~ hslocal_store hst.(hsi_local) chunk addr args src;; + RET (Some (hsist_set_local hst pc' next)) | Icond cond args ifso ifnot _ => let prev := hst.(hsi_local) in - let vargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in + DO vargs <~ hlist_args prev args ;; let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in - Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |} - | _ => None (* TODO jumptable ? *) + RET (Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |}) + | _ => RET None (* TODO jumptable ? *) end. -Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): option hsistate := +Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): ?? hsistate := match path with - | O => Some hst + | O => RET hst | S p => - SOME i <- (fn_code f)!(hst.(hsi_pc)) IN - SOME hst1 <- hsiexec_inst i hst IN + DO i <~ match (fn_code f)!(hst.(hsi_pc)) with Some i => RET i | _ => FAILWITH "hsiexec_path: invalid node" end;; + DO ohst1 <~ hsiexec_inst i hst;; + DO hst1 <~ match ohst1 with Some x => RET x | _ => FAILWITH "hsiexec_path: invalid path" end;; hsiexec_path p f hst1 end. -(* Definition hfinal_refines hfv fv := forall pge ge sp npc stk f rs0 m0 rs' m' t s', - ssem_final pge ge sp npc stk f rs0 m0 hfv rs' m' t s' <-> ssem_final pge ge sp npc stk f rs0 m0 fv rs' m' t s'. *) - -(* FIXME - too strong let's change it later.. *) -Definition hfinal_refines (hfv fv: sfval) := hfv = fv. - +(* Record hsstate := { hinternal:> hsistate; hfinal: sfval }. Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := -- cgit From 964a04ac1d646afd50cbabc2de96480e45d6f43b Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 23 Jul 2020 18:18:17 +0200 Subject: full symbolic execution with hash-consing --- kvx/lib/RTLpathSE_impl_junk.v | 130 ++++++++++++++++++++++++++++++------------ 1 file changed, 95 insertions(+), 35 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index 264706d5..708c98f5 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -80,6 +80,17 @@ Definition hsmem_get_sid (hsm: hsmem ) (hid: hashcode): hsmem := | HSstore hsm chunk addr hlsv srce _ => HSstore hsm chunk addr hlsv srce hid end. +(* Symbolic final value -- from hash-consed values +It does not seem useful to hash-consed these final values (because they are final). +*) +Inductive hsfval := + | HSnone + | HScall (sig:signature) (svos: hsval + ident) (lsv:hlist_sval) (res:reg) (pc:node) + | HStailcall (sig:signature) (svos: hsval + ident) (lsv:hlist_sval) + | HSbuiltin (ef:external_function) (sargs: list (builtin_arg hsval)) (res: builtin_res reg) (pc:node) + | HSjumptable (sv: hsval) (tbl: list node) + | HSreturn (res:option hsval) +. Section CanonBuilding. @@ -281,7 +292,7 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): ?? (option hsistate) : | Istore chunk addr args src pc' => DO next <~ hslocal_store hst.(hsi_local) chunk addr args src;; RET (Some (hsist_set_local hst pc' next)) - | Icond cond args ifso ifnot _ => + | Icond cond args ifso ifnot _ => let prev := hst.(hsi_local) in DO vargs <~ hlist_args prev args ;; let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in @@ -289,56 +300,105 @@ Definition hsiexec_inst (i: instruction) (hst: hsistate): ?? (option hsistate) : | _ => RET None (* TODO jumptable ? *) end. +Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A := + match o with + | Some x => RET x + | None => FAILWITH msg + end. + Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): ?? hsistate := match path with | O => RET hst | S p => - DO i <~ match (fn_code f)!(hst.(hsi_pc)) with Some i => RET i | _ => FAILWITH "hsiexec_path: invalid node" end;; + DO i <~ some_or_fail ((fn_code f)!(hst.(hsi_pc))) "hsiexec_path.internal_error.1";; DO ohst1 <~ hsiexec_inst i hst;; - DO hst1 <~ match ohst1 with Some x => RET x | _ => FAILWITH "hsiexec_path: invalid path" end;; + DO hst1 <~ some_or_fail ohst1 "hsiexec_path.internal_error.2";; hsiexec_path p f hst1 end. -(* -Record hsstate := { hinternal:> hsistate; hfinal: sfval }. +Record hsstate := { hinternal:> hsistate; hfinal: hsfval }. + +Fixpoint hbuiltin_arg (hst: PTree.t hsval) (arg : builtin_arg reg): ?? builtin_arg hsval := + match arg with + | BA r => + DO v <~ hsi_sreg_get hst r;; + RET (BA v) + | BA_int n => RET (BA_int (A:=hsval) n) + | BA_long n => RET (BA_long (A:=hsval) n) + | BA_float f0 => RET (BA_float (A:=hsval) f0) + | BA_single s => RET (BA_single (A:=hsval) s) + | BA_loadstack chunk ptr => RET (BA_loadstack (A:=hsval) chunk ptr) + | BA_addrstack ptr => RET (BA_addrstack (A:=hsval) ptr) + | BA_loadglobal chunk id ptr => RET (BA_loadglobal (A:=hsval) chunk id ptr) + | BA_addrglobal id ptr => RET (BA_addrglobal (A:=hsval) id ptr) + | BA_splitlong ba1 ba2 => + DO v1 <~ hbuiltin_arg hst ba1;; + DO v2 <~ hbuiltin_arg hst ba2;; + RET (BA_splitlong v1 v2) + | BA_addptr ba1 ba2 => + DO v1 <~ hbuiltin_arg hst ba1;; + DO v2 <~ hbuiltin_arg hst ba2;; + RET (BA_addptr v1 v2) + end. + +Fixpoint hbuiltin_args (hst: PTree.t hsval) (args: list (builtin_arg reg)): ?? list (builtin_arg hsval) := + match args with + | nil => RET nil + | a::l => + DO ha <~ hbuiltin_arg hst a;; + DO hl <~ hbuiltin_args hst l;; + RET (ha::hl) + end. + +Definition hsum_left (hst: PTree.t hsval) (ros: reg + ident): ?? (hsval + ident) := + match ros with + | inl r => DO hr <~ hsi_sreg_get hst r;; RET (inl hr) + | inr s => RET (inr s) + end. -Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := +Definition hsexec_final (i: instruction) (hst: PTree.t hsval): ?? hsfval := match i with | Icall sig ros args res pc => - let svos := sum_left_map (hsi_sreg_get prev) ros in - let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - Scall sig svos sargs res pc + DO svos <~ hsum_left hst ros;; + DO sargs <~ hlist_args hst args;; + RET (HScall sig svos sargs res pc) | Itailcall sig ros args => - let svos := sum_left_map (hsi_sreg_get prev) ros in - let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in - Stailcall sig svos sargs + DO svos <~ hsum_left hst ros;; + DO sargs <~ hlist_args hst args;; + RET (HStailcall sig svos sargs) | Ibuiltin ef args res pc => - let sargs := List.map (builtin_arg_inj (hsi_sreg_get prev)) args in - Sbuiltin ef sargs res pc - | Ireturn or => - let sor := SOME r <- or IN Some (hsi_sreg_get prev r) in - Sreturn sor + DO sargs <~ hbuiltin_args hst args;; + RET (HSbuiltin ef sargs res pc) | Ijumptable reg tbl => - let sv := hsi_sreg_get prev reg in - Sjumptable sv tbl - | _ => Snone + DO sv <~ hsi_sreg_get hst reg;; + RET (HSjumptable sv tbl) + | Ireturn or => + match or with + | Some r => DO hr <~ hsi_sreg_get hst r;; RET (HSreturn (Some hr)) + | None => RET (HSreturn None) + end + | _ => RET (HSnone) end. -Definition init_hsistate_local := {| hsi_lsmem := Sinit::nil; hsi_ok_lsval := nil; hsi_sreg := PTree.empty sval |}. - -Definition init_hsistate pc := {| hsi_pc := pc; hsi_exits := nil; hsi_local := init_hsistate_local |}. - -Definition hsexec (f: function) (pc:node): option hsstate := - SOME path <- (fn_path f)!pc IN - SOME hst <- hsiexec_path path.(psize) f (init_hsistate pc) IN - SOME i <- (fn_code f)!(hst.(hsi_pc)) IN - Some (match hsiexec_inst i hst with - | Some hst' => {| hinternal := hst'; hfinal := Snone |} - | None => {| hinternal := hst; hfinal := hsexec_final i hst.(hsi_local) |} - end). - - -*) +Definition init_hsistate_local (_:unit): ?? hsistate_local + := DO hm <~ hSinit ();; + RET {| hsi_smem := hm; hsi_ok_lsval := nil; hsi_sreg := PTree.empty hsval |}. + +Definition init_hsistate pc: ?? hsistate + := DO hst <~ init_hsistate_local ();; + RET {| hsi_pc := pc; hsi_exits := nil; hsi_local := hst |}. + +Definition hsexec (f: function) (pc:node): ?? hsstate := + DO path <~ some_or_fail ((fn_path f)!pc) "hsexec.internal_error.1";; + DO hinit <~ init_hsistate pc;; + 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 ccf42e43c0270de82d53b1f0e4195f3573f3b088 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 23 Jul 2020 18:26:28 +0200 Subject: generalize builtin_arg_inj into builtin_arg_map --- kvx/lib/RTLpathSE_impl.v | 2 +- kvx/lib/RTLpathSE_theory.v | 17 +++++++++-------- 2 files changed, 10 insertions(+), 9 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v index 8adc92f9..afc9785e 100644 --- a/kvx/lib/RTLpathSE_impl.v +++ b/kvx/lib/RTLpathSE_impl.v @@ -1174,7 +1174,7 @@ Definition hsexec_final (i: instruction) (prev: hsistate_local): sfval := let sargs := list_sval_inj (List.map (hsi_sreg_get prev) args) in Stailcall sig svos sargs | Ibuiltin ef args res pc => - let sargs := List.map (builtin_arg_inj (hsi_sreg_get prev)) args in + let sargs := List.map (builtin_arg_map (hsi_sreg_get prev)) args in Sbuiltin ef sargs res pc | Ireturn or => let sor := SOME r <- or IN Some (hsi_sreg_get prev r) in diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v index 81351f6a..5002b7c5 100644 --- a/kvx/lib/RTLpathSE_theory.v +++ b/kvx/lib/RTLpathSE_theory.v @@ -998,7 +998,8 @@ Inductive ssem pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m ssem pge ge sp st stack f rs0 m0 t s . -Fixpoint builtin_arg_inj (f: reg -> sval) (arg: builtin_arg reg) : builtin_arg sval := +(* 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 @@ -1009,14 +1010,14 @@ Fixpoint builtin_arg_inj (f: reg -> sval) (arg: builtin_arg reg) : builtin_arg s | 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_inj f ba1) (builtin_arg_inj f ba2) - | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_inj f ba1) (builtin_arg_inj f ba2) + | 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, seval_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_inj sreg 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). @@ -1029,7 +1030,7 @@ Qed. Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs: (forall r, seval_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_inj sreg) args) vargs. + seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs. Proof. induction 2. - constructor. @@ -1039,7 +1040,7 @@ Qed. Lemma seval_builtin_arg_complete ge sp rs m rs0 m0 sreg: forall arg varg, (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - seval_builtin_arg ge sp m rs0 m0 (builtin_arg_inj sreg arg) varg -> + 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. @@ -1053,7 +1054,7 @@ Qed. Lemma seval_builtin_args_complete ge sp rs m rs0 m0 sreg: forall args vargs, (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) -> - seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_inj sreg) args) vargs -> + 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. @@ -1075,7 +1076,7 @@ Definition sexec_final (i: instruction) (prev: sistate_local): sfval := let sargs := list_sval_inj (List.map prev.(si_sreg) args) in Stailcall sig svos sargs | Ibuiltin ef args res pc => - let sargs := List.map (builtin_arg_inj prev.(si_sreg)) args in + let sargs := List.map (builtin_arg_map prev.(si_sreg)) args in Sbuiltin ef sargs res pc | Ireturn or => let sor := SOME r <- or IN Some (prev.(si_sreg) r) in -- cgit From 5492ba77fdab3776d58b838ed9a2c5b043f255f5 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Thu, 23 Jul 2020 20:08:30 +0200 Subject: starting implem of hsilocal_simu_test --- kvx/lib/RTLpathSE_impl_junk.v | 194 ++++++++++++++++-------------------------- 1 file changed, 75 insertions(+), 119 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index 708c98f5..0bf67110 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -74,7 +74,7 @@ Definition hlist_sval_set_hid (hlsv: hlist_sval) (hid: hashcode): hlist_sval := | HScons hsv hlsv _ => HScons hsv hlsv hid end. -Definition hsmem_get_sid (hsm: hsmem ) (hid: hashcode): hsmem := +Definition hsmem_set_hid (hsm: hsmem ) (hid: hashcode): hsmem := match hsm with | HSinit _ => HSinit hid | HSstore hsm chunk addr hlsv srce _ => HSstore hsm chunk addr hlsv srce hid @@ -402,133 +402,89 @@ Definition hsexec (f: function) (pc:node): ?? hsstate := End CanonBuilding. -(* -(** * The simulation test of concrete symbolic execution *) +(** * The simulation test of concrete hash-consed symbolic execution *) + +(* Now, we build the hash-Cons value from a "hash_eq". + +Informal specification: + [hash_eq] must be consistent with the "hashed" constructors defined above. -(* TODO - generalize it with a list of registers to test ? *) -Definition hsilocal_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate_local) := OK tt. +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 revmap_check_single (dm: PTree.t node) (n tn: node) : res unit := - match dm ! tn with - | None => Error (msg "revmap_check_single: no mapping for tn") - | Some res => if (Pos.eq_dec n res) then OK tt - else Error (msg "revmap_check_single: n and res do not match") +Definition hsval_hash_eq (sv1 sv2: hsval): ?? bool := + match sv1, sv2 with + | HSinput r1 _, HSinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *) + | HSop op1 lsv1 sm1 _, HSop op2 lsv2 sm2 _ => + DO b1 <~ phys_eq lsv1 lsv2;; + DO b2 <~ phys_eq sm1 sm2;; + if b1 && b2 + then struct_eq op1 op2 (* NB: really need a struct_eq here ? *) + else RET false + | HSload sm1 trap1 chk1 addr1 lsv1 _, HSload 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. -(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) -Fixpoint sval_simub (sv1 sv2: sval) := - match sv1 with - | Sinput r => - match sv2 with - | Sinput r' => if (Pos.eq_dec r r') then OK tt else Error (msg "sval_simub: Sinput different registers") - | _ => Error (msg "sval_simub: Sinput expected") - end - | Sop op lsv sm => - match sv2 with - | Sop op' lsv' sm' => - if (eq_operation op op') then - do _ <- list_sval_simub lsv lsv'; - smem_simub sm sm' - else Error (msg "sval_simub: different operations in Sop") - | _ => Error (msg "sval_simub: Sop expected") - end - | Sload sm trap chk addr lsv => - match sv2 with - | Sload sm' trap' chk' addr' lsv' => - if (trapping_mode_eq trap trap') then - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - list_sval_simub lsv lsv' - else Error (msg "sval_simub: addressings do not match") - else Error (msg "sval_simub: chunks do not match") - else Error (msg "sval_simub: trapping modes do not match") - (* FIXME - should be refined once we introduce non trapping loads *) - | _ => Error (msg "sval_simub: Sload expected") - end - end -with list_sval_simub (lsv1 lsv2: list_sval) := - match lsv1 with - | Snil => - match lsv2 with - | Snil => OK tt - | _ => Error (msg "list_sval_simub: lists of different lengths (lsv2 is bigger)") - end - | Scons sv1 lsv1 => - match lsv2 with - | Snil => Error (msg "list_sval_simub: lists of different lengths (lsv1 is bigger)") - | Scons sv2 lsv2 => - do _ <- sval_simub sv1 sv2; - list_sval_simub lsv1 lsv2 - end - end -with smem_simub (sm1 sm2: smem) := - match sm1 with - | Sinit => - match sm2 with - | Sinit => OK tt - | _ => Error (msg "smem_simub: Sinit expected") - end - | Sstore sm chk addr lsv sv => - match sm2 with - | Sstore sm' chk' addr' lsv' sv' => - if (chunk_eq chk chk') then - if (eq_addressing addr addr') then - do _ <- smem_simub sm sm'; - do _ <- list_sval_simub lsv lsv'; - sval_simub sv sv' - else Error (msg "smem_simub: addressings do not match") - else Error (msg "smem_simub: chunks not match") - | _ => Error (msg "smem_simub: Sstore expected") - end +Definition hlist_sval_hash_eq (lsv1 lsv2: hlist_sval): ?? bool := + match lsv1, lsv2 with + | HSnil _, HSnil _ => RET true + | HScons sv1 lsv1' _, HScons sv2 lsv2' _ => + DO b <~ phys_eq lsv1' lsv2';; + if b + then phys_eq sv1 sv2 + else RET false + | _,_ => RET false end. -Lemma sval_simub_correct sv1: forall sv2, - sval_simub sv1 sv2 = OK tt -> sv1 = sv2. -Proof. - induction sv1 using sval_mut with - (P0 := fun lsv1 => forall lsv2, list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2) - (P1 := fun sm1 => forall sm2, smem_simub sm1 sm2 = OK tt -> sm1 = sm2); simpl; auto. - (* Sinput *) - - destruct sv2; try discriminate. - destruct (Pos.eq_dec r r0); [congruence|discriminate]. - (* Sop *) - - destruct sv2; try discriminate. - destruct (eq_operation _ _); [|discriminate]. subst. - intro. explore. apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sload *) - - destruct sv2; try discriminate. - destruct (trapping_mode_eq _ _); [|discriminate]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. assert (sm = sm0) by auto. assert (lsv = lsv0) by auto. - congruence. - (* Snil *) - - destruct lsv2; [|discriminate]. congruence. - (* Scons *) - - destruct lsv2; [discriminate|]. intro. explore. - apply IHsv1 in EQ. apply IHsv0 in EQ0. congruence. - (* Sinit *) - - destruct sm2; [|discriminate]. congruence. - (* Sstore *) - - destruct sm2; [discriminate|]. - destruct (chunk_eq _ _); [|discriminate]. - destruct (eq_addressing _ _); [|discriminate]. - intro. explore. - assert (sm = sm2) by auto. assert (lsv = lsv0) by auto. assert (sv1 = srce) by auto. - congruence. -Qed. +Definition hsmem_hash_eq (sm1 sm2: hsmem): ?? bool := + match sm1, sm2 with + | HSinit _, HSinit _ => RET true + | HSstore sm1 chk1 addr1 lsv1 sv1 _, HSstore 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 list_sval_simub_correct lsv1: forall lsv2, - list_sval_simub lsv1 lsv2 = OK tt -> lsv1 = lsv2. -Proof. - induction lsv1; simpl; auto. - - destruct lsv2; try discriminate. reflexivity. - - destruct lsv2; try discriminate. intro. explore. - apply sval_simub_correct in EQ. assert (lsv1 = lsv2) by auto. - congruence. +Definition hSVAL: hashP hsval := {| hash_eq := hsval_hash_eq; get_hid:=hsval_get_hid; set_hid:=hsval_set_hid |}. +Definition hLSVAL: hashP hlist_sval := {| hash_eq := hlist_sval_hash_eq; get_hid:= hlist_sval_get_hid; set_hid:= hlist_sval_set_hid |}. +Definition hSMEM: hashP hsmem := {| hash_eq := hsmem_hash_eq; get_hid:= hsmem_get_hid; set_hid:= hsmem_set_hid |}. + +Program Definition mk_hash_params: Dict.hash_params hsval := + {| + Dict.test_eq := phys_eq; + Dict.hashing := fun (ht: hsval) => RET (hsval_get_hid ht); + Dict.log := fun _ => RET () (* NB no log *) |}. +Obligation 1. + wlp_simplify. Qed. +Definition revmap_check_single (dm: PTree.t node) (n tn: node) : ?? unit := + DO res <~ some_or_fail (dm ! tn) "revmap_check_single: no mapping for tn";; + DO b <~ phys_eq n res;; (* NB: phys_eq on nodes is probably sufficient *) + assert_b b "revmap_check_single: n and res are physically different";; + RET tt. + +Definition hsilocal_simu_test hst1 hst2 : ?? unit := + DO b <~ phys_eq (hsi_smem hst2) (hsi_smem hst1);; + assert_b b "hsi_smem sets aren't equiv";; + Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);; + (* TODO - compare on the whole ptree *) RET tt. + +(* Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); -- cgit From ed743b5cfaeea5f23d6b00e7ee57359f7bce87d6 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Jul 2020 09:03:26 +0200 Subject: Implem of hsiexit_simu_check --- kvx/lib/RTLpathSE_impl_junk.v | 269 ++++++++++++++++++++++++++---------------- 1 file changed, 167 insertions(+), 102 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index 0bf67110..359f4578 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -24,7 +24,9 @@ Local Open Scope impure. Import ListNotations. Local Open Scope list_scope. -(** * Refinement of symbolic values/symbolic memories with hash-consing data *) +(** * Implementation of Data-structure use in Hash-consing *) + +(** ** Implementation of symbolic values/symbolic memories with hash-consing data *) Inductive hsval := | HSinput (r: reg) (hid:hashcode) @@ -80,6 +82,75 @@ Definition hsmem_set_hid (hsm: hsmem ) (hid: hashcode): hsmem := | HSstore hsm chunk addr hlsv srce _ => HSstore hsm chunk addr hlsv 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 hsval_hash_eq (sv1 sv2: hsval): ?? bool := + match sv1, sv2 with + | HSinput r1 _, HSinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *) + | HSop op1 lsv1 sm1 _, HSop op2 lsv2 sm2 _ => + DO b1 <~ phys_eq lsv1 lsv2;; + DO b2 <~ phys_eq sm1 sm2;; + if b1 && b2 + then struct_eq op1 op2 (* NB: really need a struct_eq here ? *) + else RET false + | HSload sm1 trap1 chk1 addr1 lsv1 _, HSload 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. + +Definition hlist_sval_hash_eq (lsv1 lsv2: hlist_sval): ?? bool := + match lsv1, lsv2 with + | HSnil _, HSnil _ => RET true + | HScons sv1 lsv1' _, HScons sv2 lsv2' _ => + DO b <~ phys_eq lsv1' lsv2';; + if b + then phys_eq sv1 sv2 + else RET false + | _,_ => RET false + end. + +Definition hsmem_hash_eq (sm1 sm2: hsmem): ?? bool := + match sm1, sm2 with + | HSinit _, HSinit _ => RET true + | HSstore sm1 chk1 addr1 lsv1 sv1 _, HSstore 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. + +Definition hSVAL: hashP hsval := {| hash_eq := hsval_hash_eq; get_hid:=hsval_get_hid; set_hid:=hsval_set_hid |}. +Definition hLSVAL: hashP hlist_sval := {| hash_eq := hlist_sval_hash_eq; get_hid:= hlist_sval_get_hid; set_hid:= hlist_sval_set_hid |}. +Definition hSMEM: hashP hsmem := {| hash_eq := hsmem_hash_eq; get_hid:= hsmem_get_hid; set_hid:= hsmem_set_hid |}. + +Program Definition mk_hash_params: Dict.hash_params hsval := + {| + Dict.test_eq := phys_eq; + Dict.hashing := fun (ht: hsval) => RET (hsval_get_hid ht); + Dict.log := fun _ => RET () (* NB no log *) |}. +Obligation 1. + wlp_simplify. +Qed. + + (* Symbolic final value -- from hash-consed values It does not seem useful to hash-consed these final values (because they are final). *) @@ -92,6 +163,35 @@ Inductive hsfval := | HSreturn (res:option hsval) . +(** ** Implementation of symbolic states +*) + +(** name : Hash-consed Symbolic Internal state local. *) +Record hsistate_local := + { + (** [hsi_smem] represents the current smem symbolic evaluations. + (we can recover the previous one from smem) *) + hsi_smem:> hsmem; + (** For the values in registers: + 1) we store a list of sval evaluations + 2) we encode the symbolic regset by a PTree *) + hsi_ok_lsval: list hsval; + hsi_sreg:> PTree.t hsval + }. + +(* Syntax and semantics of symbolic exit states *) +Record hsistate_exit := mk_hsistate_exit + { hsi_cond: condition; hsi_scondargs: hlist_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. + + +(** ** Syntax and Semantics of symbolic internal state *) +Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. + +(** ** Syntax and Semantics of symbolic state *) +Record hsstate := { hinternal:> hsistate; hfinal: hsfval }. + + +(** * Implementation of symbolic execution *) Section CanonBuilding. Variable hC_hsval: hashinfo hsval -> ?? hsval. @@ -160,23 +260,6 @@ Definition hSstore (hsm: hsmem) (chunk:memory_chunk) (addr:addressing) (hlsv:hli hC_hsmem {| hdata:=HSstore hsm chunk addr hlsv srce unknown_hid; hcodes := hv |}. -(** * Implementation of local symbolic internal states -*) - -(** name : Hash-consed Symbolic Internal state local. Later on we will use the - refinement to introduce hash consing *) -Record hsistate_local := - { - (** [hsi_smem] represents the current smem symbolic evaluations. - (we can recover the previous one from smem) *) - hsi_smem:> hsmem; - (** For the values in registers: - 1) we store a list of sval evaluations - 2) we encode the symbolic regset by a PTree *) - hsi_ok_lsval: list hsval; - hsi_sreg:> PTree.t hsval - }. - Definition hsi_sreg_get (hst: PTree.t hsval) r: ?? hsval := match PTree.get r hst with | None => hSinput r @@ -192,15 +275,6 @@ Fixpoint hlist_args (hst: PTree.t hsval) (l: list reg): ?? hlist_sval := hScons v hlsv end. - -(* Syntax and semantics of symbolic exit states *) -Record hsistate_exit := mk_hsistate_exit - { hsi_cond: condition; hsi_scondargs: hlist_sval; hsi_elocal: hsistate_local; hsi_ifso: node }. - - -(** * Syntax and Semantics of symbolic internal state *) -Record hsistate := { hsi_pc: node; hsi_exits: list hsistate_exit; hsi_local: hsistate_local }. - (** ** Assignment of memory *) Definition hslocal_store (hst:hsistate_local) chunk addr args src: ?? hsistate_local := let pt := hst.(hsi_sreg) in @@ -316,8 +390,6 @@ Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): ?? hsistate := hsiexec_path p f hst1 end. -Record hsstate := { hinternal:> hsistate; hfinal: hsfval }. - Fixpoint hbuiltin_arg (hst: PTree.t hsval) (arg : builtin_arg reg): ?? builtin_arg hsval := match arg with | BA r => @@ -404,95 +476,88 @@ End CanonBuilding. (** * The simulation test of concrete hash-consed symbolic execution *) -(* 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: +Definition phys_check {A} (x y:A) (msg: pstring): ?? unit := + DO b <~ phys_eq x y;; + assert_b b msg;; + RET tt. - (hash_eq (hdata x) (hdata y) ~> true) <-> (hcodes x)=(hcodes y) -*) +Definition struct_check {A} (x y:A) (msg: pstring): ?? unit := + DO b <~ struct_eq x y;; + assert_b b msg;; + RET tt. -Definition hsval_hash_eq (sv1 sv2: hsval): ?? bool := - match sv1, sv2 with - | HSinput r1 _, HSinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *) - | HSop op1 lsv1 sm1 _, HSop op2 lsv2 sm2 _ => - DO b1 <~ phys_eq lsv1 lsv2;; - DO b2 <~ phys_eq sm1 sm2;; - if b1 && b2 - then struct_eq op1 op2 (* NB: really need a struct_eq here ? *) - else RET false - | HSload sm1 trap1 chk1 addr1 lsv1 _, HSload 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 +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. -Definition hlist_sval_hash_eq (lsv1 lsv2: hlist_sval): ?? bool := - match lsv1, lsv2 with - | HSnil _, HSnil _ => RET true - | HScons sv1 lsv1' _, HScons sv2 lsv2' _ => - DO b <~ phys_eq lsv1' lsv2';; - if b - then phys_eq sv1 sv2 - else RET false - | _,_ => RET false +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. +Global 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. -Definition hsmem_hash_eq (sm1 sm2: hsmem): ?? bool := - match sm1, sm2 with - | HSinit _, HSinit _ => RET true - | HSstore sm1 chk1 addr1 lsv1 sv1 _, HSstore 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 +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. + +Fixpoint PTree_frame_eq_check {A} (frame: list positive) (d1 d2: PTree.t A): ?? unit := + match frame with + | nil => RET tt + | k::l => + option_eq_check (PTree.get k d1) (PTree.get k d2);; + PTree_frame_eq_check l d1 d2 end. -Definition hSVAL: hashP hsval := {| hash_eq := hsval_hash_eq; get_hid:=hsval_get_hid; set_hid:=hsval_set_hid |}. -Definition hLSVAL: hashP hlist_sval := {| hash_eq := hlist_sval_hash_eq; get_hid:= hlist_sval_get_hid; set_hid:= hlist_sval_set_hid |}. -Definition hSMEM: hashP hsmem := {| hash_eq := hsmem_hash_eq; get_hid:= hsmem_get_hid; set_hid:= hsmem_set_hid |}. - -Program Definition mk_hash_params: Dict.hash_params hsval := - {| - Dict.test_eq := phys_eq; - Dict.hashing := fun (ht: hsval) => RET (hsval_get_hid ht); - Dict.log := fun _ => RET () (* NB no log *) |}. -Obligation 1. - wlp_simplify. +Lemma PTree_frame_eq_check_correct A l (d1 d2: t A): + WHEN PTree_frame_eq_check l d1 d2 ~> _ THEN forall x, List.In x l -> PTree.get x d1 = PTree.get x d2. +Proof. + induction l as [|k l]; simpl; wlp_simplify. + subst; auto. Qed. +Global Opaque PTree_frame_eq_check. + +Definition hsilocal_simu_check hst1 hst2 : ?? unit := + phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_simu_check: hsi_smem sets aren't equiv";; + Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);; + PTree_eq_check (hsi_sreg hst1) (hsi_sreg hst2). Definition revmap_check_single (dm: PTree.t node) (n tn: node) : ?? unit := DO res <~ some_or_fail (dm ! tn) "revmap_check_single: no mapping for tn";; - DO b <~ phys_eq n res;; (* NB: phys_eq on nodes is probably sufficient *) - assert_b b "revmap_check_single: n and res are physically different";; - RET tt. + phys_check n res "revmap_check_single: n and res are physically different". -Definition hsilocal_simu_test hst1 hst2 : ?? unit := - DO b <~ phys_eq (hsi_smem hst2) (hsi_smem hst1);; - assert_b b "hsi_smem sets aren't equiv";; +Definition hsilocal_frame_simu_check frame hst1 hst2 : ?? unit := + phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_frame_simu_check: hsi_smem sets aren't equiv";; Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);; - (* TODO - compare on the whole ptree *) RET tt. + PTree_frame_eq_check frame (hsi_sreg hst1) (hsi_sreg hst2). -(* -Definition hsiexit_simub (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit) := - if (eq_condition (hsi_cond hse1) (hsi_cond hse2)) then - do _ <- list_sval_simub (hsi_scondargs hse1) (hsi_scondargs hse2); - do _ <- hsilocal_simub dm f (hsi_elocal hse1) (hsi_elocal hse2); - revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2) - else Error (msg "siexit_simub: conditions do not match") -. +Definition hsiexit_simu_check (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit): ?? unit := + struct_check (hsi_cond hse1) (hsi_cond hse2) "hsiexit_simu_check: conditions do not match";; + phys_check (hsi_scondargs hse1) (hsi_scondargs hse2) "hsiexit_simu_check: args do not match";; + revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2);; + DO path <~ some_or_fail ((fn_path f) ! (hsi_ifso hse1)) "hsiexit_simu_check: internal error";; + hsilocal_frame_simu_check (Regset.elements path.(input_regs)) (hsi_elocal hse1) (hsi_elocal hse2). +(* Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := match lhse1 with | nil => -- cgit From 05cca27e2beacc7949aa54a35ac6528858402116 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Jul 2020 09:48:32 +0200 Subject: hsstate_simu_check --- kvx/lib/RTLpathSE_impl_junk.v | 309 ++++++++++++------------------------------ 1 file changed, 89 insertions(+), 220 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index 359f4578..f08998b4 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -395,14 +395,14 @@ Fixpoint hbuiltin_arg (hst: PTree.t hsval) (arg : builtin_arg reg): ?? builtin_a | BA r => DO v <~ hsi_sreg_get hst r;; RET (BA v) - | BA_int n => RET (BA_int (A:=hsval) n) - | BA_long n => RET (BA_long (A:=hsval) n) - | BA_float f0 => RET (BA_float (A:=hsval) f0) - | BA_single s => RET (BA_single (A:=hsval) s) - | BA_loadstack chunk ptr => RET (BA_loadstack (A:=hsval) chunk ptr) - | BA_addrstack ptr => RET (BA_addrstack (A:=hsval) ptr) - | BA_loadglobal chunk id ptr => RET (BA_loadglobal (A:=hsval) chunk id ptr) - | BA_addrglobal id ptr => RET (BA_addrglobal (A:=hsval) id ptr) + | 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 hst ba1;; DO v2 <~ hbuiltin_arg hst ba2;; @@ -557,218 +557,87 @@ Definition hsiexit_simu_check (dm: PTree.t node) (f: RTLpath.function) (hse1 hse DO path <~ some_or_fail ((fn_path f) ! (hsi_ifso hse1)) "hsiexit_simu_check: internal error";; hsilocal_frame_simu_check (Regset.elements path.(input_regs)) (hsi_elocal hse1) (hsi_elocal hse2). -(* -Fixpoint hsiexits_simub (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := - match lhse1 with - | nil => - match lhse2 with - | nil => OK tt - | _ => Error (msg "siexists_simub: sle1 and sle2 lengths do not match") - end - | hse1 :: lhse1 => - match lhse2 with - | nil => Error (msg "siexits_simub: sle1 and sle2 lengths do not match") - | hse2 :: lhse2 => - do _ <- hsiexit_simub dm f hse1 hse2; - do _ <- hsiexits_simub dm f lhse1 lhse2; - OK tt - end - end. - -Definition hsistate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := - do _ <- hsilocal_simub dm f (hsi_local hst1) (hsi_local hst2); - do _ <- hsiexits_simub dm f (hsi_exits hst1) (hsi_exits hst2); - OK tt. - -Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): res unit := - match ln with - | nil => - match ln' with - | nil => OK tt - | _ => Error (msg "revmap_check_list: lists have different lengths") - end - | n::ln => - match ln' with - | nil => Error (msg "revmap_check_list: lists have different lengths") - | n'::ln' => do _ <- revmap_check_single dm n n'; revmap_check_list dm ln ln' - end - end. - -Definition svos_simub (svos1 svos2: sval + ident) := - match svos1 with - | inl sv1 => - match svos2 with - | inl sv2 => sval_simub sv1 sv2 - | _ => Error (msg "svos_simub: expected sval") - end - | inr id1 => - match svos2 with - | inr id2 => - if (ident_eq id1 id2) then OK tt - else Error (msg "svos_simub: ids do not match") - | _ => Error (msg "svos_simub: expected id") - end - end. - -Fixpoint builtin_arg_simub (bs bs': builtin_arg sval) := - match bs with - | BA sv => - match bs' with - | BA sv' => sval_simub sv sv' - | _ => Error (msg "builtin_arg_simub: BA expected") - end - | BA_int n => - match bs' with - | BA_int n' => if (Integers.int_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") - | _ => Error (msg "buitin_arg_simub: BA_int expected") - end - | BA_long n => - match bs' with - | BA_long n' => if (int64_eq n n') then OK tt else Error (msg "builtin_arg_simub: integers do not match") - | _ => Error (msg "buitin_arg_simub: BA_long expected") - end - | BA_float f => - match bs' with - | BA_float f' => if (float_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") - | _ => Error (msg "builtin_arg_simub: BA_float expected") - end - | BA_single f => - match bs' with - | BA_single f' => if (float32_eq f f') then OK tt else Error (msg "builtin_arg_simub: floats do not match") - | _ => Error (msg "builtin_arg_simub: BA_single expected") - end - | BA_loadstack chk ptr => - match bs' with - | BA_loadstack chk' ptr' => - if (chunk_eq chk chk') then - if (ptrofs_eq ptr ptr') then OK tt - else Error (msg "builtin_arg_simub: ptr do not match") - else Error (msg "builtin_arg_simub: chunks do not match") - | _ => Error (msg "builtin_arg_simub: BA_loadstack expected") - end - | BA_addrstack ptr => - match bs' with - | BA_addrstack ptr' => if (ptrofs_eq ptr ptr') then OK tt else Error (msg "builtin_arg_simub: ptr do not match") - | _ => Error (msg "builtin_arg_simub: BA_addrstack expected") - end - | BA_loadglobal chk id ofs => - match bs' with - | BA_loadglobal chk' id' ofs' => - if (chunk_eq chk chk') then - if (ident_eq id id') then - if (ptrofs_eq ofs ofs') then OK tt - else Error (msg "builtin_arg_simub: offsets do not match") - else Error (msg "builtin_arg_simub: identifiers do not match") - else Error (msg "builtin_arg_simub: chunks do not match") - | _ => Error (msg "builtin_arg_simub: BA_loadglobal expected") - end - | BA_addrglobal id ofs => - match bs' with - | BA_addrglobal id' ofs' => - if (ident_eq id id') then - if (ptrofs_eq ofs ofs') then OK tt - else Error (msg "builtin_arg_simub: offsets do not match") - else Error (msg "builtin_arg_simub: identifiers do not match") - | _ => Error (msg "builtin_arg_simub: BA_addrglobal expected") - end - | BA_splitlong lo hi => - match bs' with - | BA_splitlong lo' hi' => do _ <- builtin_arg_simub lo lo'; builtin_arg_simub hi hi' - | _ => Error (msg "builtin_arg_simub: BA_splitlong expected") - end - | BA_addptr b1 b2 => - match bs' with - | BA_addptr b1' b2' => do _ <- builtin_arg_simub b1 b1'; builtin_arg_simub b2 b2' - | _ => Error (msg "builtin_arg_simub: BA_addptr expected") - end - end. - -Fixpoint list_builtin_arg_simub lbs1 lbs2 := - match lbs1 with - | nil => - match lbs2 with - | nil => OK tt - | _ => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs2 is bigger)") - end - | bs1::lbs1 => - match lbs2 with - | nil => Error (msg "list_builtin_arg_simub: lists of different lengths (lbs1 is bigger)") - | bs2::lbs2 => - do _ <- builtin_arg_simub bs1 bs2; - list_builtin_arg_simub lbs1 lbs2 - end - end. - -(* WARNING: ce code va bouger pas mal quand on aura le hash-consing ! *) -Definition sfval_simub (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: sfval) := - match fv1 with - | Snone => - match fv2 with - | Snone => revmap_check_single dm pc1 pc2 - | _ => Error (msg "sfval_simub: Snone expected") - end - | Scall sig svos lsv res pc1 => - match fv2 with - | Scall sig2 svos2 lsv2 res2 pc2 => - do _ <- revmap_check_single dm pc1 pc2; - if (signature_eq sig sig2) then - if (Pos.eq_dec res res2) then - do _ <- svos_simub svos svos2; - list_sval_simub lsv lsv2 - else Error (msg "sfval_simub: Scall res do not match") - else Error (msg "sfval_simub: Scall different signatures") - | _ => Error (msg "sfval_simub: Scall expected") - end - | Stailcall sig svos lsv => - match fv2 with - | Stailcall sig' svos' lsv' => - if (signature_eq sig sig') then - do _ <- svos_simub svos svos'; - list_sval_simub lsv lsv' - else Error (msg "sfval_simub: signatures do not match") - | _ => Error (msg "sfval_simub: Stailcall expected") - end - | Sbuiltin ef lbs br pc => - match fv2 with - | Sbuiltin ef' lbs' br' pc' => - if (external_function_eq ef ef') then - if (builtin_res_eq_pos br br') then - do _ <- revmap_check_single dm pc pc'; - list_builtin_arg_simub lbs lbs' - else Error (msg "sfval_simub: builtin res do not match") - else Error (msg "sfval_simub: external functions do not match") - | _ => Error (msg "sfval_simub: Sbuiltin expected") - end - | Sjumptable sv ln => - match fv2 with - | Sjumptable sv' ln' => - do _ <- revmap_check_list dm ln ln'; sval_simub sv sv' - | _ => Error (msg "sfval_simub: Sjumptable expected") - end - | Sreturn osv => - match fv2 with - | Sreturn osv' => - match osv with - | Some sv => - match osv' with - | Some sv' => sval_simub sv sv' - | None => Error (msg "sfval_simub sv' expected") - end - | None => - match osv' with - | Some sv' => Error (msg "None expected") - | None => OK tt - end - end - | _ => Error (msg "sfval_simub: Sreturn expected") - end - end. - -Definition hsstate_simub (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := - do u1 <- hsistate_simub dm f (hinternal hst1) (hinternal hst2); - do u2 <- sfval_simub dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2); - OK tt. +Fixpoint hsiexits_simu_check (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) := + match lhse1,lhse2 with + | nil, nil => RET tt + | hse1 :: lhse1, hse2 :: lhse2 => + hsiexit_simu_check dm f hse1 hse2;; + hsiexits_simu_check dm f lhse1 lhse2 + | _, _ => FAILWITH "siexists_simu_check: lengths do not match" + end. + +Definition hsistate_simu_check (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) := + hsilocal_simu_check (hsi_local hst1) (hsi_local hst2);; + hsiexits_simu_check dm f (hsi_exits hst1) (hsi_exits hst2). + +Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): ?? unit := + match ln, ln' with + | nil, nil => RET tt + | n::ln, n'::ln' => + revmap_check_single dm n n';; + revmap_check_list dm ln ln' + | _, _ => FAILWITH "revmap_check_list: lists have different lengths" + end. + +Definition svos_simu_check (svos1 svos2: hsval + 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. + +Fixpoint builtin_arg_simu_check (bs bs': builtin_arg hsval) := + match bs, bs' with + | BA sv, BA sv' => phys_check sv sv' "builtin_arg_simu_check: sval mismatch" + | BA_splitlong lo hi, BA_splitlong lo' hi' => + builtin_arg_simu_check lo lo';; + builtin_arg_simu_check hi hi' + | BA_addptr b1 b2, BA_addptr b1' b2' => + builtin_arg_simu_check b1 b1';; + builtin_arg_simu_check b2 b2' + | _, _ => struct_check bs bs' "builtin_arg_simu_check: basic mismatch" + end. + +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. + +Definition sfval_simu_check (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: hsfval) := + match fv1, fv2 with + | HSnone, HSnone => revmap_check_single dm pc1 pc2 + | HScall sig1 svos1 lsv1 res1 pc1, HScall sig2 svos2 lsv2 res2 pc2 => + revmap_check_single dm pc1 pc2;; + 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" + | HStailcall sig1 svos1 lsv1, HStailcall 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" + | HSbuiltin ef1 lbs1 br1 pc1, HSbuiltin ef2 lbs2 br2 pc2 => + revmap_check_single dm pc1 pc2;; + phys_check ef1 ef2 "sfval_simu_check: builtin ef do not match";; + phys_check br1 br2 "sfval_simu_check: builtin br do not match";; + list_builtin_arg_simu_check lbs1 lbs2 + | HSjumptable sv ln, HSjumptable sv' ln' => + revmap_check_list dm ln ln';; + phys_check sv sv' "sfval_simu_check: Sjumptable sval do not match" + | HSreturn osv1, HSreturn osv2 => + option_eq_check osv1 osv2 + | _, _ => FAILWITH "sfval_simu_check: structure mismatch" + end. + +Definition hsstate_simu_check (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) := + hsistate_simu_check dm f (hinternal hst1) (hinternal hst2);; + sfval_simu_check dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2). +(* Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := let (pc2, pc1) := m in match (hsexec f pc1) with @@ -776,7 +645,7 @@ Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpa | Some hst1 => match (hsexec tf pc2) with | None => Error (msg "simu_check_single: hsexec tf pc2 failed") - | Some hst2 => hsstate_simub dm f hst1 hst2 + | Some hst2 => hsstate_simu_check dm f hst1 hst2 end end. -- cgit From 2a48e12f8c62e33e6a9f3c172a6fb9dfc85e887d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Jul 2020 12:31:47 +0200 Subject: first draft of simu_check --- kvx/lib/RTLpathSE_impl_junk.v | 102 ++++++++++++++++++++++++++++++++++-------- 1 file changed, 83 insertions(+), 19 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index f08998b4..e46dddba 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -637,35 +637,99 @@ Definition hsstate_simu_check (dm: PTree.t node) (f: RTLpath.function) (hst1 hst hsistate_simu_check dm f (hinternal hst1) (hinternal hst2);; sfval_simu_check dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2). -(* -Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node) := +Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node): ?? unit := let (pc2, pc1) := m in - match (hsexec f pc1) with - | None => Error (msg "simu_check_single: hsexec f pc1 failed") - | Some hst1 => - match (hsexec tf pc2) with - | None => Error (msg "simu_check_single: hsexec tf pc2 failed") - | Some hst2 => hsstate_simu_check dm f hst1 hst2 - end - end. + (* creating the hash-consing tables *) + DO hC_sval <~ hCons hSVAL;; + DO hC_hlist_sval <~ hCons hLSVAL;; + DO hC_hsmem <~ hCons hSMEM;; + let hsexec := hsexec hC_sval.(hC) hC_hlist_sval.(hC) hC_hsmem.(hC) in + (* performing the hash-consed executions *) + DO hst1 <~ hsexec f pc1;; + DO hst2 <~ hsexec tf pc2;; + (* comparing the executions *) + hsstate_simu_check dm f hst1 hst2. + +Lemma simu_check_single_correct dm tf f pc1 pc2: + WHEN simu_check_single dm f tf (pc2, pc1) ~> _ THEN + sexec_simu dm f tf pc1 pc2. +Admitted. +Global Opaque simu_check_single. +Global Hint Resolve simu_check_single_correct: wlp. -Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm := +Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm : ?? unit := match lm with - | nil => OK tt - | m :: lm => do u1 <- simu_check_single dm f tf m; - do u2 <- simu_check_rec dm f tf lm; - OK tt + | nil => RET tt + | m :: lm => + simu_check_single dm f tf m;; + simu_check_rec dm f tf lm + 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. +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 simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := +Definition imp_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? unit := simu_check_rec dm f tf (PTree.elements dm). -*) +Local Hint Resolve PTree.elements_correct: core. +Lemma imp_simu_check_correct dm f tf: + WHEN imp_simu_check dm f tf ~> _ THEN + forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2. +Proof. + wlp_simplify. +Qed. +Global Opaque imp_simu_check. +Global Hint Resolve imp_simu_check_correct: wlp. + +Program Definition aux_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? bool := + DO r <~ + (TRY + imp_simu_check dm f tf;; + RET true + CATCH_FAIL s, _ => + println ("simu_check_failure:" +; s);; + RET false + ENSURE (fun b => b=true -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2));; + RET (`r). +Obligation 1. + split; wlp_simplify. discriminate. +Qed. + +Lemma aux_simu_check_correct dm f tf: + WHEN aux_simu_check dm f tf ~> b THEN + b=true -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2. +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). *) -Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) := OK tt. +Import UnsafeImpure. + +Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) : res unit := + match unsafe_coerce (aux_simu_check dm f tf) with + | Some true => OK tt + | _ => Error (msg "simu_check has failed") + end. Lemma simu_check_correct dm f tf: simu_check dm f tf = OK tt -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2. -Admitted. \ No newline at end of file +Proof. + unfold simu_check. + destruct (unsafe_coerce (aux_simu_check dm f tf)) as [[|]|] eqn:Hres; simpl; try discriminate. + intros; eapply aux_simu_check_correct; eauto. + eapply unsafe_coerce_not_really_correct; eauto. +Qed. \ No newline at end of file -- cgit From 0e6cf8850f4963cc842f701471d82e5381662f5c Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Jul 2020 12:35:55 +0200 Subject: fix phys_eq -> struct_eq --- kvx/lib/RTLpathSE_impl_junk.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index e46dddba..e82cca10 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -543,7 +543,7 @@ Definition hsilocal_simu_check hst1 hst2 : ?? unit := Definition revmap_check_single (dm: PTree.t node) (n tn: node) : ?? unit := DO res <~ some_or_fail (dm ! tn) "revmap_check_single: no mapping for tn";; - phys_check n res "revmap_check_single: n and res are physically different". + struct_check n res "revmap_check_single: n and res are physically different". Definition hsilocal_frame_simu_check frame hst1 hst2 : ?? unit := phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_frame_simu_check: hsi_smem sets aren't equiv";; -- cgit From 3c8429b030efcaba205b8d0faf7fdd1f174ea1e7 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 24 Jul 2020 14:04:27 +0200 Subject: Temporary prepass flags in test/regression --- test/regression/Makefile | 2 ++ 1 file changed, 2 insertions(+) diff --git a/test/regression/Makefile b/test/regression/Makefile index 744a2c03..87827282 100644 --- a/test/regression/Makefile +++ b/test/regression/Makefile @@ -1,6 +1,8 @@ include ../../Makefile.config CCOMP=../../ccomp +# TODO - temporary +CCOMPOPTS:=$(CCOMPOPTS) -fall-loads-nontrap -fduplicate 2 -fprepass CCOMPFLAGS=$(CCOMPOPTS) -stdlib ../../runtime \ -dparse -dc -dclight -dasm -fall \ -DARCH_$(ARCH) -DMODEL_$(MODEL) -- cgit From 9a8fcdd29d763d2b5650166ddd5b359a8bfed373 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 24 Jul 2020 14:33:57 +0200 Subject: trace quand le simulateur est appele --- kvx/lib/RTLpathSE_impl_junk.v | 5 +++-- test/c/Makefile | 2 ++ test/regression/Makefile | 2 +- 3 files changed, 6 insertions(+), 3 deletions(-) diff --git a/kvx/lib/RTLpathSE_impl_junk.v b/kvx/lib/RTLpathSE_impl_junk.v index e82cca10..1b4efad7 100644 --- a/kvx/lib/RTLpathSE_impl_junk.v +++ b/kvx/lib/RTLpathSE_impl_junk.v @@ -678,8 +678,9 @@ Qed. Global Opaque simu_check_rec. Global Hint Resolve simu_check_rec_correct: wlp. -Definition imp_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? unit := - simu_check_rec dm f tf (PTree.elements dm). +Definition imp_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? unit := + simu_check_rec dm f tf (PTree.elements dm);; + println("simu_check OK!"). Local Hint Resolve PTree.elements_correct: core. Lemma imp_simu_check_correct dm f tf: diff --git a/test/c/Makefile b/test/c/Makefile index 726631d2..a728d182 100644 --- a/test/c/Makefile +++ b/test/c/Makefile @@ -1,6 +1,8 @@ include ../../Makefile.config CCOMP=../../ccomp +# TODO - temporary +# CCOMPOPTS:=$(CCOMPOPTS) -fall-loads-nontrap -fduplicate 2 -fprepass CCOMPFLAGS=$(CCOMPOPTS) -stdlib ../../runtime -dc -dclight -dasm CFLAGS+=-O2 -Wall diff --git a/test/regression/Makefile b/test/regression/Makefile index 87827282..8b2f4021 100644 --- a/test/regression/Makefile +++ b/test/regression/Makefile @@ -2,7 +2,7 @@ include ../../Makefile.config CCOMP=../../ccomp # TODO - temporary -CCOMPOPTS:=$(CCOMPOPTS) -fall-loads-nontrap -fduplicate 2 -fprepass +# CCOMPOPTS:=$(CCOMPOPTS) -fall-loads-nontrap -fduplicate 2 -fprepass CCOMPFLAGS=$(CCOMPOPTS) -stdlib ../../runtime \ -dparse -dc -dclight -dasm -fall \ -DARCH_$(ARCH) -DMODEL_$(MODEL) -- cgit From 1e93ef5453d4f9a6c3108b1d43b3b9f3f4fb9d13 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 24 Jul 2020 15:01:39 +0200 Subject: Coq 8.11.2 --- test/monniaux/picosat-965/small.dat | 2 ++ 1 file changed, 2 insertions(+) create mode 100644 test/monniaux/picosat-965/small.dat diff --git a/test/monniaux/picosat-965/small.dat b/test/monniaux/picosat-965/small.dat new file mode 100644 index 00000000..ff3d7081 --- /dev/null +++ b/test/monniaux/picosat-965/small.dat @@ -0,0 +1,2 @@ +p cnf 2 1 +1 2 0 -- cgit From d073f1901e1ddf012bc2ac607e230d8156eeefe1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 24 Jul 2020 17:36:56 +0200 Subject: smaller example --- test/monniaux/picosat-965/small.dat | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/test/monniaux/picosat-965/small.dat b/test/monniaux/picosat-965/small.dat index ff3d7081..accb9054 100644 --- a/test/monniaux/picosat-965/small.dat +++ b/test/monniaux/picosat-965/small.dat @@ -1,2 +1,2 @@ -p cnf 2 1 -1 2 0 +p cnf 1 1 +1 0 -- cgit From 18988711111bd57900567b032aead1c17997cb7f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 24 Jul 2020 17:40:29 +0200 Subject: even smaller --- test/monniaux/picosat-965/tiny.dat | 2 ++ 1 file changed, 2 insertions(+) create mode 100644 test/monniaux/picosat-965/tiny.dat diff --git a/test/monniaux/picosat-965/tiny.dat b/test/monniaux/picosat-965/tiny.dat new file mode 100644 index 00000000..030df0e5 --- /dev/null +++ b/test/monniaux/picosat-965/tiny.dat @@ -0,0 +1,2 @@ +p cnf 1 1 +0 -- cgit From cffd805ea0c06710605565fae3d0e885a7943f74 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 24 Jul 2020 17:41:31 +0200 Subject: also crashes with 0 variables --- test/monniaux/picosat-965/tiny.dat | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/picosat-965/tiny.dat b/test/monniaux/picosat-965/tiny.dat index 030df0e5..1d89b303 100644 --- a/test/monniaux/picosat-965/tiny.dat +++ b/test/monniaux/picosat-965/tiny.dat @@ -1,2 +1,2 @@ -p cnf 1 1 +p cnf 0 1 0 -- cgit From 9464136be3904db1b227327cec0d8670b1fc164f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 24 Jul 2020 17:53:38 +0200 Subject: picosat in one file for testing purposes --- test/monniaux/picosat-965/onefile/picosat.c | 9765 +++++++++++++++++++++++++++ 1 file changed, 9765 insertions(+) create mode 100644 test/monniaux/picosat-965/onefile/picosat.c diff --git a/test/monniaux/picosat-965/onefile/picosat.c b/test/monniaux/picosat-965/onefile/picosat.c new file mode 100644 index 00000000..b255f3ee --- /dev/null +++ b/test/monniaux/picosat-965/onefile/picosat.c @@ -0,0 +1,9765 @@ +#define NALARM 1 +#define NZIP 1 +#define NGETRUSAGE 1 +#define NDEBUG 1 + +#include "picosat.h" + +#include +#include +#include +#include +#include +#include +#include + +#define GUNZIP "gunzip -c %s" +#define BUNZIP2 "bzcat %s" +#define GZIP "gzip -c -f > %s" + +#ifndef NZIP +FILE * popen (const char *, const char*); +int pclose (FILE *); +#endif + +static PicoSAT * picosat; + +static int lineno; +static FILE *input; +static int inputid; +static FILE *output; +static int verbose; +static int sargc; +static char ** sargv; +static char buffer[100]; +static char *bhead = buffer; +static const char *eob = buffer + 80; +static FILE * incremental_rup_file; +static signed char * sol; + +extern void picosat_enter (PicoSAT *); +extern void picosat_leave (PicoSAT *); + +static int +next (void) +{ + int res = getc (input); + if (res == '\n') + lineno++; + + return res; +} + +static const char * +parse (PicoSAT * picosat, int force) +{ + int ch, sign, lit, vars, clauses; + + lineno = 1; + /* DM inputid = fileno (input); */ + +SKIP_COMMENTS: + ch = next (); + if (ch == 'c') + { + while ((ch = next ()) != EOF && ch != '\n') + ; + goto SKIP_COMMENTS; + } + + if (isspace (ch)) + goto SKIP_COMMENTS; + + if (ch != 'p') +INVALID_HEADER: + return "missing or invalid 'p cnf ' header"; + + if (!isspace (next ())) + goto INVALID_HEADER; + + while (isspace (ch = next ())) + ; + + if (ch != 'c' || next () != 'n' || next () != 'f' || !isspace (next ())) + goto INVALID_HEADER; + + while (isspace (ch = next ())) + ; + + if (!isdigit (ch)) + goto INVALID_HEADER; + + vars = ch - '0'; + while (isdigit (ch = next ())) + vars = 10 * vars + (ch - '0'); + + if (!isspace (ch)) + goto INVALID_HEADER; + + while (isspace (ch = next ())) + ; + + if (!isdigit (ch)) + goto INVALID_HEADER; + + clauses = ch - '0'; + while (isdigit (ch = next ())) + clauses = 10 * clauses + (ch - '0'); + + if (!isspace (ch) && ch != '\n' ) + goto INVALID_HEADER; + + if (verbose) + { + fprintf (output, "c parsed header 'p cnf %d %d'\n", vars, clauses); + fflush (output); + } + + picosat_adjust (picosat, vars); + + if (incremental_rup_file) + picosat_set_incremental_rup_file (picosat, incremental_rup_file, vars, clauses); + + lit = 0; +READ_LITERAL: + ch = next (); + + if (ch == 'c') + { + while ((ch = next ()) != EOF && ch != '\n') + ; + goto READ_LITERAL; + } + + if (ch == EOF) + { + if (lit) + return "trailing 0 missing"; + + if (clauses && !force) + return "clause missing"; + + return 0; + } + + if (isspace (ch)) + goto READ_LITERAL; + + sign = 1; + if (ch == '-') + { + sign = -1; + ch = next (); + } + + if (!isdigit (ch)) + return "expected number"; + + lit = ch - '0'; + while (isdigit (ch = next ())) + lit = 10 * lit + (ch - '0'); + + if (!clauses && !force) + return "too many clauses"; + + if (lit) + { + if (lit > vars && !force) + return "maximal variable index exceeded"; + + lit *= sign; + } + else + clauses--; + + picosat_add (picosat, lit); + + goto READ_LITERAL; +} + +static void +bflush (void) +{ + *bhead = 0; + fputs (buffer, output); + fputc ('\n', output); + bhead = buffer; +} + +static void +printi (int i) +{ + char *next; + int l; + +REENTER: + if (bhead == buffer) + *bhead++ = 'v'; + + l = sprintf (bhead, " %d", i); + next = bhead + l; + + if (next >= eob) + { + bflush (); + goto REENTER; + } + else + bhead = next; +} + +static void +printa (PicoSAT * picosat, int partial) +{ + int max_idx = picosat_variables (picosat), i, lit, val; + + assert (bhead == buffer); + + for (i = 1; i <= max_idx; i++) + { + if (partial) + { + val = picosat_deref_partial (picosat, i); + if (!val) + continue; + } + else + val = picosat_deref (picosat, i); + lit = (val > 0) ? i : -i; + printi (lit); + } + + printi (0); + if (bhead > buffer) + bflush (); +} + +static void +blocksol (PicoSAT * picosat) +{ + int max_idx = picosat_variables (picosat), i; + + if (!sol) + { + sol = malloc (max_idx + 1); + memset (sol, 0, max_idx + 1); + } + + for (i = 1; i <= max_idx; i++) + sol[i] = (picosat_deref (picosat, i) > 0) ? 1 : -1; + + for (i = 1; i <= max_idx; i++) + picosat_add (picosat, (sol[i] < 0) ? i : -i); + + picosat_add (picosat, 0); +} + +static int +has_suffix (const char *str, const char *suffix) +{ + const char *tmp = strstr (str, suffix); + if (!tmp) + return 0; + + return str + strlen (str) - strlen (suffix) == tmp; +} + +static void +write_core_variables (PicoSAT * picosat, FILE * file) +{ + int i, max_idx = picosat_variables (picosat), count = 0; + for (i = 1; i <= max_idx; i++) + if (picosat_corelit (picosat, i)) + { + fprintf (file, "%d\n", i); + count++; + } + + if (verbose) + fprintf (output, "c found and wrote %d core variables\n", count); +} + +static int +next_assumption (int start) +{ + char * arg, c; + int res; + res = start + 1; + while (res < sargc) + { + arg = sargv[res++]; + if (!strcmp (arg, "-a")) + { + assert (res < sargc); + break; + } + + if (arg[0] == '-') { + c = arg[1]; + if (c == 'l' || c == 'i' || c == 's' || c == 'o' || c == 't' || + c == 'T' || c == 'r' || c == 'R' || c == 'c' || c == 'V' || + c == 'U' || c == 'A') res++; + } + } + if (res >= sargc) res = 0; + return res; +} + +static void +write_failed_assumptions (PicoSAT * picosat, FILE * file) +{ + int i, lit, count = 0; +#ifndef NDEBUG + int max_idx = picosat_variables (picosat); +#endif + i = 0; + while ((i = next_assumption (i))) { + lit = atoi (sargv[i]); + if (!picosat_failed_assumption (picosat, lit)) continue; + fprintf (file, "%d\n", lit); + count++; + } + if (verbose) + fprintf (output, "c found and wrote %d failed assumptions\n", count); +#ifndef NDEBUG + for (i = 1; i <= max_idx; i++) + if (picosat_failed_assumption (picosat, i)) + count--; +#endif + assert (!count); +} + +static void +write_to_file (PicoSAT * picosat, + const char *name, + const char *type, + void (*writer) (PicoSAT *, FILE *)) +{ + int pclose_file, zipped = has_suffix (name, ".gz"); + FILE *file; + char *cmd; + + if (zipped) + { +#ifdef NZIP + file = NULL; +#else + cmd = malloc (strlen (GZIP) + strlen (name)); + sprintf (cmd, GZIP, name); + file = popen (cmd, "w"); + free (cmd); + pclose_file = 1; +#endif + } + else + { + file = fopen (name, "w"); + pclose_file = 0; + } + + if (file) + { + if (verbose) + fprintf (output, + "c\nc writing %s%s to '%s'\n", + zipped ? "gzipped " : "", type, name); + + writer (picosat, file); + +#ifndef NZIP + if (pclose_file) + pclose (file); + else +#endif + fclose (file); + } + else + fprintf (output, "*** picosat: can not write to '%s'\n", name); +} + +static int catched; + +static void (*sig_int_handler); +static void (*sig_segv_handler); +static void (*sig_abrt_handler); +static void (*sig_term_handler); +#ifndef NALLSIGNALS +static void (*sig_kill_handler); +static void (*sig_xcpu_handler); +static void (*sig_xfsz_handler); +#endif + +static void +resetsighandlers (void) +{ + (void) signal (SIGINT, sig_int_handler); + (void) signal (SIGSEGV, sig_segv_handler); + (void) signal (SIGABRT, sig_abrt_handler); + (void) signal (SIGTERM, sig_term_handler); +#ifndef NALLSIGNALS + (void) signal (SIGKILL, sig_kill_handler); + (void) signal (SIGXCPU, sig_xcpu_handler); + (void) signal (SIGXFSZ, sig_xfsz_handler); +#endif +} + +static int time_limit_in_seconds; +static void (*sig_alarm_handler); +static int ought_to_be_interrupted, interrupt_notified; + +static void +alarm_triggered (int sig) +{ + (void) sig; + assert (sig == SIGALRM); + assert (time_limit_in_seconds); + assert (!ought_to_be_interrupted); + ought_to_be_interrupted = 1; + assert (!interrupt_notified); +} + +static int +interrupt_call_back (void * dummy) +{ + (void) dummy; + if (!ought_to_be_interrupted) + return 0; + if (!interrupt_notified) + { + if (verbose) + { + picosat_message (picosat, 1, ""); + picosat_message (picosat, 1, + "*** TIME LIMIT OF %d SECONDS REACHED ***", + time_limit_in_seconds); + picosat_message (picosat, 1, ""); + } + interrupt_notified = 1; + } + return 1; +} + +static void +setalarm () +{ +#ifndef NALARM + assert (time_limit_in_seconds > 0); + sig_alarm_handler = signal (SIGALRM, alarm_triggered); + alarm (time_limit_in_seconds); + assert (picosat); + picosat_set_interrupt (picosat, 0, interrupt_call_back); +#endif +} + +static void +resetalarm () +{ + assert (time_limit_in_seconds > 0); + (void) signal (SIGALRM, sig_term_handler); +} + +static void +message (int sig) +{ + picosat_message (picosat, 1, ""); + picosat_message (picosat, 1, "*** CAUGHT SIGNAL %d ***", sig); + picosat_message (picosat, 1, ""); +} + +static void +catch (int sig) +{ + if (!catched) + { + message (sig); + catched = 1; + picosat_stats (picosat); + message (sig); + } + + resetsighandlers (); + raise (sig); +} + +static void +setsighandlers (void) +{ + sig_int_handler = signal (SIGINT, catch); + sig_segv_handler = signal (SIGSEGV, catch); + sig_abrt_handler = signal (SIGABRT, catch); + sig_term_handler = signal (SIGTERM, catch); +#ifndef NALLSIGNALS + sig_kill_handler = signal (SIGKILL, catch); + sig_xcpu_handler = signal (SIGXCPU, catch); + sig_xfsz_handler = signal (SIGXFSZ, catch); +#endif +} + +#define USAGE \ +"usage: picosat [