aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling
diff options
context:
space:
mode:
authorDavid Monniaux <David.Monniaux@univ-grenoble-alpes.fr>2021-09-01 19:58:50 +0200
committerDavid Monniaux <David.Monniaux@univ-grenoble-alpes.fr>2021-09-01 19:58:50 +0200
commitd0819d14a514c81cd9437de467d4880965bc3a7f (patch)
tree1ac783da1a94f65073b69886fe288262ad1708c0 /scheduling
parentd1fe9e79ad19feff22f9e319dfafc36a534d9479 (diff)
parentddc17a17408541efa8b23afa3e6ccad1e6ce0b6e (diff)
downloadcompcert-kvx-d0819d14a514c81cd9437de467d4880965bc3a7f.tar.gz
compcert-kvx-d0819d14a514c81cd9437de467d4880965bc3a7f.zip
Merge branch 'kvx-work' of gricad-gitlab.univ-grenoble-alpes.fr:sixcy/CompCert into kvx-work
Diffstat (limited to 'scheduling')
-rw-r--r--scheduling/BTL.v758
-rw-r--r--scheduling/BTLRenumber.ml112
-rw-r--r--scheduling/BTLScheduleraux.ml337
-rw-r--r--scheduling/BTL_Livecheck.v690
-rw-r--r--scheduling/BTL_SEimpl.v1508
-rw-r--r--scheduling/BTL_SEsimuref.v852
-rw-r--r--scheduling/BTL_SEtheory.v1336
-rw-r--r--scheduling/BTL_Scheduler.v176
-rw-r--r--scheduling/BTL_Schedulerproof.v434
-rw-r--r--scheduling/BTLcommonaux.ml87
-rw-r--r--scheduling/BTLmatchRTL.v606
-rw-r--r--scheduling/BTLroadmap.md454
-rw-r--r--scheduling/BTLtoRTL.v26
-rw-r--r--scheduling/BTLtoRTLaux.ml88
-rw-r--r--scheduling/BTLtoRTLproof.v405
-rw-r--r--scheduling/BTLtypes.ml15
-rw-r--r--scheduling/InstructionScheduler.mli2
-rw-r--r--scheduling/PrintBTL.ml118
-rw-r--r--scheduling/RTLpath.v1066
-rw-r--r--scheduling/RTLpathCommon.ml14
-rw-r--r--scheduling/RTLpathLivegen.v325
-rw-r--r--scheduling/RTLpathLivegenaux.ml290
-rw-r--r--scheduling/RTLpathLivegenproof.v760
-rw-r--r--scheduling/RTLpathSE_impl.v1664
-rw-r--r--scheduling/RTLpathSE_simu_specs.v937
-rw-r--r--scheduling/RTLpathSE_theory.v1876
-rw-r--r--scheduling/RTLpathScheduler.v329
-rw-r--r--scheduling/RTLpathScheduleraux.ml497
-rw-r--r--scheduling/RTLpathSchedulerproof.v509
-rw-r--r--scheduling/RTLpathWFcheck.v187
-rw-r--r--scheduling/RTLpathproof.v50
-rw-r--r--scheduling/RTLtoBTL.v27
-rw-r--r--scheduling/RTLtoBTLaux.ml119
-rw-r--r--scheduling/RTLtoBTLproof.v758
34 files changed, 8907 insertions, 8505 deletions
diff --git a/scheduling/BTL.v b/scheduling/BTL.v
new file mode 100644
index 00000000..f832085c
--- /dev/null
+++ b/scheduling/BTL.v
@@ -0,0 +1,758 @@
+(** The BTL intermediate language: abstract syntax and semantics.
+
+ BTL stands for "Block Transfer Language".
+
+ Informally, a block is a piece of "loop-free" code, with a single entry-point,
+ hence, such that transformation preserving locally the semantics of each block,
+ preserve also globally the semantics of the function.
+
+ a BTL function is a CFG where each node is such a block, represented by structured code.
+
+ BTL gives a structured view of RTL code.
+ It is dedicated to optimizations validated by "symbolic simulation" over blocks.
+
+
+*)
+
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad.
+Require Import Lia.
+
+Import ListNotations.
+
+(** * Abstract syntax *)
+
+Definition exit := node. (* we may generalize this with register renamings at exit,
+ like in "phi" nodes of SSA-form *)
+
+(* inst_info is a ghost record to provide instruction information through oracles *)
+Parameter inst_info: Set.
+Extract Constant inst_info => "BTLtypes.inst_info".
+
+(* block_info is a ghost record to provide block information through oracles *)
+Parameter block_info: Set.
+Extract Constant block_info => "BTLtypes.block_info".
+
+(** final instructions (that stops block execution) *)
+Inductive final: Type :=
+ | Bgoto (succ:exit) (** No operation -- just branch to [succ]. *)
+ | Breturn (res: option reg)
+ (** terminates the execution of the current function. It returns the value of the given
+ register, or [Vundef] if none is given. *)
+ | Bcall (sig: signature) (fn: reg + ident) (args:list reg) (dest:reg) (succ:exit)
+ (** invokes the function determined by [fn] (either a function pointer found in a register or a
+ function name), giving it the values of registers [args] as arguments.
+ It stores the return value in [dest] and branches to [succ]. *)
+ | Btailcall (sig:signature) (fn: reg + ident) (args: list reg)
+ (** performs a function invocation in tail-call position
+ (the current function terminates after the call, returning the result of the callee)
+ *)
+ | Bbuiltin (ef:external_function) (args:list (builtin_arg reg)) (dest:builtin_res reg) (succ:exit)
+ (** calls the built-in function identified by [ef], giving it the values of [args] as arguments.
+ It stores the return value in [dest] and branches to [succ]. *)
+ | Bjumptable (arg:reg) (tbl:list exit)
+ (** [Bjumptable arg tbl] transitions to the node that is the [n]-th
+ element of the list [tbl], where [n] is the unsigned integer value of register [arg]. *)
+ .
+
+(* instruction block *)
+Inductive iblock: Type :=
+(* final instructions that stops block execution *)
+ | BF (fi: final) (iinfo: inst_info)
+(* basic instructions that continues block execution, except when aborting *)
+ | Bnop (oiinfo: option inst_info) (* nop instruction *)
+ | Bop (op:operation) (args:list reg) (dest:reg) (iinfo: inst_info)
+ (** performs the arithmetic operation [op] over the values of registers [args], stores the result in [dest] *)
+ | Bload (trap:trapping_mode) (chunk:memory_chunk) (addr:addressing) (args:list reg) (dest:reg) (iinfo: inst_info)
+ (** loads a [chunk] quantity from the address determined by the addressing mode [addr]
+ and the values of the [args] registers, stores the quantity just read into [dest].
+ If trap=NOTRAP, then failures lead to a default value written to [dest]. *)
+ | Bstore (chunk:memory_chunk) (addr:addressing) (args:list reg) (src:reg) (iinfo: inst_info)
+ (** stores the value of register [src] in the [chunk] quantity at the
+ the address determined by the addressing mode [addr] and the
+ values of the [args] registers. *)
+(* composed instructions *)
+ | Bseq (b1 b2: iblock)
+ (** starts by running [b1] and stops here if execution of [b1] has reached a final instruction or aborted
+ or continue with [b2] otherwise *)
+ | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (iinfo: inst_info)
+ (** evaluates the boolean condition [cond] over the values of registers [args].
+ If the condition is true, it continues on [ifso].
+ If the condition is false, it continues on [ifnot].
+ [info] is a ghost field there to provide information relative to branch prediction. *)
+ .
+Coercion BF: final >-> Funclass.
+
+
+(** NB: - a RTL [(Inop pc)] ending a branch of block is encoded by [(Bseq Bnop (Bgoto pc))].
+ - a RTL [(Inop pc)] in the middle of a branch is simply encoded by [Bnop].
+ - the same trick appears for all "basic" instructions and [Icond].
+*)
+
+Record iblock_info := {
+ entry: iblock;
+ input_regs: Regset.t; (* extra liveness information for BTL functional semantics *)
+ binfo: block_info (* Ghost field used in oracles *)
+}.
+
+Definition code: Type := PTree.t iblock_info.
+
+Record function: Type := mkfunction {
+ fn_sig: signature;
+ fn_params: list reg;
+ fn_stacksize: Z;
+ fn_code: code;
+ fn_entrypoint: node
+}.
+
+(** A function description comprises a control-flow graph (CFG) [fn_code]
+ (a partial finite mapping from nodes to instructions). As in Cminor,
+ [fn_sig] is the function signature and [fn_stacksize] the number of bytes
+ for its stack-allocated activation record. [fn_params] is the list
+ of registers that are bound to the values of arguments at call time.
+ [fn_entrypoint] is the node of the first instruction of the function
+ in the CFG. *)
+
+Definition fundef := AST.fundef function.
+
+Definition program := AST.program fundef unit.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => fn_sig f
+ | External ef => ef_sig ef
+ end.
+
+(** * Operational semantics *)
+
+Definition genv := Genv.t fundef unit.
+
+(** The dynamic semantics of BTL is similar to RTL,
+ except that the step of one instruction is generalized into the run of one [iblock].
+*)
+
+Inductive stackframe : Type :=
+ | Stackframe:
+ forall (res: reg) (**r where to store the result *)
+ (f: function) (**r calling function *)
+ (sp: val) (**r stack pointer in calling function *)
+ (succ: exit) (**r program point in calling function *)
+ (rs: regset), (**r register state in calling function *)
+ stackframe.
+
+Inductive state : Type :=
+ | State:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: function) (**r current function *)
+ (sp: val) (**r stack pointer *)
+ (pc: node) (**r current program point in [c] *)
+ (rs: regset) (**r register state *)
+ (m: mem), (**r memory state *)
+ state
+ | Callstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: fundef) (**r function to call *)
+ (args: list val) (**r arguments to the call *)
+ (m: mem), (**r memory state *)
+ state
+ | Returnstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (v: val) (**r return value for the call *)
+ (m: mem), (**r memory state *)
+ state.
+
+(** outcome of a block execution *)
+Record outcome := out {
+ _rs: regset;
+ _m: mem;
+ _fin: option final
+}.
+
+(* follows RTL semantics to set the result of builtin *)
+Definition reg_builtin_res (res: builtin_res reg): option reg :=
+ match res with
+ | BR r => Some r
+ | _ => None
+ end.
+
+Section RELSEM.
+
+(** [step] (and in particular [final_step]) is parametrized by function [tr_exit] to transfer registers on each iblock exit.
+
+ In particular, [tr_exit f lpc or rs] computes from [rs] the [rs'] on which the execution from iblock [pc] in [lpc] in [f] may start.
+
+ Here, [or] is an optional register that will be assigned after exiting the iblock, but before entering in [pc]: e.g. the register set by a function call
+ before entering in return address.
+
+ See [tr_inputs] implementation below.
+
+*)
+
+Variable tr_exit: function -> list exit -> option reg -> regset -> regset.
+
+Variable ge: genv.
+
+Definition find_function (ros: reg + ident) (rs: regset) : option fundef :=
+ match ros with
+ | inl r => Genv.find_funct ge rs#r
+ | inr symb =>
+ match Genv.find_symbol ge symb with
+ | None => None
+ | Some b => Genv.find_funct_ptr ge b
+ end
+ end.
+
+Local Open Scope option_monad_scope.
+
+(* TODO: a new (hopefully simpler) scheme to support "NOTRAP" wrt current scheme of RTL *)
+
+Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop :=
+ | has_loaded_normal a trap
+ (EVAL: eval_addressing ge sp addr rs##args = Some a)
+ (LOAD: Mem.loadv chunk m a = Some v)
+ : has_loaded sp rs m chunk addr args v trap
+ | has_loaded_default
+ (LOAD: forall a, eval_addressing ge sp addr rs##args = Some a -> Mem.loadv chunk m a = None)
+ (DEFAULT: v = Vundef)
+ : has_loaded sp rs m chunk addr args v NOTRAP
+ .
+Local Hint Constructors has_loaded: core.
+
+(* TODO: move this scheme in "Memory" module if this scheme is useful ! *)
+
+(** internal big-step execution of one iblock *)
+Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option final -> Prop :=
+ | exec_final rs m fin iinfo: iblock_istep sp rs m (BF fin iinfo) rs m (Some fin)
+| exec_nop rs m oiinfo: iblock_istep sp rs m (Bnop oiinfo) rs m None
+ | exec_op rs m op args res v iinfo
+ (EVAL: eval_operation ge sp op rs##args m = Some v)
+ : iblock_istep sp rs m (Bop op args res iinfo) (rs#res <- v) m None
+ | exec_load rs m trap chunk addr args dst v iinfo
+ (LOAD: has_loaded sp rs m chunk addr args v trap)
+ : iblock_istep sp rs m (Bload trap chunk addr args dst iinfo) (rs#dst <- v) m None
+ | exec_store rs m chunk addr args src a m' iinfo
+ (EVAL: eval_addressing ge sp addr rs##args = Some a)
+ (STORE: Mem.storev chunk m a rs#src = Some m')
+ : iblock_istep sp rs m (Bstore chunk addr args src iinfo) rs m' None
+ | exec_seq_stop rs m b1 b2 rs' m' fin
+ (EXEC: iblock_istep sp rs m b1 rs' m' (Some fin))
+ : iblock_istep sp rs m (Bseq b1 b2) rs' m' (Some fin)
+ | exec_seq_continue rs m b1 b2 rs1 m1 rs' m' ofin
+ (EXEC1: iblock_istep sp rs m b1 rs1 m1 None)
+ (EXEC2: iblock_istep sp rs1 m1 b2 rs' m' ofin)
+ : iblock_istep sp rs m (Bseq b1 b2) rs' m' ofin
+ | exec_cond rs m cond args ifso ifnot b rs' m' ofin iinfo
+ (EVAL: eval_condition cond rs##args m = Some b)
+ (EXEC: iblock_istep sp rs m (if b then ifso else ifnot) rs' m' ofin)
+ : iblock_istep sp rs m (Bcond cond args ifso ifnot iinfo) rs' m' ofin
+ .
+Local Hint Constructors iblock_istep: core.
+
+(** A functional variant of [iblock_istep_run] of [iblock_istep].
+Lemma [iblock_istep_run_equiv] below provides a proof that "relation" [iblock_istep] is a "partial function".
+*)
+Fixpoint iblock_istep_run sp ib rs m: option outcome :=
+ match ib with
+ | BF fin _ =>
+ Some {| _rs := rs; _m := m; _fin := Some fin |}
+ (* basic instructions *)
+ | Bnop _ =>
+ Some {| _rs := rs; _m:= m; _fin := None |}
+ | Bop op args res _ =>
+ SOME v <- eval_operation ge sp op rs##args m IN
+ Some {| _rs := rs#res <- v; _m:= m; _fin := None |}
+ | Bload TRAP chunk addr args dst _ =>
+ SOME a <- eval_addressing ge sp addr rs##args IN
+ SOME v <- Mem.loadv chunk m a IN
+ Some {| _rs := rs#dst <- v; _m:= m; _fin := None |}
+ | Bload NOTRAP chunk addr args dst _ =>
+ match eval_addressing ge sp addr rs##args with
+ | Some a =>
+ match Mem.loadv chunk m a with
+ | Some v => Some {| _rs := rs#dst <- v; _m:= m; _fin := None |}
+ | None =>
+ Some {| _rs := rs#dst <- Vundef; _m:= m; _fin := None |}
+ end
+ | None =>
+ Some {| _rs := rs#dst <- Vundef; _m:= m; _fin := None |}
+ end
+ | Bstore chunk addr args src _ =>
+ SOME a <- eval_addressing ge sp addr rs##args IN
+ SOME m' <- Mem.storev chunk m a rs#src IN
+ Some {| _rs := rs; _m:= m'; _fin := None |}
+ (* composed instructions *)
+ | Bseq b1 b2 =>
+ SOME out1 <- iblock_istep_run sp b1 rs m IN
+ match out1.(_fin) with
+ | None => iblock_istep_run sp b2 out1.(_rs) out1.(_m)
+ | _ => Some out1 (* stop execution on the 1st final instruction *)
+ end
+ | Bcond cond args ifso ifnot _ =>
+ SOME b <- eval_condition cond rs##args m IN
+ iblock_istep_run sp (if b then ifso else ifnot) rs m
+ end.
+
+Lemma iblock_istep_run_equiv_load sp ib v rs rs' m trap chunk addr args dst iinfo ofin:
+ ib = (Bload trap chunk addr args dst iinfo) ->
+ rs' = rs # dst <- v ->
+ has_loaded sp rs m chunk addr args v trap ->
+ iblock_istep sp rs m ib rs' m ofin <->
+ iblock_istep_run sp ib rs m = Some {| _rs := rs'; _m := m; _fin := ofin |}.
+Proof.
+ split; subst; inv H1; simpl in *; intros.
+ - destruct trap; inv H; simpl in *; repeat autodestruct.
+ - inv H; autodestruct; rewrite LOAD; auto.
+ - destruct trap; inv H; simpl in *;
+ rewrite EVAL, LOAD in H1; inv H1; repeat econstructor; eauto.
+ - generalize H; autodestruct.
+ + rewrite LOAD in *; inv H; auto; constructor.
+ eapply has_loaded_default; eauto; try_simplify_someHyps.
+ + inv H; constructor. eapply has_loaded_default; eauto; try_simplify_someHyps.
+Qed.
+Local Hint Resolve iblock_istep_run_equiv_load: core.
+
+Lemma iblock_istep_run_equiv sp rs m ib rs' m' ofin:
+ iblock_istep sp rs m ib rs' m' ofin <-> iblock_istep_run sp ib rs m = Some {| _rs := rs'; _m:= m'; _fin := ofin |}.
+Proof.
+ constructor.
+ - induction 1; try (eapply iblock_istep_run_equiv_load; eauto; fail);
+ simpl; try autodestruct; try_simplify_someHyps.
+ - generalize rs m rs' m' ofin; clear rs m rs' m' ofin.
+ induction ib; simpl; repeat (try autodestruct; try_simplify_someHyps).
+ 1,2: constructor; eapply has_loaded_default; try_simplify_someHyps.
+ destruct o; try_simplify_someHyps; subst; eauto.
+Qed.
+
+Local Open Scope list_scope.
+
+Inductive final_step stack f sp rs m: final -> trace -> state -> Prop :=
+ | exec_Bgoto pc:
+ final_step stack f sp rs m (Bgoto pc) E0
+ (State stack f sp pc (tr_exit f [pc] None rs) m)
+ | exec_Breturn or stk m':
+ sp = (Vptr stk Ptrofs.zero) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ final_step stack f sp rs m (Breturn or)
+ E0 (Returnstate stack (regmap_optget or Vundef rs) m')
+ | exec_Bcall sig ros args res pc' fd:
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ final_step stack f sp rs m (Bcall sig ros args res pc')
+ E0 (Callstate (Stackframe res f sp pc' (tr_exit f [pc'] (Some res) rs) :: stack) fd rs##args m)
+ | exec_Btailcall sig ros args stk m' fd:
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ sp = (Vptr stk Ptrofs.zero) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ final_step stack f sp rs m (Btailcall sig ros args)
+ E0 (Callstate stack fd rs##args m')
+ | exec_Bbuiltin ef args res pc' vargs t vres m':
+ eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ final_step stack f sp rs m (Bbuiltin ef args res pc')
+ t (State stack f sp pc' (regmap_setres res vres (tr_exit f [pc'] (reg_builtin_res res) rs)) m')
+ | exec_Bjumptable arg tbl n pc':
+ rs#arg = Vint n ->
+ list_nth_z tbl (Int.unsigned n) = Some pc' ->
+ final_step stack f sp rs m (Bjumptable arg tbl)
+ E0 (State stack f sp pc' (tr_exit f tbl None rs) m)
+.
+
+(** big-step execution of one iblock *)
+Definition iblock_step stack f sp rs m ib t s: Prop :=
+ exists rs' m' fin, iblock_istep sp rs m ib rs' m' (Some fin) /\ final_step stack f sp rs' m' fin t s.
+
+(** The transitions are presented as an inductive predicate
+ [step ge st1 t st2], where [ge] is the global environment,
+ [st1] the initial state, [st2] the final state, and [t] the trace
+ of system calls performed during this transition. *)
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_iblock stack ib f sp pc rs m t s
+ (PC: (fn_code f)!pc = Some ib)
+ (STEP: iblock_step stack f sp rs m ib.(entry) t s)
+ :step (State stack f sp pc rs m) t s
+ | exec_function_internal stack f args m m' stk
+ (ALLOC: Mem.alloc m 0 f.(fn_stacksize) = (m', stk))
+ :step (Callstate stack (Internal f) args m)
+ E0 (State stack
+ f
+ (Vptr stk Ptrofs.zero)
+ f.(fn_entrypoint)
+ (init_regs args f.(fn_params))
+ m')
+ | exec_function_external stack ef args res t m m'
+ (EXTCALL: external_call ef ge args m t res m')
+ :step (Callstate stack (External ef) args m)
+ t (Returnstate stack res m')
+ | exec_return stack res f sp pc rs vres m
+ :step (Returnstate (Stackframe res f sp pc rs :: stack) vres m)
+ E0 (State stack f sp pc (rs#res <- vres) m)
+.
+
+End RELSEM.
+
+(** Execution of whole programs are described as sequences of transitions
+ from an initial state to a final state. An initial state is a [Callstate]
+ corresponding to the invocation of the ``main'' function of the program
+ without arguments and with an empty call stack. *)
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall b f m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ Genv.find_symbol ge p.(prog_main) = Some b ->
+ Genv.find_funct_ptr ge b = Some f ->
+ funsig f = signature_main ->
+ initial_state p (Callstate nil f nil m0).
+
+(** A final state is a [Returnstate] with an empty call stack. *)
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall r m,
+ final_state (Returnstate nil (Vint r) m) r.
+
+(** The full "functional" small-step semantics for a BTL program.
+ at each exit, we only transfer register in "input_regs" (i.e. "alive" registers).
+*)
+Definition transfer_regs (inputs: list reg) (rs: regset): regset :=
+ init_regs (rs##inputs) inputs.
+
+Lemma transfer_regs_inputs inputs rs r:
+ List.In r inputs -> (transfer_regs inputs rs)#r = rs#r.
+Proof.
+ unfold transfer_regs; induction inputs; simpl; intuition subst.
+ - rewrite Regmap.gss; auto.
+ - destruct (Pos.eq_dec a r).
+ + subst; rewrite Regmap.gss; auto.
+ + rewrite Regmap.gso; auto.
+Qed.
+
+Lemma transfer_regs_dead inputs rs r:
+ ~List.In r inputs -> (transfer_regs inputs rs)#r = Vundef.
+Proof.
+ unfold transfer_regs; induction inputs; simpl; intuition subst.
+ - rewrite Regmap.gi; auto.
+ - rewrite Regmap.gso; auto.
+Qed.
+
+Fixpoint union_inputs (f:function) (tbl: list exit): Regset.t :=
+ match tbl with
+ | nil => Regset.empty
+ | pc::l => let rs:= union_inputs f l in
+ match f.(fn_code)!pc with
+ | None => rs
+ | Some ib => Regset.union rs ib.(input_regs)
+ end
+ end.
+
+(* TODO: lemma not yet used
+
+Lemma union_inputs_In f tbl r:
+ Regset.In r (union_inputs f tbl) -> (exists pc, List.In pc tbl /\ exists ib, f.(fn_code)!pc = Some ib /\ Regset.In r ib.(input_regs)).
+Proof.
+ induction tbl as [|pc l]; simpl; intros.
+ - exploit Regset.empty_1; eauto. intuition.
+ - destruct ((fn_code f) ! pc) eqn:ATpc.
+ + exploit Regset.union_1; eauto.
+ destruct 1 as [X1|X1].
+ * destruct IHl as (pc' & H1 & H2); eauto.
+ * eexists; intuition eauto.
+ + destruct IHl as (pc' & H1 & H2); eauto.
+Qed.
+
+Lemma union_inputs_notIn f tbl r:
+ (forall pc ib, List.In pc tbl -> f.(fn_code)!pc = Some ib -> ~Regset.In r ib.(input_regs))
+ -> ~Regset.In r (union_inputs f tbl).
+Proof.
+ induction tbl as [|pc l]; simpl; intuition eauto.
+ - exploit Regset.empty_1; eauto.
+ - destruct ((fn_code f) ! pc) eqn:ATpc; intuition eauto.
+ exploit Regset.union_1; intuition eauto.
+Qed.
+*)
+
+(* This function computes the union of the inputs associated to the exits
+ minus the optional register in [or] (typically the result of a call or a builtin).
+ i.e. similar to [pre_output_regs] in RTLpath.
+*)
+Definition pre_inputs (f:function) (tbl: list exit) (or: option reg): Regset.t :=
+ let rs := union_inputs f tbl in
+ match or with
+ | Some r => Regset.remove r rs
+ | None => rs
+ end
+ .
+
+(* TODO: lemma pre_inputs_In + pre_inputs_notIn ? *)
+
+Definition tr_inputs (f:function) (tbl: list exit) (or:option reg): regset -> regset
+ := transfer_regs (Regset.elements (pre_inputs f tbl or)).
+
+
+(* TODO: move this elsewhere *)
+Lemma SetoidList_InA_eq_equiv A (l: list A): forall x,
+ SetoidList.InA (fun x y => x = y) x l <-> List.In x l.
+Proof.
+ intros x; split.
+ - induction 1; simpl; eauto.
+ - induction l; simpl; intuition.
+Qed.
+
+Lemma tr_pre_inputs f tbl or rs r:
+ Regset.In r (pre_inputs f tbl or) -> (tr_inputs f tbl or rs)#r = rs#r.
+Proof.
+ intros; eapply transfer_regs_inputs.
+ rewrite <- SetoidList_InA_eq_equiv.
+ eapply Regset.elements_1; eauto.
+Qed.
+
+Lemma tr_inputs_dead f tbl or rs r:
+ ~(Regset.In r (pre_inputs f tbl or)) -> (tr_inputs f tbl or rs)#r = Vundef.
+Proof.
+ intros X; eapply transfer_regs_dead; intuition eauto.
+ eapply X. eapply Regset.elements_2.
+ rewrite -> SetoidList_InA_eq_equiv; eauto.
+Qed.
+
+Local Hint Resolve tr_pre_inputs Regset.mem_2 Regset.mem_1: core.
+
+Lemma tr_inputs_get f tbl or rs r:
+ (tr_inputs f tbl or rs)#r = if Regset.mem r (pre_inputs f tbl or) then rs#r else Vundef.
+Proof.
+ autodestruct; eauto.
+ intros; apply tr_inputs_dead; eauto.
+ intro X; exploit Regset.mem_1; eauto.
+ congruence.
+Qed.
+
+(* TODO: not yet used
+Lemma tr_inputs_ext f tbl or rs1 rs2:
+ (forall r, Regset.In r (pre_inputs f tbl or) -> rs1#r = rs2#r) ->
+ (forall r, (tr_inputs f tbl or rs1)#r = (tr_inputs f tbl or rs2)#r).
+Proof.
+ intros EQ r. rewrite !tr_inputs_get.
+ autodestruct; auto.
+Qed.
+*)
+
+Definition fsem (p: program) :=
+ Semantics (step tr_inputs) (initial_state p) final_state (Genv.globalenv p).
+
+Local Open Scope list_scope.
+
+Definition poly_tr {A} (tr: function -> list exit -> option reg -> A) f (i: final): A :=
+ match i with
+ | Bgoto pc => tr f [pc] None
+ | Bcall _ _ _ res pc => tr f [pc] (Some res)
+ | Btailcall _ _ args => tr f [] None
+ | Bbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res)
+ | Breturn _ => tr f [] None
+ | Bjumptable _ tbl => tr f tbl None
+ end.
+
+Definition tr_regs: function -> final -> regset -> regset :=
+ poly_tr tr_inputs.
+
+(* TODO: NOT USEFUL ?
+Definition liveout: function -> final -> Regset.t :=
+ poly_tr pre_inputs.
+
+Lemma tr_regs_liveout_equiv f fi : tr_regs f fi = transfer_regs (Regset.elements (liveout f fi)).
+Proof.
+ destruct fi; simpl; auto.
+Qed.
+
+Local Hint Resolve tr_inputs_get: core.
+
+Lemma tr_regs_get f fi rs r: (tr_regs f fi rs)#r = if Regset.mem r (liveout f fi) then rs#r else Vundef.
+Proof.
+ Local Opaque pre_inputs.
+ destruct fi; simpl; auto.
+Qed.
+*)
+
+(* * Comparing BTL semantics modulo extensionality of [regset].
+
+NB: This is at least used in BTL_SEtheory (and probably in the future BTL_SchedulerProof).
+
+*)
+Inductive equiv_stackframe: stackframe -> stackframe -> Prop :=
+ | equiv_stackframe_intro res f sp pc (rs1 rs2: regset)
+ (EQUIV: forall r, rs1#r = rs2#r)
+ :equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2)
+ .
+
+Inductive equiv_state: state -> state -> Prop :=
+ | State_equiv stk1 stk2 f sp pc rs1 m rs2
+ (STACKS: list_forall2 equiv_stackframe stk1 stk2)
+ (EQUIV: forall r, rs1#r = rs2#r)
+ :equiv_state (State stk1 f sp pc rs1 m) (State stk2 f sp pc rs2 m)
+ | Call_equiv stk1 stk2 f args m
+ (STACKS: list_forall2 equiv_stackframe stk1 stk2)
+ :equiv_state (Callstate stk1 f args m) (Callstate stk2 f args m)
+ | Return_equiv stk1 stk2 v m
+ (STACKS: list_forall2 equiv_stackframe stk1 stk2)
+ :equiv_state (Returnstate stk1 v m) (Returnstate stk2 v m)
+ .
+Local Hint Constructors equiv_stackframe equiv_state: core.
+
+Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf.
+Proof.
+ destruct stf; eauto.
+Qed.
+
+Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk.
+Proof.
+ Local Hint Resolve equiv_stackframe_refl: core.
+ induction stk; simpl; constructor; auto.
+Qed.
+Local Hint Resolve equiv_stack_refl: core.
+
+Lemma equiv_state_refl s: equiv_state s s.
+Proof.
+ 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.
+Local Hint Resolve equiv_stackframe_trans: core.
+
+Lemma equiv_stack_trans stk1 stk2:
+ list_forall2 equiv_stackframe stk1 stk2 ->
+ forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
+ list_forall2 equiv_stackframe stk1 stk3.
+Proof.
+ induction 1; intros stk3 EQ; inv EQ; econstructor; eauto.
+Qed.
+
+Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3.
+Proof.
+ Local Hint Resolve equiv_stack_trans: core.
+ destruct 1; intros EQ; inv EQ; econstructor; eauto.
+ intros; eapply eq_trans; eauto.
+Qed.
+
+Lemma regmap_setres_eq (rs rs': regset) res vres:
+ (forall r, rs # r = rs' # r) ->
+ forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r.
+Proof.
+ intros RSEQ r. destruct res; simpl; try congruence.
+ destruct (peq x r).
+ - subst. repeat (rewrite Regmap.gss). reflexivity.
+ - repeat (rewrite Regmap.gso); auto.
+Qed.
+
+(* * Comparing BTL semantics modulo [regs_lessdef].
+
+This extends the previous [equiv_*] stuff for [Val.lessdef]. Here, we need to also compare memories
+for Mem.extends, because Vundef values are propagated in memory through stores, allocation, etc.
+
+*)
+
+Inductive lessdef_stackframe: stackframe -> stackframe -> Prop :=
+ | lessdef_stackframe_intro res f sp pc rs1 rs2
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ :lessdef_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2)
+ .
+
+Inductive lessdef_state: state -> state -> Prop :=
+ | State_lessdef stk1 stk2 f sp pc rs1 m1 rs2 m2
+ (STACKS: list_forall2 lessdef_stackframe stk1 stk2)
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ (MEMS: Mem.extends m1 m2)
+ :lessdef_state (State stk1 f sp pc rs1 m1) (State stk2 f sp pc rs2 m2)
+ | Call_lessdef stk1 stk2 f args1 args2 m1 m2
+ (STACKS: list_forall2 lessdef_stackframe stk1 stk2)
+ (ARGS: Val.lessdef_list args1 args2)
+ (MEMS: Mem.extends m1 m2)
+ :lessdef_state (Callstate stk1 f args1 m1) (Callstate stk2 f args2 m2)
+ | Return_lessdef stk1 stk2 v1 v2 m1 m2
+ (STACKS: list_forall2 lessdef_stackframe stk1 stk2)
+ (REGS: Val.lessdef v1 v2)
+ (MEMS: Mem.extends m1 m2)
+ :lessdef_state (Returnstate stk1 v1 m1) (Returnstate stk2 v2 m2)
+ .
+Local Hint Constructors lessdef_stackframe lessdef_state: core.
+
+Lemma lessdef_stackframe_refl stf: lessdef_stackframe stf stf.
+Proof.
+ destruct stf; eauto.
+Qed.
+
+Local Hint Resolve lessdef_stackframe_refl: core.
+Lemma lessdef_stack_refl stk: list_forall2 lessdef_stackframe stk stk.
+Proof.
+ induction stk; simpl; constructor; auto.
+Qed.
+
+Local Hint Resolve lessdef_stack_refl: core.
+
+Lemma lessdef_list_refl args: Val.lessdef_list args args.
+Proof.
+ induction args; simpl; auto.
+Qed.
+
+Local Hint Resolve lessdef_list_refl Mem.extends_refl: core.
+
+Lemma lessdef_state_refl s: lessdef_state s s.
+Proof.
+ induction s; simpl; constructor; auto.
+Qed.
+
+Local Hint Resolve Val.lessdef_trans: core.
+Lemma lessdef_stackframe_trans stf1 stf2 stf3:
+ lessdef_stackframe stf1 stf2 -> lessdef_stackframe stf2 stf3 -> lessdef_stackframe stf1 stf3.
+Proof.
+ destruct 1. intros LD; inv LD; try econstructor; eauto.
+Qed.
+
+Local Hint Resolve lessdef_stackframe_trans: core.
+Lemma lessdef_stack_trans stk1 stk2:
+ list_forall2 lessdef_stackframe stk1 stk2 ->
+ forall stk3, list_forall2 lessdef_stackframe stk2 stk3 ->
+ list_forall2 lessdef_stackframe stk1 stk3.
+Proof.
+ induction 1; intros stk3 LD; inv LD; econstructor; eauto.
+Qed.
+
+Local Hint Resolve lessdef_stack_trans Mem.extends_extends_compose Val.lessdef_list_trans: core.
+Lemma lessdef_state_trans s1 s2 s3: lessdef_state s1 s2 -> lessdef_state s2 s3 -> lessdef_state s1 s3.
+Proof.
+ destruct 1; intros LD; inv LD; econstructor; eauto.
+Qed.
+
+Lemma init_regs_lessdef_preserv args1 args2
+ (ARGS : Val.lessdef_list args1 args2)
+ : forall rl r, Val.lessdef (init_regs args1 rl)#r (init_regs args2 rl)#r.
+Proof.
+ induction ARGS; simpl; auto.
+ intros rl r1; destruct rl; simpl in *; auto.
+ eapply set_reg_lessdef; eauto.
+ intros r2; eauto.
+Qed.
+
+Lemma find_function_lessdef ge ros rs1 rs2 fd
+ (FIND: find_function ge ros rs1 = Some fd)
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ : find_function ge ros rs2 = Some fd.
+Proof.
+ destruct ros; simpl in *; auto.
+ exploit Genv.find_funct_inv; eauto.
+ intros (b & EQ).
+ destruct (REGS r); try congruence.
+Qed.
+
+(** * Auxiliary general purpose functions on BTL *)
+
+Definition is_goto (ib: iblock): bool :=
+ match ib with
+ | BF (Bgoto _) _ => true
+ | _ => false
+ end.
+
diff --git a/scheduling/BTLRenumber.ml b/scheduling/BTLRenumber.ml
new file mode 100644
index 00000000..2f4f9203
--- /dev/null
+++ b/scheduling/BTLRenumber.ml
@@ -0,0 +1,112 @@
+open Maps
+open BTL
+open RTLcommonaux
+open BTLcommonaux
+open BTLtypes
+open DebugPrint
+open PrintBTL
+
+let recompute_inumbs btl entry =
+ let btl = reset_visited_ib (reset_visited_ibf btl) in
+ let last_used = ref 0 in
+ let ibf = get_some @@ PTree.get entry btl in
+ let ipos () =
+ last_used := !last_used + 1;
+ !last_used
+ in
+ let rec walk ib k =
+ (* heuristic: try to explore the lower numbered branch first *)
+ let walk_smallest_child s1 s2 ib1 ib2 =
+ if s1 < s2 then (
+ walk ib1 None;
+ walk ib2 None)
+ else (
+ walk ib2 None;
+ walk ib1 None)
+ in
+ if jump_visit ib then ()
+ else
+ match ib with
+ | BF (Bcall (_, _, _, _, s), iinfo) | BF (Bbuiltin (_, _, _, s), iinfo) ->
+ let ib' = (get_some @@ PTree.get s btl).entry in
+ walk ib' None;
+ iinfo.inumb <- ipos ()
+ | BF (Bgoto s, _) ->
+ let ib' = (get_some @@ PTree.get s btl).entry in
+ walk ib' None
+ | BF (Bjumptable (_, tbl), iinfo) ->
+ List.iter
+ (fun s ->
+ let ib' = (get_some @@ PTree.get s btl).entry in
+ walk ib' None)
+ tbl;
+ iinfo.inumb <- ipos ()
+ | BF (Btailcall (_, _, _), iinfo) | BF (Breturn _, iinfo) ->
+ iinfo.inumb <- ipos ()
+ | Bnop None ->
+ failwith
+ "recompute_inumbs: Bnop None can only be in the right child of \
+ Bcond"
+ | Bnop (Some iinfo)
+ | Bop (_, _, _, iinfo)
+ | Bload (_, _, _, _, _, iinfo)
+ | Bstore (_, _, _, _, iinfo) ->
+ let succ = get_some @@ k in
+ walk succ None;
+ iinfo.inumb <- ipos ()
+ | Bseq (ib1, ib2) -> walk ib1 (Some ib2)
+ | Bcond (_, _, BF (Bgoto s1, iinfoL), Bnop None, iinfoF) ->
+ iinfoL.visited <- true;
+ let ib1 = get_some @@ PTree.get s1 btl in
+ let ib2 = get_some @@ k in
+ walk_smallest_child (p2i s1) (get_inumb_or_next ib2) ib1.entry ib2;
+ iinfoF.inumb <- ipos ()
+ | Bcond (_, _, _, _, _) -> failwith "recompute_inumbs: unsupported Bcond"
+ in
+ walk ibf.entry None;
+ btl
+
+let regenerate_btl_tree btl entry =
+ let new_entry = ref entry in
+ let rec renumber_iblock ib =
+ let get_new_succ s =
+ let sentry = get_some @@ PTree.get s btl in
+ i2p (get_inumb_or_next sentry.entry)
+ in
+ match ib with
+ | BF (Bcall (sign, fn, lr, rd, s), iinfo) ->
+ BF (Bcall (sign, fn, lr, rd, get_new_succ s), iinfo)
+ | BF (Bbuiltin (sign, fn, lr, s), iinfo) ->
+ BF (Bbuiltin (sign, fn, lr, get_new_succ s), iinfo)
+ | BF (Bgoto s, iinfo) -> BF (Bgoto (get_new_succ s), iinfo)
+ | BF (Bjumptable (arg, tbl), iinfo) ->
+ let tbl' = List.map (fun s -> get_new_succ s) tbl in
+ BF (Bjumptable (arg, tbl'), iinfo)
+ | Bcond (cond, lr, ib1, ib2, iinfo) ->
+ Bcond (cond, lr, renumber_iblock ib1, renumber_iblock ib2, iinfo)
+ | Bseq (ib1, ib2) -> Bseq (renumber_iblock ib1, renumber_iblock ib2)
+ | _ -> ib
+ in
+ let dm = ref PTree.empty in
+ let ord_btl =
+ PTree.fold
+ (fun ord_btl old_n old_ibf ->
+ let ib = renumber_iblock old_ibf.entry in
+ let n = get_inumb_or_next ib in
+ let n_pos = i2p n in
+ let bi = mk_binfo n old_ibf.binfo.s_output_regs old_ibf.binfo.typing in
+ let ibf = { entry = ib; input_regs = old_ibf.input_regs; binfo = bi } in
+ if old_n = entry then new_entry := n_pos;
+ dm := PTree.set old_n n_pos !dm;
+ PTree.set n_pos ibf ord_btl)
+ btl PTree.empty
+ in
+ debug "Renumbered BTL with new_entry=%d:\n" (p2i !new_entry);
+ print_btl_code stderr ord_btl;
+ ((ord_btl, !new_entry), !dm)
+
+let renumber btl entry =
+ (*debug_flag := true;*)
+ let btl' = recompute_inumbs btl entry in
+ (*debug_flag := false;*)
+ regenerate_btl_tree btl' entry
diff --git a/scheduling/BTLScheduleraux.ml b/scheduling/BTLScheduleraux.ml
new file mode 100644
index 00000000..75672243
--- /dev/null
+++ b/scheduling/BTLScheduleraux.ml
@@ -0,0 +1,337 @@
+open Maps
+open BTL
+open BTLtypes
+open Machine
+open DebugPrint
+open PrintBTL
+open Registers
+open RTLcommonaux
+open BTLcommonaux
+open ExpansionOracle
+open PrepassSchedulingOracle
+
+let flatten_blk_basics ibf =
+ let ib = ibf.entry in
+ let last = ref None in
+ let rec traverse_blk ib =
+ match ib with
+ | BF (_, _) ->
+ last := Some ib;
+ []
+ | Bseq ((Bcond (_, _, _, _, iinfo) as ib1), ib2) -> (
+ match iinfo.pcond with
+ | Some _ -> [ ib1 ] @ traverse_blk ib2
+ | None ->
+ last := Some ib;
+ [])
+ | Bseq (ib1, ib2) -> traverse_blk ib1 @ traverse_blk ib2
+ | _ -> [ ib ]
+ in
+ let ibl = traverse_blk ib in
+ (Array.of_list ibl, !last)
+
+let is_a_cb = function Bcond _ -> true | _ -> false
+
+let is_a_load = function Bload _ -> true | _ -> false
+
+module SI = Set.Make (Int)
+
+let find_array arr n =
+ let index = ref None in
+ (try
+ Array.iteri
+ (fun i n' ->
+ match !index with
+ | Some _ -> raise Exit
+ | None -> if n = n' then index := Some i)
+ arr
+ with Exit -> ());
+ get_some @@ !index
+
+let count_cbs bseq olast indexes =
+ let current_cbs = ref SI.empty in
+ let cbs_above = Hashtbl.create 100 in
+ let update_cbs n ib =
+ print_btl_inst stderr ib;
+ if is_a_cb ib then current_cbs := SI.add indexes.(n) !current_cbs
+ else if is_a_load ib then Hashtbl.add cbs_above indexes.(n) !current_cbs
+ in
+ Array.iteri (fun n ib -> update_cbs n ib) bseq;
+ (match olast with
+ | Some last -> update_cbs (Array.length bseq) last
+ | None -> ());
+ cbs_above
+
+let apply_schedule bseq olast positions =
+ let bseq_new = Array.map (fun i -> bseq.(i)) positions in
+ (if !config.has_non_trapping_loads then
+ let fmap n = find_array positions n in
+ let seq = Array.init (Array.length positions) (fun i -> i) in
+ let fseq = Array.map fmap seq in
+ let cbs_above_old = count_cbs bseq olast fseq in
+ let cbs_above_new = count_cbs bseq_new olast seq in
+ Array.iteri
+ (fun n ib ->
+ let n' = fseq.(n) in
+ match ib with
+ | Bload (t, a, b, c, d, e) ->
+ let set_old = Hashtbl.find cbs_above_old n' in
+ let set_new = Hashtbl.find cbs_above_new n' in
+ if SI.subset set_old set_new then
+ bseq_new.(n') <- Bload (AST.TRAP, a, b, c, d, e)
+ else assert !config.has_non_trapping_loads
+ | _ -> ())
+ bseq);
+ let ibl = Array.to_list bseq_new in
+ let rec build_iblock = function
+ | [] -> failwith "build_iblock: empty list"
+ | [ ib ] -> ( match olast with Some last -> Bseq (ib, last) | None -> ib)
+ | ib1 :: ib2 :: k -> Bseq (ib1, build_iblock (ib2 :: k))
+ in
+ build_iblock ibl
+
+(** the useful one. Returns a hashtable with bindings of shape
+ ** [(r,(t, n)], where [r] is a pseudo-register (Registers.reg),
+ ** [t] is its class (according to [typing]), and [n] the number of
+ ** times it's referenced as an argument in instructions of [seqa] ;
+ ** and an arrray containg the list of regs referenced by each
+ ** instruction, with a boolean to know whether it's as arg or dest *)
+let reference_counting (seqa : (iblock * Regset.t) array) (out_regs : Regset.t)
+ (typing : RTLtyping.regenv) :
+ (reg, int * int) Hashtbl.t * (reg * bool) list array =
+ let retl = Hashtbl.create 42 in
+ let retr = Array.make (Array.length seqa) [] in
+ (* retr.(i) : (r, b) -> (r', b') -> ...
+ * where b = true if seen as arg, false if seen as dest
+ *)
+ List.iter
+ (fun reg ->
+ Hashtbl.add retl reg (Machregsaux.class_of_type (typing reg), 1))
+ (Regset.elements out_regs);
+ let add_reg reg =
+ match Hashtbl.find_opt retl reg with
+ | Some (t, n) -> Hashtbl.add retl reg (t, n + 1)
+ | None -> Hashtbl.add retl reg (Machregsaux.class_of_type (typing reg), 1)
+ in
+ let map_true = List.map (fun r -> (r, true)) in
+ Array.iteri
+ (fun i (ins, _) ->
+ match ins with
+ | Bop (_, args, dest, _) | Bload (_, _, _, args, dest, _) ->
+ List.iter add_reg args;
+ retr.(i) <- (dest, false) :: map_true args
+ | Bcond (_, args, _, _, _) ->
+ List.iter add_reg args;
+ retr.(i) <- map_true args
+ | Bstore (_, _, args, src, _) ->
+ List.iter add_reg args;
+ add_reg src;
+ retr.(i) <- (src, true) :: map_true args
+ | BF (Bcall (_, fn, args, dest, _), _) ->
+ List.iter add_reg args;
+ retr.(i) <-
+ (match fn with
+ | Datatypes.Coq_inl reg ->
+ add_reg reg;
+ (dest, false) :: (reg, true) :: map_true args
+ | _ -> (dest, false) :: map_true args)
+ | BF (Btailcall (_, fn, args), _) ->
+ List.iter add_reg args;
+ retr.(i) <-
+ (match fn with
+ | Datatypes.Coq_inl reg ->
+ add_reg reg;
+ (reg, true) :: map_true args
+ | _ -> map_true args)
+ | BF (Bbuiltin (_, args, dest, _), _) ->
+ let rec bar = function
+ | AST.BA r ->
+ add_reg r;
+ retr.(i) <- (r, true) :: retr.(i)
+ | AST.BA_splitlong (hi, lo) | AST.BA_addptr (hi, lo) ->
+ bar hi;
+ bar lo
+ | _ -> ()
+ in
+ List.iter bar args;
+ let rec bad = function
+ | AST.BR r -> retr.(i) <- (r, false) :: retr.(i)
+ | AST.BR_splitlong (hi, lo) ->
+ bad hi;
+ bad lo
+ | _ -> ()
+ in
+ bad dest
+ | BF (Bjumptable (reg, _), _) | BF (Breturn (Some reg), _) ->
+ add_reg reg;
+ retr.(i) <- [ (reg, true) ]
+ | _ -> ())
+ seqa;
+ (* print_string "mentions\n";
+ * Array.iteri (fun i l ->
+ * print_int i;
+ * print_string ": [";
+ * List.iter (fun (r, b) ->
+ * print_int (Camlcoq.P.to_int r);
+ * print_string ":";
+ * print_string (if b then "a:" else "d");
+ * if b then print_int (snd (Hashtbl.find retl r));
+ * print_string ", "
+ * ) l;
+ * print_string "]\n";
+ * flush stdout;
+ * ) retr; *)
+ (retl, retr)
+
+let get_live_regs_entry seqa ibf btl =
+ if !Clflags.option_debug_compcert > 6 then debug_flag := true;
+ let ret =
+ Array.fold_right
+ (fun (ins, liveins) regset_i ->
+ let regset = Regset.union liveins regset_i in
+ match ins with
+ | Bnop _ -> regset
+ | Bop (_, args, dest, _) | Bload (_, _, _, args, dest, _) ->
+ List.fold_left
+ (fun set reg -> Regset.add reg set)
+ (Regset.remove dest regset)
+ args
+ | Bstore (_, _, args, src, _) ->
+ List.fold_left
+ (fun set reg -> Regset.add reg set)
+ (Regset.add src regset) args
+ | BF (Bcall (_, fn, args, dest, _), _) ->
+ List.fold_left
+ (fun set reg -> Regset.add reg set)
+ ((match fn with
+ | Datatypes.Coq_inl reg -> Regset.add reg
+ | Datatypes.Coq_inr _ -> fun x -> x)
+ (Regset.remove dest regset))
+ args
+ | BF (Btailcall (_, fn, args), _) ->
+ List.fold_left
+ (fun set reg -> Regset.add reg set)
+ (match fn with
+ | Datatypes.Coq_inl reg -> Regset.add reg regset
+ | Datatypes.Coq_inr _ -> regset)
+ args
+ | BF (Bbuiltin (_, args, dest, _), _) ->
+ List.fold_left
+ (fun set arg ->
+ let rec add reg set =
+ match reg with
+ | AST.BA r -> Regset.add r set
+ | AST.BA_splitlong (hi, lo) | AST.BA_addptr (hi, lo) ->
+ add hi (add lo set)
+ | _ -> set
+ in
+ add arg set)
+ (let rec rem dest set =
+ match dest with
+ | AST.BR r -> Regset.remove r set
+ | AST.BR_splitlong (hi, lo) -> rem hi (rem lo set)
+ | _ -> set
+ in
+ rem dest regset)
+ args
+ | BF (Bjumptable (reg, _), _) | BF (Breturn (Some reg), _) ->
+ Regset.add reg regset
+ | Bcond (_, args, _, _, _) ->
+ List.fold_left (fun set reg -> Regset.add reg set) regset args
+ | _ -> regset)
+ seqa ibf.binfo.s_output_regs
+ in
+ debug "live in regs: ";
+ print_regset ret;
+ debug "\n";
+ debug_flag := false;
+ ret
+
+let schedule_blk n ibf btl =
+ if not !Clflags.option_fprepass then btl
+ else
+ let bseq, olast = flatten_blk_basics ibf in
+ let seqa =
+ Array.mapi
+ (fun i inst -> (inst, get_liveins inst))
+ bseq
+ in
+ let live_regs_entry = get_live_regs_entry seqa ibf btl in
+ let refcount =
+ reference_counting seqa ibf.binfo.s_output_regs ibf.binfo.typing
+ in
+ match
+ schedule_sequence seqa btl live_regs_entry ibf.binfo.typing refcount
+ with
+ | Some positions ->
+ let new_ib = apply_schedule bseq olast positions in
+ let new_ibf =
+ { entry = new_ib; binfo = ibf.binfo; input_regs = ibf.input_regs }
+ in
+ PTree.set n new_ibf btl
+ | None -> btl
+
+let turn_all_loads_nontrap n ibf btl =
+ if not !config.has_non_trapping_loads then btl
+ else
+ let rec traverse_rec ib =
+ match ib with
+ | Bseq (ib1, ib2) -> Bseq (traverse_rec ib1, traverse_rec ib2)
+ | Bload (t, a, b, c, d, e) -> Bload (AST.NOTRAP, a, b, c, d, e)
+ | _ -> ib
+ in
+ let ib' = traverse_rec ibf.entry in
+ let ibf' =
+ { entry = ib'; input_regs = ibf.input_regs; binfo = ibf.binfo }
+ in
+ PTree.set n ibf' btl
+
+let compute_liveins n ibf btl =
+ let rec traverse_rec ib =
+ match ib with
+ | Bseq (ib1, ib2) ->
+ traverse_rec ib1;
+ traverse_rec ib2
+ | BF (Bgoto succ, iinfo)
+ | BF (Bcall (_, _, _, _, succ), iinfo)
+ | BF (Bbuiltin (_, _, _, succ), iinfo) ->
+ let lives = (get_some @@ PTree.get succ btl).input_regs in
+ iinfo.liveins <- lives
+ | BF (Bjumptable (_, tbl), iinfo) ->
+ List.iter
+ (fun ex ->
+ let lives = (get_some @@ PTree.get ex btl).input_regs in
+ iinfo.liveins <- Regset.union iinfo.liveins lives)
+ tbl
+ | Bcond (_, _, BF (Bgoto succ, _), Bnop None, iinfo) -> (
+ match iinfo.pcond with
+ | Some predb ->
+ assert (predb = false);
+ let lives = (get_some @@ PTree.get succ btl).input_regs in
+ iinfo.liveins <- lives
+ | None -> ())
+ | _ -> ()
+ in
+ traverse_rec ibf.entry
+
+let rec do_schedule btl = function
+ | [] -> btl
+ | (n, ibf) :: blks ->
+ compute_liveins n ibf btl;
+ let code_exp = expanse n ibf btl in
+ let ibf_exp = get_some @@ PTree.get n code_exp in
+ let code_nt = turn_all_loads_nontrap n ibf_exp code_exp in
+ let ibf_nt = get_some @@ PTree.get n code_nt in
+ let btl' = schedule_blk n ibf_nt code_nt in
+ (*debug_flag := true;*)
+ print_btl_code stderr btl;
+ print_btl_code stderr btl';
+ (*debug_flag := false;*)
+ do_schedule btl' blks
+
+let btl_scheduler f =
+ let btl = f.fn_code in
+ let elts = PTree.elements btl in
+ find_last_reg elts;
+ let btl' = do_schedule btl elts in
+ btl'
diff --git a/scheduling/BTL_Livecheck.v b/scheduling/BTL_Livecheck.v
new file mode 100644
index 00000000..d200b9bd
--- /dev/null
+++ b/scheduling/BTL_Livecheck.v
@@ -0,0 +1,690 @@
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep Op Registers OptionMonad.
+Require Import Errors RTL BTL BTLmatchRTL.
+
+
+Local Open Scope lazy_bool_scope.
+
+Local Open Scope option_monad_scope.
+
+Fixpoint list_mem (rl: list reg) (alive: Regset.t) {struct rl}: bool :=
+ match rl with
+ | nil => true
+ | r1 :: rs => Regset.mem r1 alive &&& list_mem rs alive
+ end.
+
+Definition reg_option_mem (or: option reg) (alive: Regset.t) :=
+ match or with None => true | Some r => Regset.mem r alive end.
+
+Definition reg_sum_mem (ros: reg + ident) (alive: Regset.t) :=
+ match ros with inl r => Regset.mem r alive | inr s => true end.
+
+(* NB: definition following [regmap_setres] in [RTL.step] semantics *)
+Definition reg_builtin_res (res: builtin_res reg) (alive: Regset.t): Regset.t :=
+ match res with
+ | BR r => Regset.add r alive
+ | _ => alive
+ end.
+
+Definition exit_checker (btl: code) (alive: Regset.t) (s: node): option unit :=
+ SOME next <- btl!s IN
+ ASSERT Regset.subset next.(input_regs) alive IN
+ Some tt.
+
+Fixpoint exit_list_checker (btl: code) (alive: Regset.t) (l: list node): bool :=
+ match l with
+ | nil => true
+ | s :: l' => exit_checker btl alive s &&& exit_list_checker btl alive l'
+ end.
+
+Definition final_inst_checker (btl: code) (alive: Regset.t) (fin: final): option unit :=
+ match fin with
+ | Bgoto s =>
+ exit_checker btl alive s
+ | Breturn oreg =>
+ ASSERT reg_option_mem oreg alive IN Some tt
+ | Bcall _ ros args res s =>
+ ASSERT list_mem args alive IN
+ ASSERT reg_sum_mem ros alive IN
+ exit_checker btl (Regset.add res alive) s
+ | Btailcall _ ros args =>
+ ASSERT list_mem args alive IN
+ ASSERT reg_sum_mem ros alive IN Some tt
+ | Bbuiltin _ args res s =>
+ ASSERT list_mem (params_of_builtin_args args) alive IN
+ exit_checker btl (reg_builtin_res res alive) s
+ | Bjumptable arg tbl =>
+ ASSERT Regset.mem arg alive IN
+ ASSERT exit_list_checker btl alive tbl IN Some tt
+ end.
+
+(* This definition is the meet (infimum) subset of alive registers,
+ used for conditions by the below checker.
+ A None argument represents the neutral element for intersection. *)
+Definition meet (o1 o2: option Regset.t): option Regset.t :=
+ match o1, o2 with
+ | None, _ => o2
+ | _, None => o1
+ | Some alive1, Some alive2 => Some (Regset.inter alive1 alive2)
+ end.
+
+Fixpoint body_checker (btl: code) (ib: iblock) (alive: Regset.t): option (option Regset.t) :=
+ match ib with
+ | Bseq ib1 ib2 =>
+ SOME oalive1 <- body_checker btl ib1 alive IN
+ SOME alive1 <- oalive1 IN
+ body_checker btl ib2 alive1
+ | Bnop _ => Some (Some alive)
+ | Bop _ args dest _ =>
+ ASSERT list_mem args alive IN
+ Some (Some (Regset.add dest alive))
+ | Bload _ _ _ args dest _ =>
+ ASSERT list_mem args alive IN
+ Some (Some (Regset.add dest alive))
+ | Bstore _ _ args src _ =>
+ ASSERT Regset.mem src alive IN
+ ASSERT list_mem args alive IN
+ Some (Some alive)
+ | Bcond _ args ib1 ib2 _ =>
+ ASSERT list_mem args alive IN
+ SOME oalive1 <- body_checker btl ib1 alive IN
+ SOME oalive2 <- body_checker btl ib2 alive IN
+ Some (meet oalive1 oalive2)
+ | BF fin _ =>
+ SOME _ <- final_inst_checker btl alive fin IN
+ Some None
+ end.
+
+(* This definition simply convert the result in an option unit *)
+Definition iblock_checker (btl: code) (ib: iblock) (alive: Regset.t): option unit :=
+ SOME _ <- body_checker btl ib alive IN Some tt.
+
+Fixpoint list_iblock_checker (btl: code) (l: list (node*iblock_info)): bool :=
+ match l with
+ | nil => true
+ | (_, ibf) :: l' => iblock_checker btl ibf.(entry) ibf.(input_regs) &&& list_iblock_checker btl l'
+ end.
+
+Lemma lazy_and_Some_true A (o: option A) (b: bool): o &&& b = true <-> (exists v, o = Some v) /\ b = true.
+Proof.
+ destruct o; simpl; intuition.
+ - eauto.
+ - firstorder. try_simplify_someHyps.
+Qed.
+
+Lemma lazy_and_Some_tt_true (o: option unit) (b: bool): o &&& b = true <-> o = Some tt /\ b = true.
+Proof.
+ intros; rewrite lazy_and_Some_true; firstorder.
+ destruct x; auto.
+Qed.
+
+Lemma list_iblock_checker_correct btl l:
+ list_iblock_checker btl l = true ->
+ forall e, List.In e l -> iblock_checker btl (snd e).(entry) (snd e).(input_regs) = Some tt.
+Proof.
+ intros CHECKER e H; induction l as [|(n & ibf) l]; intuition.
+ simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto).
+Qed.
+
+Definition liveness_checker_bool (f: BTL.function): bool :=
+ f.(fn_code)!(f.(fn_entrypoint)) &&& list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)).
+
+Definition liveness_checker (f: BTL.function): res unit :=
+ match liveness_checker_bool f with
+ | true => OK tt
+ | false => Error (msg "BTL_Livecheck: liveness_checker failed")
+ end.
+
+Lemma decomp_liveness_checker f:
+ liveness_checker f = OK tt ->
+ exists ibf, f.(fn_code)!(f.(fn_entrypoint)) = Some ibf /\
+ list_iblock_checker f.(fn_code) (PTree.elements f.(fn_code)) = true.
+Proof.
+ intros LIVE; unfold liveness_checker in LIVE.
+ destruct liveness_checker_bool eqn:EQL; try congruence.
+ clear LIVE. unfold liveness_checker_bool in EQL.
+ rewrite lazy_and_Some_true in EQL; destruct EQL as [[ibf ENTRY] LIST].
+ eexists; split; eauto.
+Qed.
+
+Lemma liveness_checker_correct f n ibf:
+ liveness_checker f = OK tt ->
+ f.(fn_code)!n = Some ibf ->
+ iblock_checker f.(fn_code) ibf.(entry) ibf.(input_regs) = Some tt.
+Proof.
+ intros LIVE PC.
+ apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]].
+ exploit list_iblock_checker_correct; eauto.
+ - eapply PTree.elements_correct; eauto.
+ - simpl; auto.
+Qed.
+
+Lemma liveness_checker_entrypoint f:
+ liveness_checker f = OK tt ->
+ f.(fn_code)!(f.(fn_entrypoint)) <> None.
+Proof.
+ intros LIVE; apply decomp_liveness_checker in LIVE; destruct LIVE as [ibf' [ENTRY LIST]].
+ unfold not; intros CONTRA. congruence.
+Qed.
+
+Definition liveness_ok_function (f: BTL.function): Prop := liveness_checker f = OK tt.
+
+Inductive liveness_ok_fundef: fundef -> Prop :=
+ | liveness_ok_Internal f: liveness_ok_function f -> liveness_ok_fundef (Internal f)
+ | liveness_ok_External ef: liveness_ok_fundef (External ef).
+
+
+Local Notation ext alive := (fun r => Regset.In r alive).
+
+Definition ext_opt (oalive: option Regset.t): Regset.elt -> Prop :=
+ match oalive with
+ | Some alive => ext alive
+ | None => fun _ => True
+ end.
+
+Lemma ext_opt_meet: forall r oalive1 oalive2,
+ ext_opt (meet oalive1 oalive2) r ->
+ ext_opt oalive1 r /\ ext_opt oalive2 r.
+Proof.
+ intros. destruct oalive1, oalive2;
+ simpl in *; intuition.
+ eapply Regset.inter_1; eauto.
+ eapply Regset.inter_2; eauto.
+Qed.
+
+Lemma regset_add_spec live r1 r2: Regset.In r1 (Regset.add r2 live) <-> (r1 = r2 \/ Regset.In r1 live).
+Proof.
+ destruct (Pos.eq_dec r1 r2).
+ - subst. intuition; eapply Regset.add_1; auto.
+ - intuition.
+ * right. eapply Regset.add_3; eauto.
+ * eapply Regset.add_2; auto.
+Qed.
+
+Local Hint Resolve Regset.mem_2 Regset.subset_2: core.
+
+Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true.
+Proof.
+ destruct b1; simpl; intuition.
+Qed.
+
+Lemma list_mem_correct (rl: list reg) (alive: Regset.t):
+ list_mem rl alive = true -> forall r, List.In r rl -> ext alive r.
+Proof.
+ induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto.
+Qed.
+
+Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop :=
+ forall r, (alive r) -> rs1#r = rs2#r.
+
+Lemma eqlive_reg_update (alive: Regset.elt -> Prop) rs1 rs2 r v: eqlive_reg (fun r1 => r1 <> r /\ alive r1) rs1 rs2 -> eqlive_reg alive (rs1 # r <- v) (rs2 # r <- v).
+Proof.
+ unfold eqlive_reg; intros EQLIVE r0 ALIVE.
+ destruct (Pos.eq_dec r r0) as [H|H].
+ - subst. rewrite! Regmap.gss. auto.
+ - rewrite! Regmap.gso; auto.
+Qed.
+
+Lemma eqlive_reg_monotonic (alive1 alive2: Regset.elt -> Prop) rs1 rs2: eqlive_reg alive2 rs1 rs2 -> (forall r, alive1 r -> alive2 r) -> eqlive_reg alive1 rs1 rs2.
+Proof.
+ unfold eqlive_reg; intuition.
+Qed.
+
+Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args.
+Proof.
+ induction args; simpl; auto.
+ intros EQLIVE ALIVE; rewrite IHargs; auto.
+ unfold eqlive_reg in EQLIVE.
+ rewrite EQLIVE; auto.
+Qed.
+
+Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args.
+Proof.
+ intros; eapply eqlive_reg_list; eauto.
+ intros; eapply list_mem_correct; eauto.
+Qed.
+
+Inductive eqlive_stackframes: stackframe -> stackframe -> Prop :=
+ | eqlive_stackframes_intro ibf res f sp pc rs1 rs2
+ (LIVE: liveness_ok_function f)
+ (ENTRY: f.(fn_code)!pc = Some ibf)
+ (EQUIV: forall v, eqlive_reg (ext ibf.(input_regs)) (rs1 # res <- v) (rs2 # res <- v)):
+ eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2).
+
+Inductive eqlive_states: state -> state -> Prop :=
+ | eqlive_states_intro
+ ibf st1 st2 f sp pc rs1 rs2 m
+ (STACKS: list_forall2 eqlive_stackframes st1 st2)
+ (LIVE: liveness_ok_function f)
+ (PATH: f.(fn_code)!pc = Some ibf)
+ (EQUIV: eqlive_reg (ext ibf.(input_regs)) rs1 rs2):
+ eqlive_states (State st1 f sp pc rs1 m) (State st2 f sp pc rs2 m)
+ | eqlive_states_call st1 st2 f args m
+ (LIVE: liveness_ok_fundef f)
+ (STACKS: list_forall2 eqlive_stackframes st1 st2):
+ eqlive_states (Callstate st1 f args m) (Callstate st2 f args m)
+ | eqlive_states_return st1 st2 v m
+ (STACKS: list_forall2 eqlive_stackframes st1 st2):
+ eqlive_states (Returnstate st1 v m) (Returnstate st2 v m).
+
+Section FSEM_SIMULATES_CFGSEM.
+
+Variable prog: BTL.program.
+
+Let ge := Genv.globalenv prog.
+
+Hypothesis all_fundef_liveness_ok: forall b f, Genv.find_funct_ptr ge b = Some f -> liveness_ok_fundef f.
+
+Local Hint Constructors eqlive_stackframes eqlive_states final_step list_forall2 step: core.
+
+Lemma eqlive_reg_update_gso alive rs1 rs2 res r: forall v : val,
+ eqlive_reg (ext alive) rs1 # res <- v rs2 # res <- v ->
+ res <> r -> Regset.In r alive ->
+ rs1 # r = rs2 # r.
+Proof.
+ intros v REGS NRES INR. unfold eqlive_reg in REGS.
+ specialize REGS with r. apply REGS in INR.
+ rewrite !Regmap.gso in INR; auto.
+Qed.
+
+Lemma find_funct_liveness_ok v fd:
+ Genv.find_funct ge v = Some fd -> liveness_ok_fundef fd.
+Proof.
+ unfold Genv.find_funct.
+ destruct v; try congruence.
+ destruct (Integers.Ptrofs.eq_dec _ _); try congruence.
+ eapply all_fundef_liveness_ok; eauto.
+Qed.
+
+Lemma find_function_liveness_ok ros rs f:
+ find_function ge ros rs = Some f -> liveness_ok_fundef f.
+Proof.
+ destruct ros as [r|i]; simpl.
+ - intros; eapply find_funct_liveness_ok; eauto.
+ - destruct (Genv.find_symbol ge i); try congruence.
+ eapply all_fundef_liveness_ok; eauto.
+Qed.
+
+Lemma find_function_eqlive alive ros rs1 rs2:
+ eqlive_reg (ext alive) rs1 rs2 ->
+ reg_sum_mem ros alive = true ->
+ find_function ge ros rs1 = find_function ge ros rs2.
+Proof.
+ intros EQLIVE.
+ destruct ros; simpl; auto.
+ intros H; erewrite (EQLIVE r); eauto.
+Qed.
+
+Lemma exit_checker_eqlive (btl: code) (alive: Regset.t) (pc: node) rs1 rs2:
+ exit_checker btl alive pc = Some tt ->
+ eqlive_reg (ext alive) rs1 rs2 ->
+ exists ibf, btl!pc = Some ibf /\ eqlive_reg (ext ibf.(input_regs)) rs1 rs2.
+Proof.
+ unfold exit_checker.
+ inversion_SOME next.
+ inversion_ASSERT. try_simplify_someHyps.
+ repeat (econstructor; eauto).
+ intros; eapply eqlive_reg_monotonic; eauto.
+ intros; exploit Regset.subset_2; eauto.
+Qed.
+
+Lemma exit_list_checker_eqlive (btl: code) (alive: Regset.t) (tbl: list node) rs1 rs2 pc: forall n,
+ exit_list_checker btl alive tbl = true ->
+ eqlive_reg (ext alive) rs1 rs2 ->
+ list_nth_z tbl n = Some pc ->
+ exists ibf, btl!pc = Some ibf /\ eqlive_reg (ext ibf.(input_regs)) rs1 rs2.
+Proof.
+ induction tbl; simpl.
+ - intros; try congruence.
+ - intros n; rewrite lazy_and_Some_tt_true; destruct (zeq n 0) eqn: Hn.
+ * try_simplify_someHyps; intuition.
+ exploit exit_checker_eqlive; eauto.
+ * intuition. eapply IHtbl; eauto.
+Qed.
+
+Lemma exit_checker_eqlive_update (btl: code) (alive: Regset.t) (pc: node) r rs1 rs2:
+ exit_checker btl (Regset.add r alive) pc = Some tt ->
+ eqlive_reg (ext alive) rs1 rs2 ->
+ exists ibf, btl!pc = Some ibf /\ (forall v, eqlive_reg (ext ibf.(input_regs)) (rs1 # r <- v) (rs2 # r <- v)).
+Proof.
+ unfold exit_checker.
+ inversion_SOME next.
+ inversion_ASSERT. try_simplify_someHyps.
+ repeat (econstructor; eauto).
+ intros; eapply eqlive_reg_update; eauto.
+ eapply eqlive_reg_monotonic; eauto.
+ intros r0 [X1 X2]; exploit Regset.subset_2; eauto.
+ rewrite regset_add_spec. intuition subst.
+Qed.
+
+Lemma exit_checker_eqlive_builtin_res (btl: code) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg):
+ exit_checker btl (reg_builtin_res res alive) pc = Some tt ->
+ eqlive_reg (ext alive) rs1 rs2 ->
+ exists ibf, btl!pc = Some ibf /\ (forall vres, eqlive_reg (ext ibf.(input_regs)) (regmap_setres res vres rs1) (regmap_setres res vres rs2)).
+Proof.
+ destruct res; simpl.
+ - intros; exploit exit_checker_eqlive_update; eauto.
+ - intros; exploit exit_checker_eqlive; eauto.
+ intros (ibf & PC & REGS).
+ eexists; intuition eauto.
+ - intros; exploit exit_checker_eqlive; eauto.
+ intros (ibf & PC & REGS).
+ eexists; intuition eauto.
+Qed.
+
+Local Hint Resolve in_or_app: local.
+Lemma eqlive_eval_builtin_args alive rs1 rs2 sp m args vargs:
+ eqlive_reg alive rs1 rs2 ->
+ Events.eval_builtin_args ge (fun r => rs1 # r) sp m args vargs ->
+ (forall r, List.In r (params_of_builtin_args args) -> alive r) ->
+ Events.eval_builtin_args ge (fun r => rs2 # r) sp m args vargs.
+Proof.
+ unfold Events.eval_builtin_args.
+ intros EQLIVE; induction 1 as [|a1 al b1 bl EVAL1 EVALL]; simpl.
+ { econstructor; eauto. }
+ intro X.
+ assert (X1: eqlive_reg (fun r => In r (params_of_builtin_arg a1)) rs1 rs2).
+ { eapply eqlive_reg_monotonic; eauto with local. }
+ lapply IHEVALL; eauto with local.
+ clear X IHEVALL; intro X. econstructor; eauto.
+ generalize X1; clear EVALL X1 X.
+ induction EVAL1; simpl; try (econstructor; eauto; fail).
+ - intros X1; erewrite X1; [ econstructor; eauto | eauto ].
+ - intros; econstructor.
+ + eapply IHEVAL1_1; eauto.
+ eapply eqlive_reg_monotonic; eauto.
+ simpl; intros; eauto with local.
+ + eapply IHEVAL1_2; eauto.
+ eapply eqlive_reg_monotonic; eauto.
+ simpl; intros; eauto with local.
+ - intros; econstructor.
+ + eapply IHEVAL1_1; eauto.
+ eapply eqlive_reg_monotonic; eauto.
+ simpl; intros; eauto with local.
+ + eapply IHEVAL1_2; eauto.
+ eapply eqlive_reg_monotonic; eauto.
+ simpl; intros; eauto with local.
+Qed.
+
+Lemma tr_inputs_eqlive_None f pc tbl ibf rs1 rs2
+ (PC: (fn_code f) ! pc = Some ibf)
+ (REGS: eqlive_reg (ext (input_regs ibf)) rs1 rs2)
+ :eqlive_reg (ext (input_regs ibf)) (tid f (pc :: tbl) None rs1)
+ (tr_inputs f (pc :: tbl) None rs2).
+Proof.
+ unfold eqlive_reg. intros r INR.
+ unfold tid. rewrite tr_inputs_get.
+ simpl. rewrite PC.
+ exploit Regset.union_3. eapply INR.
+ intros INRU. eapply Regset.mem_1 in INRU.
+ erewrite INRU; eauto.
+Qed.
+
+Lemma tr_inputs_eqlive_list_None tbl: forall f pc n alive ibf rs1 rs2
+ (REGS1: eqlive_reg (ext alive) rs1 rs2)
+ (EXIT_LIST: exit_list_checker (fn_code f) alive tbl = true)
+ (LIST: list_nth_z tbl n = Some pc)
+ (PC: (fn_code f) ! pc = Some ibf)
+ (REGS2: eqlive_reg (ext (input_regs ibf)) rs1 rs2),
+ eqlive_reg (ext (input_regs ibf)) (tid f tbl None rs1)
+ (tr_inputs f tbl None rs2).
+Proof.
+ induction tbl as [| pc' tbl IHtbl]; try_simplify_someHyps.
+ autodestruct; try_simplify_someHyps.
+ - intros; eapply tr_inputs_eqlive_None; eauto.
+ - rewrite lazy_and_Some_tt_true in EXIT_LIST.
+ destruct EXIT_LIST as [EXIT EXIT_REM].
+ intros. unfold eqlive_reg. intros r INR.
+ exploit (IHtbl f pc (Z.pred n) alive ibf rs1 rs2); eauto.
+ unfold tid. rewrite !tr_inputs_get.
+ exploit exit_checker_eqlive; eauto.
+ intros (ibf' & PC' & REGS3).
+ simpl; rewrite PC'. autodestruct.
+ + intro INRU. apply Regset.mem_2 in INRU.
+ intros EQR. eapply Regset.union_2 in INRU.
+ eapply Regset.mem_1 in INRU. erewrite INRU; auto.
+ + intros. autodestruct.
+ rewrite (REGS2 r); auto.
+Qed.
+
+Lemma tr_inputs_eqlive_update f pc ibf rs1 rs2 res
+ (PC: (fn_code f) ! pc = Some ibf)
+ :forall (v: val)
+ (REGS: eqlive_reg (ext (input_regs ibf)) rs1 # res <- v rs2 # res <- v),
+ eqlive_reg (ext (input_regs ibf))
+ (tid f (pc :: nil) (Some res) rs1) # res <- v
+ (tr_inputs f (pc :: nil) (Some res) rs2) # res <- v.
+Proof.
+ intros. apply eqlive_reg_update.
+ unfold eqlive_reg. intros r (NRES & INR).
+ unfold tid. rewrite tr_inputs_get.
+ simpl. rewrite PC. assert (NRES': res <> r) by auto.
+ clear NRES. exploit Regset.union_3. eapply INR.
+ intros INRU. exploit Regset.remove_2; eauto.
+ intros INRU_RES. eapply Regset.mem_1 in INRU_RES.
+ erewrite INRU_RES. eapply eqlive_reg_update_gso; eauto.
+Qed.
+
+Local Hint Resolve tr_inputs_eqlive_None tr_inputs_eqlive_update: core.
+Lemma cfgsem2fsem_finalstep_simu sp f stk1 stk2 s t fin alive rs1 m rs2
+ (FSTEP: final_step tid ge stk1 f sp rs1 m fin t s)
+ (LIVE: liveness_ok_function f)
+ (REGS: eqlive_reg (ext alive) rs1 rs2)
+ (FCHK: final_inst_checker (fn_code f) alive fin = Some tt)
+ (STACKS: list_forall2 eqlive_stackframes stk1 stk2)
+ :exists s',
+ final_step tr_inputs ge stk2 f sp rs2 m fin t s'
+ /\ eqlive_states s s'.
+Proof.
+ destruct FSTEP; try_simplify_someHyps; repeat inversion_ASSERT; intros.
+ - (* Bgoto *)
+ eexists; split.
+ + econstructor; eauto.
+ + exploit exit_checker_eqlive; eauto.
+ intros (ibf & PC & REGS').
+ econstructor; eauto.
+ - (* Breturn *)
+ eexists; split. econstructor; eauto.
+ destruct or; simpl in *;
+ try erewrite (REGS r); eauto.
+ - (* Bcall *)
+ exploit exit_checker_eqlive_update; eauto.
+ intros (ibf & PC & REGS').
+ eexists; split.
+ + econstructor; eauto.
+ erewrite <- find_function_eqlive; eauto.
+ + erewrite eqlive_reg_listmem; eauto.
+ eapply eqlive_states_call; eauto.
+ eapply find_function_liveness_ok; eauto.
+ - (* Btailcall *)
+ eexists; split.
+ + econstructor; eauto.
+ erewrite <- find_function_eqlive; eauto.
+ + erewrite eqlive_reg_listmem; eauto.
+ eapply eqlive_states_call; eauto.
+ eapply find_function_liveness_ok; eauto.
+ - (* Bbuiltin *)
+ exploit exit_checker_eqlive_builtin_res; eauto.
+ intros (ibf & PC & REGS').
+ eexists; split.
+ + econstructor; eauto.
+ eapply eqlive_eval_builtin_args; eauto.
+ intros; eapply list_mem_correct; eauto.
+ + repeat (econstructor; simpl; eauto).
+ unfold regmap_setres. destruct res; simpl in *; eauto.
+ - (* Bjumptable *)
+ exploit exit_list_checker_eqlive; eauto.
+ intros (ibf & PC & REGS').
+ eexists; split.
+ + econstructor; eauto.
+ erewrite <- REGS; eauto.
+ + repeat (econstructor; simpl; eauto).
+ apply (tr_inputs_eqlive_list_None tbl f pc' (Int.unsigned n) alive ibf rs1 rs2);
+ auto.
+Qed.
+
+Lemma cfgsem2fsem_ibistep_simu_None sp f ib: forall rs1 m rs1' m'
+ (ISTEP: iblock_istep ge sp rs1 m ib rs1' m' None)
+ alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2)
+ (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2)),
+ exists rs2',
+ iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' None)
+ /\ eqlive_reg (ext_opt oalive2) rs1' rs2'.
+Proof.
+ induction ib; intros; try_simplify_someHyps;
+ repeat inversion_ASSERT; intros; inv ISTEP.
+ - (* Bnop *)
+ inv BDY; eauto.
+ - (* Bop *)
+ erewrite <- eqlive_reg_listmem; eauto.
+ try_simplify_someHyps; intros.
+ repeat econstructor.
+ apply eqlive_reg_update.
+ eapply eqlive_reg_monotonic; eauto.
+ intros r0; rewrite regset_add_spec.
+ intuition.
+ - (* Bload *)
+ erewrite <- eqlive_reg_listmem; eauto.
+ try_simplify_someHyps; intros.
+ destruct trap; inv LOAD;
+ rewrite EVAL, LOAD0 || (autodestruct; try rewrite LOAD0; auto).
+ all:
+ repeat econstructor;
+ apply eqlive_reg_update;
+ eapply eqlive_reg_monotonic; eauto;
+ intros r0; rewrite regset_add_spec;
+ intuition.
+ - (* Bstore *)
+ erewrite <- eqlive_reg_listmem; eauto.
+ rewrite <- (REGS src); auto.
+ try_simplify_someHyps; intros.
+ rewrite STORE; eauto.
+ - (* Bseq continue *)
+ destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate.
+ generalize BDY; clear BDY.
+ inversion_SOME aliveMid; intros OALIVE BDY2. inv OALIVE.
+ exploit IHib1; eauto.
+ intros (rs2' & ISTEP1 & REGS1). rewrite ISTEP1; simpl.
+ eapply IHib2; eauto.
+ - (* Bcond *)
+ generalize BDY; clear BDY.
+ inversion_SOME oaliveSo; inversion_SOME oaliveNot; intros BDY1 BDY2 JOIN.
+ erewrite <- eqlive_reg_listmem; eauto.
+ rewrite EVAL.
+ destruct b; [ exploit IHib1; eauto | exploit IHib2; eauto].
+ all:
+ intros (rs2' & ISTEP1 & REGS1);
+ econstructor; split; eauto; inv JOIN;
+ eapply eqlive_reg_monotonic; eauto;
+ intros r EXTM; apply ext_opt_meet in EXTM; intuition.
+Qed.
+
+Lemma cfgsem2fsem_ibistep_simu_Some sp f stk1 stk2 ib: forall s t rs1 m rs1' m' fin
+ (ISTEP: iblock_istep ge sp rs1 m ib rs1' m' (Some fin))
+ (FSTEP: final_step tid ge stk1 f sp rs1' m' fin t s)
+ alive1 oalive2 rs2 (REGS: eqlive_reg (ext alive1) rs1 rs2)
+ (BDY: body_checker (fn_code f) ib alive1 = Some (oalive2))
+ (LIVE: liveness_ok_function f)
+ (STACKS: list_forall2 eqlive_stackframes stk1 stk2),
+ exists rs2' s',
+ iblock_istep_run ge sp ib rs2 m = Some (out rs2' m' (Some fin))
+ /\ final_step tr_inputs ge stk2 f sp rs2' m' fin t s'
+ /\ eqlive_states s s'.
+Proof.
+ induction ib; simpl; try_simplify_someHyps;
+ repeat inversion_ASSERT; intros; inv ISTEP.
+ - (* BF *)
+ generalize BDY; clear BDY.
+ inversion_SOME x; try_simplify_someHyps; intros FCHK.
+ destruct x; exploit cfgsem2fsem_finalstep_simu; eauto.
+ intros (s2 & FSTEP' & STATES); eauto.
+ - (* Bseq stop *)
+ destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate.
+ generalize BDY; clear BDY.
+ inversion_SOME aliveMid. intros OALIVE BDY2. inv OALIVE.
+ exploit IHib1; eauto. intros (rs2' & s' & ISTEP1 & FSTEP1 & STATES).
+ rewrite ISTEP1; simpl.
+ do 2 eexists; intuition eauto.
+ - (* Bseq continue *)
+ destruct (body_checker _ _ _) eqn:BDY1 in BDY; try discriminate.
+ generalize BDY; clear BDY.
+ inversion_SOME aliveMid; intros OALIVE BDY2. inv OALIVE.
+ exploit cfgsem2fsem_ibistep_simu_None; eauto.
+ intros (rs2' & ISTEP1 & REGS'). rewrite ISTEP1; simpl; eauto.
+ - (* Bcond *)
+ generalize BDY; clear BDY.
+ inversion_SOME oaliveSo; inversion_SOME oaliveNot; intros BDY1 BDY2 JOIN.
+ erewrite <- eqlive_reg_listmem; eauto. rewrite EVAL.
+ destruct b; eauto.
+Qed.
+
+Lemma cfgsem2fsem_ibstep_simu stk1 stk2 f sp rs1 m rs2 ibf pc s1 t:
+ iblock_step tid (Genv.globalenv prog) stk1 f sp rs1 m ibf.(entry) t s1 ->
+ list_forall2 eqlive_stackframes stk1 stk2 ->
+ eqlive_reg (ext (input_regs ibf)) rs1 rs2 ->
+ liveness_ok_function f ->
+ (fn_code f) ! pc = Some ibf ->
+ exists s2 : state,
+ iblock_step tr_inputs (Genv.globalenv prog) stk2 f sp rs2 m ibf.(entry) t s2 /\
+ eqlive_states s1 s2.
+Proof.
+ intros STEP STACKS EQLIVE LIVE PC.
+ assert (CHECKER: liveness_ok_function f) by auto.
+ unfold liveness_ok_function in CHECKER.
+ apply decomp_liveness_checker in CHECKER; destruct CHECKER as [ibf' [ENTRY LIST]].
+ eapply PTree.elements_correct in PC as PCIN.
+ eapply list_iblock_checker_correct in LIST as IBC; eauto.
+ unfold iblock_checker in IBC. generalize IBC; clear IBC.
+ inversion_SOME alive; intros BODY _.
+ destruct STEP as (rs1' & m1' & fin' & ISTEP & FSTEP).
+ exploit cfgsem2fsem_ibistep_simu_Some; eauto.
+ intros (rs2' & s' & ISTEP' & FSTEP' & REGS).
+ rewrite <- iblock_istep_run_equiv in ISTEP'. clear ISTEP.
+ unfold iblock_step. repeat eexists; eauto.
+Qed.
+
+Local Hint Constructors step: core.
+
+Lemma cfgsem2fsem_step_simu s1 s1' t s2:
+ step tid (Genv.globalenv prog) s1 t s1' ->
+ eqlive_states s1 s2 ->
+ exists s2' : state,
+ step tr_inputs (Genv.globalenv prog) s2 t s2' /\
+ eqlive_states s1' s2'.
+Proof.
+ destruct 1 as [stack ibf f sp n rs m t s ENTRY STEP | | | ]; intros STATES.
+ - (* iblock *)
+ inv STATES; simplify_someHyps.
+ intros.
+ exploit cfgsem2fsem_ibstep_simu; eauto.
+ intros (s2 & STEP2 & EQUIV2).
+ eexists; split; eauto.
+ - (* function internal *)
+ inv STATES; inv LIVE.
+ apply liveness_checker_entrypoint in H0 as ENTRY.
+ destruct ((fn_code f) ! (fn_entrypoint f)) eqn:EQENTRY; try congruence; eauto.
+ eexists; split; repeat econstructor; eauto.
+ - (* function external *)
+ inv STATES; inv LIVE; eexists; split; econstructor; eauto.
+ - (* return *)
+ inv STATES.
+ inversion STACKS as [|s1 st1 s' s2 STACK STACKS']; subst; clear STACKS.
+ inv STACK.
+ exists (State s2 f sp pc (rs2 # res <- vres) m); split.
+ * apply exec_return.
+ * eapply eqlive_states_intro; eauto.
+Qed.
+
+Theorem cfgsem2fsem: forward_simulation (cfgsem prog) (fsem prog).
+Proof.
+ eapply forward_simulation_step with eqlive_states; simpl; eauto.
+ - destruct 1, f; intros; eexists; intuition eauto;
+ repeat (econstructor; eauto).
+ - intros s1 s2 r STATES FINAL; destruct FINAL.
+ inv STATES; inv STACKS; constructor.
+ - intros. eapply cfgsem2fsem_step_simu; eauto.
+Qed.
+
+End FSEM_SIMULATES_CFGSEM.
+
+
diff --git a/scheduling/BTL_SEimpl.v b/scheduling/BTL_SEimpl.v
new file mode 100644
index 00000000..65d16d01
--- /dev/null
+++ b/scheduling/BTL_SEimpl.v
@@ -0,0 +1,1508 @@
+Require Import Coqlib AST Maps.
+Require Import Op Registers.
+Require Import RTL BTL Errors.
+Require Import BTL_SEsimuref BTL_SEtheory OptionMonad.
+Require Import BTL_SEsimplify.
+
+Require Import Impure.ImpHCons.
+Import Notations.
+Import HConsing.
+
+Local Open Scope option_monad_scope.
+Local Open Scope impure.
+
+Import ListNotations.
+Local Open Scope list_scope.
+
+(** Tactics *)
+
+Ltac simplify_SOME x := repeat inversion_SOME x; try_simplify_someHyps.
+
+(** Notations to make lemmas more readable *)
+Notation "'sval_refines' ctx sv1 sv2" := (eval_sval ctx sv1 = eval_sval ctx sv2)
+ (only parsing, at level 0, ctx at next level, sv1 at next level, sv2 at next level): hse.
+
+Local Open Scope hse.
+
+Notation "'list_sval_refines' ctx lsv1 lsv2" := (eval_list_sval ctx lsv1 = eval_list_sval ctx lsv2)
+ (only parsing, at level 0, ctx at next level, lsv1 at next level, lsv2 at next level): hse.
+
+Notation "'smem_refines' ctx sm1 sm2" := (eval_smem ctx sm1 = eval_smem ctx sm2)
+ (only parsing, at level 0, ctx at next level, sm1 at next level, sm2 at next level): hse.
+
+(** Debug printer *)
+Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := RET tt. (* TO REMOVE DEBUG INFO *)
+(*Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := DO s <~ k x;; println ("DEBUG simu_check:" +; s). (* TO INSERT DEBUG INFO *)*)
+
+Definition DEBUG (s: pstring): ?? unit := XDEBUG tt (fun _ => RET s).
+
+(** * Implementation of Data-structure use in Hash-consing *)
+
+Definition sval_get_hid (sv: sval): hashcode :=
+ match sv with
+ | Sundef hid => hid
+ | Sinput _ hid => hid
+ | Sop _ _ hid => hid
+ | Sload _ _ _ _ _ hid => hid
+ end.
+
+Definition list_sval_get_hid (lsv: list_sval): hashcode :=
+ match lsv with
+ | Snil hid => hid
+ | Scons _ _ hid => hid
+ end.
+
+Definition smem_get_hid (sm: smem): hashcode :=
+ match sm with
+ | Sinit hid => hid
+ | Sstore _ _ _ _ _ hid => hid
+ end.
+
+Definition sval_set_hid (sv: sval) (hid: hashcode): sval :=
+ match sv with
+ | Sundef _ => Sundef hid
+ | Sinput r _ => Sinput r hid
+ | Sop o lsv _ => Sop o lsv hid
+ | Sload sm trap chunk addr lsv _ => Sload sm trap chunk addr lsv hid
+ end.
+
+Definition list_sval_set_hid (lsv: list_sval) (hid: hashcode): list_sval :=
+ match lsv with
+ | Snil _ => Snil hid
+ | Scons sv lsv _ => Scons sv lsv hid
+ end.
+
+Definition smem_set_hid (sm: smem) (hid: hashcode): smem :=
+ match sm with
+ | Sinit _ => Sinit hid
+ | Sstore sm chunk addr lsv srce _ => Sstore sm chunk addr lsv srce hid
+ end.
+
+Lemma sval_set_hid_correct ctx x y:
+ sval_set_hid x unknown_hid = sval_set_hid y unknown_hid ->
+ sval_refines ctx x y.
+Proof.
+ destruct x, y; intro H; inversion H; subst; simpl; auto.
+Qed.
+Local Hint Resolve sval_set_hid_correct: core.
+
+Lemma list_sval_set_hid_correct ctx x y:
+ list_sval_set_hid x unknown_hid = list_sval_set_hid y unknown_hid ->
+ list_sval_refines ctx x y.
+Proof.
+ destruct x, y; intro H; inversion H; subst; simpl; auto.
+Qed.
+Local Hint Resolve list_sval_set_hid_correct: core.
+
+Lemma smem_set_hid_correct ctx x y:
+ smem_set_hid x unknown_hid = smem_set_hid y unknown_hid ->
+ smem_refines ctx x y.
+Proof.
+ destruct x, y; intro H; inversion H; subst; simpl; auto.
+Qed.
+Local Hint Resolve smem_set_hid_correct: core.
+
+(** Now, we build the hash-Cons value from a "hash_eq".
+
+ Informal specification:
+ [hash_eq] must be consistent with the "hashed" constructors defined above.
+
+ We expect that hashinfo values in the code of these "hashed" constructors verify:
+ (hash_eq (hdata x) (hdata y) ~> true) <-> (hcodes x)=(hcodes y)
+*)
+
+Definition sval_hash_eq (sv1 sv2: sval): ?? bool :=
+ match sv1, sv2 with
+ | Sinput r1 _, Sinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *)
+ | Sop op1 lsv1 _, Sop op2 lsv2 _ =>
+ DO b1 <~ phys_eq lsv1 lsv2;;
+ if b1
+ then struct_eq op1 op2 (* NB: really need a struct_eq here ? *)
+ else RET false
+ | Sload sm1 trap1 chk1 addr1 lsv1 _, Sload sm2 trap2 chk2 addr2 lsv2 _ =>
+ DO b1 <~ phys_eq lsv1 lsv2;;
+ DO b2 <~ phys_eq sm1 sm2;;
+ DO b3 <~ struct_eq trap1 trap2;;
+ DO b4 <~ struct_eq chk1 chk2;;
+ if b1 && b2 && b3 && b4
+ then struct_eq addr1 addr2
+ else RET false
+ | _,_ => RET false
+ end.
+
+Lemma and_true_split a b: a && b = true <-> a = true /\ b = true.
+Proof.
+ destruct a; simpl; intuition.
+Qed.
+
+Lemma sval_hash_eq_correct x y:
+ WHEN sval_hash_eq x y ~> b THEN
+ b = true -> sval_set_hid x unknown_hid = sval_set_hid y unknown_hid.
+Proof.
+ destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
+Qed.
+Global Opaque sval_hash_eq.
+Local Hint Resolve sval_hash_eq_correct: wlp.
+
+Definition list_sval_hash_eq (lsv1 lsv2: list_sval): ?? bool :=
+ match lsv1, lsv2 with
+ | Snil _, Snil _ => RET true
+ | Scons sv1 lsv1' _, Scons sv2 lsv2' _ =>
+ DO b <~ phys_eq lsv1' lsv2';;
+ if b
+ then phys_eq sv1 sv2
+ else RET false
+ | _,_ => RET false
+ end.
+
+Lemma list_sval_hash_eq_correct x y:
+ WHEN list_sval_hash_eq x y ~> b THEN
+ b = true -> list_sval_set_hid x unknown_hid = list_sval_set_hid y unknown_hid.
+Proof.
+ destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
+Qed.
+Global Opaque list_sval_hash_eq.
+Local Hint Resolve list_sval_hash_eq_correct: wlp.
+
+Definition smem_hash_eq (sm1 sm2: smem): ?? bool :=
+ match sm1, sm2 with
+ | Sinit _, Sinit _ => RET true
+ | Sstore sm1 chk1 addr1 lsv1 sv1 _, Sstore sm2 chk2 addr2 lsv2 sv2 _ =>
+ DO b1 <~ phys_eq lsv1 lsv2;;
+ DO b2 <~ phys_eq sm1 sm2;;
+ DO b3 <~ phys_eq sv1 sv2;;
+ DO b4 <~ struct_eq chk1 chk2;;
+ if b1 && b2 && b3 && b4
+ then struct_eq addr1 addr2
+ else RET false
+ | _,_ => RET false
+ end.
+
+Lemma smem_hash_eq_correct x y:
+ WHEN smem_hash_eq x y ~> b THEN
+ b = true -> smem_set_hid x unknown_hid = smem_set_hid y unknown_hid.
+Proof.
+ destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
+Qed.
+Global Opaque smem_hash_eq.
+Local Hint Resolve smem_hash_eq_correct: wlp.
+
+Definition hSVAL: hashP sval := {| hash_eq := sval_hash_eq; get_hid:=sval_get_hid; set_hid:=sval_set_hid |}.
+Definition hLSVAL: hashP list_sval := {| hash_eq := list_sval_hash_eq; get_hid:= list_sval_get_hid; set_hid:= list_sval_set_hid |}.
+Definition hSMEM: hashP smem := {| hash_eq := smem_hash_eq; get_hid:= smem_get_hid; set_hid:= smem_set_hid |}.
+
+Program Definition mk_hash_params: Dict.hash_params sval :=
+ {|
+ Dict.test_eq := phys_eq;
+ Dict.hashing := fun (sv: sval) => RET (sval_get_hid sv);
+ Dict.log := fun sv =>
+ DO sv_name <~ string_of_hashcode (sval_get_hid sv);;
+ println ("unexpected undef behavior of hashcode:" +; (CamlStr sv_name)) |}.
+Obligation 1.
+ wlp_simplify.
+Qed.
+
+(** * Implementation of symbolic execution *)
+Section CanonBuilding.
+
+Variable hC_sval: hashinfo sval -> ?? sval.
+
+Hypothesis hC_sval_correct: forall s,
+ WHEN hC_sval s ~> s' THEN forall ctx,
+ sval_refines ctx (hdata s) s'.
+
+Variable hC_list_sval: hashinfo list_sval -> ?? list_sval.
+Hypothesis hC_list_sval_correct: forall lh,
+ WHEN hC_list_sval lh ~> lh' THEN forall ctx,
+ list_sval_refines ctx (hdata lh) lh'.
+
+Variable hC_smem: hashinfo smem -> ?? smem.
+Hypothesis hC_smem_correct: forall hm,
+ WHEN hC_smem hm ~> hm' THEN forall ctx,
+ smem_refines ctx (hdata hm) hm'.
+
+(* First, we wrap constructors for hashed values !*)
+
+Definition reg_hcode := 1.
+Definition op_hcode := 2.
+Definition load_hcode := 3.
+Definition undef_code := 4.
+
+Definition hSinput_hcodes (r: reg) :=
+ DO hc <~ hash reg_hcode;;
+ DO sv <~ hash r;;
+ RET [hc;sv].
+Extraction Inline hSinput_hcodes.
+
+Definition hSinput (r:reg): ?? sval :=
+ DO sv <~ hSinput_hcodes r;;
+ hC_sval {| hdata:=Sinput r unknown_hid; hcodes :=sv; |}.
+
+Lemma hSinput_correct r:
+ WHEN hSinput r ~> sv THEN forall ctx,
+ sval_refines ctx sv (Sinput r unknown_hid).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hSinput.
+Local Hint Resolve hSinput_correct: wlp.
+
+Definition hSop_hcodes (op:operation) (lsv: list_sval) :=
+ DO hc <~ hash op_hcode;;
+ DO sv <~ hash op;;
+ RET [hc;sv;list_sval_get_hid lsv].
+Extraction Inline hSop_hcodes.
+
+Definition hSop (op:operation) (lsv: list_sval): ?? sval :=
+ DO sv <~ hSop_hcodes op lsv;;
+ hC_sval {| hdata:=Sop op lsv unknown_hid; hcodes :=sv |}.
+
+Lemma hSop_fSop_correct op lsv:
+ WHEN hSop op lsv ~> sv THEN forall ctx,
+ sval_refines ctx sv (fSop op lsv).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hSop.
+Local Hint Resolve hSop_fSop_correct: wlp_raw.
+
+Lemma hSop_correct op lsv1:
+ WHEN hSop op lsv1 ~> sv THEN forall ctx lsv2
+ (LR: list_sval_refines ctx lsv1 lsv2),
+ sval_refines ctx sv (Sop op lsv2 unknown_hid).
+Proof.
+ wlp_xsimplify ltac:(intuition eauto with wlp wlp_raw).
+ rewrite <- LR. erewrite H; eauto.
+Qed.
+Local Hint Resolve hSop_correct: wlp.
+
+Definition hSload_hcodes (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval):=
+ DO hc <~ hash load_hcode;;
+ DO sv1 <~ hash trap;;
+ DO sv2 <~ hash chunk;;
+ DO sv3 <~ hash addr;;
+ RET [hc; smem_get_hid sm; sv1; sv2; sv3; list_sval_get_hid lsv].
+Extraction Inline hSload_hcodes.
+
+Definition hSload (sm: smem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval): ?? sval :=
+ DO sv <~ hSload_hcodes sm trap chunk addr lsv;;
+ hC_sval {| hdata := Sload sm trap chunk addr lsv unknown_hid; hcodes := sv |}.
+
+Lemma hSload_correct sm1 trap chunk addr lsv1:
+ WHEN hSload sm1 trap chunk addr lsv1 ~> sv THEN forall ctx lsv2 sm2
+ (LR: list_sval_refines ctx lsv1 lsv2)
+ (MR: smem_refines ctx sm1 sm2),
+ sval_refines ctx sv (Sload sm2 trap chunk addr lsv2 unknown_hid).
+Proof.
+ wlp_simplify.
+ rewrite <- LR, <- MR.
+ auto.
+Qed.
+Global Opaque hSload.
+Local Hint Resolve hSload_correct: wlp.
+
+Definition hSundef_hcodes :=
+ DO hc <~ hash undef_code;;
+ RET [hc].
+Extraction Inline hSundef_hcodes.
+
+Definition hSundef : ?? sval :=
+ DO sv <~ hSundef_hcodes;;
+ hC_sval {| hdata:=Sundef unknown_hid; hcodes :=sv; |}.
+
+Lemma hSundef_correct:
+ WHEN hSundef ~> sv THEN forall ctx,
+ sval_refines ctx sv (Sundef unknown_hid).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hSundef.
+Local Hint Resolve hSundef_correct: wlp.
+
+Definition hSnil (_: unit): ?? list_sval :=
+ hC_list_sval {| hdata := Snil unknown_hid; hcodes := nil |}.
+
+Lemma hSnil_correct:
+ WHEN hSnil() ~> sv THEN forall ctx,
+ list_sval_refines ctx sv (Snil unknown_hid).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hSnil.
+Local Hint Resolve hSnil_correct: wlp.
+
+Definition hScons (sv: sval) (lsv: list_sval): ?? list_sval :=
+ hC_list_sval {| hdata := Scons sv lsv unknown_hid; hcodes := [sval_get_hid sv; list_sval_get_hid lsv] |}.
+
+Lemma hScons_correct sv1 lsv1:
+ WHEN hScons sv1 lsv1 ~> lsv1' THEN forall ctx sv2 lsv2
+ (VR: sval_refines ctx sv1 sv2)
+ (LR: list_sval_refines ctx lsv1 lsv2),
+ list_sval_refines ctx lsv1' (Scons sv2 lsv2 unknown_hid).
+Proof.
+ wlp_simplify.
+ rewrite <- VR, <- LR.
+ auto.
+Qed.
+Global Opaque hScons.
+Local Hint Resolve hScons_correct: wlp.
+
+Definition hSinit (_: unit): ?? smem :=
+ hC_smem {| hdata := Sinit unknown_hid; hcodes := nil |}.
+
+Lemma hSinit_correct:
+ WHEN hSinit() ~> hm THEN forall ctx,
+ smem_refines ctx hm (Sinit unknown_hid).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hSinit.
+Local Hint Resolve hSinit_correct: wlp.
+
+Definition hSstore_hcodes (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval):=
+ DO sv1 <~ hash chunk;;
+ DO sv2 <~ hash addr;;
+ RET [smem_get_hid sm; sv1; sv2; list_sval_get_hid lsv; sval_get_hid srce].
+Extraction Inline hSstore_hcodes.
+
+Definition hSstore (sm: smem) (chunk: memory_chunk) (addr: addressing) (lsv: list_sval) (srce: sval): ?? smem :=
+ DO sv <~ hSstore_hcodes sm chunk addr lsv srce;;
+ hC_smem {| hdata := Sstore sm chunk addr lsv srce unknown_hid; hcodes := sv |}.
+
+Lemma hSstore_correct sm1 chunk addr lsv1 sv1:
+ WHEN hSstore sm1 chunk addr lsv1 sv1 ~> sm1' THEN forall ctx lsv2 sm2 sv2
+ (LR: list_sval_refines ctx lsv1 lsv2)
+ (MR: smem_refines ctx sm1 sm2)
+ (VR: sval_refines ctx sv1 sv2),
+ smem_refines ctx sm1' (Sstore sm2 chunk addr lsv2 sv2 unknown_hid).
+Proof.
+ wlp_simplify.
+ rewrite <- LR, <- MR, <- VR.
+ auto.
+Qed.
+Global Opaque hSstore.
+Local Hint Resolve hSstore_correct: wlp.
+
+Definition hrs_sreg_get (hrs: ristate) r: ?? sval :=
+ match PTree.get r hrs with
+ | None => if ris_input_init hrs then hSinput r else hSundef
+ | Some sv => RET sv
+ end.
+
+Lemma hrs_sreg_get_correct hrs r:
+ WHEN hrs_sreg_get hrs r ~> sv THEN forall ctx (f: reg -> sval)
+ (RR: forall r, sval_refines ctx (hrs r) (f r)),
+ sval_refines ctx sv (f r).
+Proof.
+ unfold ris_sreg_get; wlp_simplify; rewrite <- RR; rewrite H; auto;
+ rewrite H0, H1; simpl; auto.
+Qed.
+Global Opaque hrs_sreg_get.
+Local Hint Resolve hrs_sreg_get_correct: wlp.
+
+Fixpoint hlist_args (hrs: ristate) (l: list reg): ?? list_sval :=
+ match l with
+ | nil => hSnil()
+ | r::l =>
+ DO v <~ hrs_sreg_get hrs r;;
+ DO lsv <~ hlist_args hrs l;;
+ hScons v lsv
+ end.
+
+Lemma hlist_args_correct hrs l:
+ WHEN hlist_args hrs l ~> lsv THEN forall ctx (f: reg -> sval)
+ (RR: forall r, sval_refines ctx (hrs r) (f r)),
+ list_sval_refines ctx lsv (list_sval_inj (List.map f l)).
+Proof.
+ induction l; wlp_simplify.
+Qed.
+Global Opaque hlist_args.
+Local Hint Resolve hlist_args_correct: wlp.
+
+(** Convert a "fake" hash-consed term into a "real" hash-consed term *)
+
+Fixpoint fsval_proj sv: ?? sval :=
+ match sv with
+ | Sundef hc =>
+ DO b <~ phys_eq hc unknown_hid;;
+ if b then hSundef else RET sv
+ | Sinput r hc =>
+ DO b <~ phys_eq hc unknown_hid;;
+ if b then hSinput r (* was not yet really hash-consed *)
+ else RET sv (* already hash-consed *)
+ | Sop op hl hc =>
+ DO b <~ phys_eq hc unknown_hid;;
+ if b then (* was not yet really hash-consed *)
+ DO hl' <~ fsval_list_proj hl;;
+ hSop op hl'
+ else RET sv (* already hash-consed *)
+ | Sload hm t chk addr hl _ => RET sv (* FIXME TODO gourdinl ? *)
+ end
+with fsval_list_proj sl: ?? list_sval :=
+ match sl with
+ | Snil hc =>
+ DO b <~ phys_eq hc unknown_hid;;
+ if b then hSnil() (* was not yet really hash-consed *)
+ else RET sl (* already hash-consed *)
+ | Scons sv hl hc =>
+ DO b <~ phys_eq hc unknown_hid;;
+ if b then (* was not yet really hash-consed *)
+ DO sv' <~ fsval_proj sv;;
+ DO hl' <~ fsval_list_proj hl;;
+ hScons sv' hl'
+ else RET sl (* already hash-consed *)
+ end.
+
+Lemma fsval_proj_correct sv:
+ WHEN fsval_proj sv ~> sv' THEN forall ctx,
+ sval_refines ctx sv sv'.
+Proof.
+ induction sv using sval_mut
+ with (P0 := fun lsv =>
+ WHEN fsval_list_proj lsv ~> lsv' THEN forall ctx,
+ eval_list_sval ctx lsv = eval_list_sval ctx lsv')
+ (P1 := fun sm => True); try (wlp_simplify; tauto).
+ - wlp_xsimplify ltac:(intuition eauto with wlp_raw wlp).
+ rewrite H, H0; auto.
+ - wlp_simplify; erewrite H0, H1; eauto.
+Qed.
+Global Opaque fsval_proj.
+Local Hint Resolve fsval_proj_correct: wlp.
+
+Lemma fsval_list_proj_correct lsv:
+ WHEN fsval_list_proj lsv ~> lsv' THEN forall ctx,
+ list_sval_refines ctx lsv lsv'.
+Proof.
+ induction lsv; wlp_simplify.
+ erewrite H0, H1; eauto.
+Qed.
+Global Opaque fsval_list_proj.
+Local Hint Resolve fsval_list_proj_correct: wlp.
+
+(** ** Assignment of memory *)
+
+Definition hrexec_store chunk addr args src hrs: ?? ristate :=
+ DO hargs <~ hlist_args hrs args;;
+ DO hsrc <~ hrs_sreg_get hrs src;;
+ DO hm <~ hSstore hrs chunk addr hargs hsrc;;
+ RET (rset_smem hm hrs).
+
+Lemma hrexec_store_correct chunk addr args src hrs:
+ WHEN hrexec_store chunk addr args src hrs ~> hrs' THEN forall ctx sis
+ (REF: ris_refines ctx hrs sis),
+ ris_refines ctx hrs' (sexec_store chunk addr args src sis).
+Proof.
+ wlp_simplify.
+ eapply rset_mem_correct; simpl; eauto.
+ - intros X; erewrite H1; eauto.
+ rewrite X. simplify_SOME z.
+ - intros X; inversion REF.
+ erewrite H1; eauto.
+Qed.
+Local Hint Resolve hrexec_store_correct: wlp.
+
+(** ** Assignment of registers *)
+
+(** locally new symbolic values during symbolic execution *)
+Inductive root_sval: Type :=
+| Rop (op: operation)
+| Rload (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing)
+.
+
+Definition root_apply (rsv: root_sval) (lr: list reg) (st: sistate): sval :=
+ let lsv := list_sval_inj (List.map (si_sreg st) lr) in
+ let sm := si_smem st in
+ match rsv with
+ | Rop op => fSop op lsv
+ | Rload trap chunk addr => fSload sm trap chunk addr lsv
+ end.
+Coercion root_apply: root_sval >-> Funclass.
+
+Definition root_happly (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval :=
+ DO lsv <~ hlist_args hrs lr;;
+ match rsv with
+ | Rop op => hSop op lsv
+ | Rload trap chunk addr => hSload hrs trap chunk addr lsv
+ end.
+
+Lemma root_happly_correct (rsv: root_sval) lr hrs:
+ WHEN root_happly rsv lr hrs ~> sv THEN forall ctx sis
+ (REF: ris_refines ctx hrs sis)
+ (OK: ris_ok ctx hrs),
+ sval_refines ctx sv (rsv lr sis).
+Proof.
+ unfold root_apply, root_happly; destruct rsv; wlp_simplify; inv REF;
+ erewrite H0, H; eauto.
+Qed.
+Global Opaque root_happly.
+Hint Resolve root_happly_correct: wlp.
+
+Local Open Scope lazy_bool_scope.
+
+(* NB: return [false] if the rsv cannot fail *)
+Definition may_trap (rsv: root_sval) (lr: list reg): bool :=
+ match rsv with
+ | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lr) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *)
+ | Rload TRAP _ _ => true
+ | _ => false
+ end.
+
+Lemma lazy_orb_negb_false (b1 b2:bool):
+ (b1 ||| negb b2) = false <-> (b1 = false /\ b2 = true).
+Proof.
+ unfold negb. repeat autodestruct; simpl; intuition (try congruence).
+Qed.
+
+Lemma eval_list_sval_length ctx (f: reg -> sval) (l:list reg):
+ forall l', eval_list_sval ctx (list_sval_inj (List.map f l)) = Some l' ->
+ Datatypes.length l = Datatypes.length l'.
+Proof.
+ induction l.
+ - simpl. intros. inv H. reflexivity.
+ - simpl. intros. destruct (eval_sval _ _); [|discriminate].
+ destruct (eval_list_sval _ _) eqn:SLS; [|discriminate]. inv H. simpl.
+ erewrite IHl; eauto.
+Qed.
+
+Lemma may_trap_correct ctx (rsv: root_sval) (lr: list reg) st:
+ may_trap rsv lr = false ->
+ eval_list_sval ctx (list_sval_inj (List.map (si_sreg st) lr)) <> None ->
+ eval_smem ctx (si_smem st) <> None ->
+ eval_sval ctx (rsv lr st) <> None.
+Proof.
+ destruct rsv; simpl; try congruence.
+ - rewrite lazy_orb_negb_false. intros (TRAP1 & LEN) OK1 OK2.
+ autodestruct; try congruence. intros.
+ eapply is_trapping_op_sound; eauto.
+ erewrite <- eval_list_sval_length; eauto.
+ apply Nat.eqb_eq in LEN.
+ assumption.
+ - intros X OK1 OK2.
+ repeat autodestruct; try congruence.
+Qed.
+
+(** simplify a symbolic value before assignment to a register *)
+Definition simplify (rsv: root_sval) (lr: list reg) (hrs: ristate): ?? sval :=
+ match rsv with
+ | Rop op =>
+ match is_move_operation op lr with
+ | Some arg => hrs_sreg_get hrs arg (* optimization of Omove *)
+ | None =>
+ match target_op_simplify op lr hrs with
+ | Some fsv => fsval_proj fsv
+ | None =>
+ DO lhsv <~ hlist_args hrs lr;;
+ hSop op lhsv
+ end
+ end
+ | Rload _ chunk addr =>
+ DO lsv <~ hlist_args hrs lr;;
+ hSload hrs NOTRAP chunk addr lsv
+ end.
+
+Lemma simplify_correct rsv lr hrs:
+ WHEN simplify rsv lr hrs ~> sv THEN forall ctx sis
+ (REF: ris_refines ctx hrs sis)
+ (OK0: ris_ok ctx hrs)
+ (OK1: eval_sval ctx (rsv lr sis) <> None),
+ sval_refines ctx sv (rsv lr sis).
+Proof.
+ destruct rsv; simpl; auto.
+ - (* Rop *)
+ destruct (is_move_operation _ _) eqn: Hmove.
+ { wlp_simplify; exploit is_move_operation_correct; eauto.
+ intros (Hop & Hlsv); subst; simpl in *. inv REF.
+ simplify_SOME z; erewrite H; eauto. }
+ destruct (target_op_simplify _ _ _) eqn: Htarget_op_simp; wlp_simplify.
+ { destruct (eval_list_sval _ _) eqn: OKlist; try congruence.
+ rewrite <- H; exploit target_op_simplify_correct; eauto. }
+ inv REF; clear Htarget_op_simp.
+ intuition eauto.
+ - (* Rload *)
+ destruct trap; wlp_simplify; inv REF.
+ + erewrite H0; eauto.
+ erewrite H; [|eapply REG_EQ; eauto].
+ erewrite MEM_EQ; eauto.
+ repeat simplify_SOME z.
+ * destruct (Memory.Mem.loadv _ _ _); try congruence.
+ * rewrite H1 in OK1; congruence.
+ + erewrite H0; eauto.
+Qed.
+Global Opaque simplify.
+Local Hint Resolve simplify_correct: wlp.
+
+Definition red_PTree_set (r: reg) (sv: sval) (hrs: ristate): PTree.t sval :=
+ match (ris_input_init hrs), sv with
+ | true, Sinput r' _ =>
+ if Pos.eq_dec r r'
+ then PTree.remove r' hrs
+ else PTree.set r sv hrs
+ | false, Sundef _ =>
+ PTree.remove r hrs
+ | _, _ => PTree.set r sv hrs
+ end.
+
+Lemma red_PTree_set_correct (r r0:reg) sv (hrs: ristate) ctx:
+ sval_refines ctx ((ris_sreg_set hrs (red_PTree_set r sv hrs)) r0) ((ris_sreg_set hrs (PTree.set r sv hrs)) r0).
+Proof.
+ unfold red_PTree_set, ris_sreg_set, ris_sreg_get; simpl.
+ destruct (ris_input_init hrs) eqn:Hinit, sv; simpl; auto.
+ 1: destruct (Pos.eq_dec r r1); auto; subst;
+ rename r1 into r.
+ all: destruct (Pos.eq_dec r r0); auto;
+ [ subst; rewrite PTree.grs, PTree.gss; simpl; auto |
+ rewrite PTree.gro, PTree.gso; simpl; auto].
+Qed.
+
+Lemma red_PTree_set_refines (r r0:reg) ctx hrs sis sv1 sv2:
+ ris_refines ctx hrs sis ->
+ sval_refines ctx sv1 sv2 ->
+ ris_ok ctx hrs ->
+ sval_refines ctx (ris_sreg_set hrs (red_PTree_set r sv1 hrs) r0) (if Pos.eq_dec r r0 then sv2 else si_sreg sis r0).
+Proof.
+ intros REF SREF OK; rewrite red_PTree_set_correct.
+ unfold ris_sreg_set, ris_sreg_get.
+ destruct (Pos.eq_dec r r0).
+ - subst; simpl. rewrite PTree.gss; simpl; auto.
+ - inv REF; simpl. rewrite PTree.gso; simpl; eauto.
+Qed.
+
+Definition cbranch_expanse (prev: ristate) (cond: condition) (args: list reg): ?? (condition * list_sval) :=
+ match target_cbranch_expanse prev cond args with
+ | Some (cond', vargs) =>
+ DO vargs' <~ fsval_list_proj vargs;;
+ RET (cond', vargs')
+ | None =>
+ DO vargs <~ hlist_args prev args ;;
+ RET (cond, vargs)
+ end.
+
+Lemma cbranch_expanse_correct hrs c l:
+ WHEN cbranch_expanse hrs c l ~> r THEN forall ctx sis
+ (REF : ris_refines ctx hrs sis)
+ (OK: ris_ok ctx hrs),
+ eval_scondition ctx (fst r) (snd r) =
+ eval_scondition ctx c (list_sval_inj (map (si_sreg sis) l)).
+Proof.
+ unfold cbranch_expanse.
+ destruct (target_cbranch_expanse _ _ _) eqn: TARGET; wlp_simplify;
+ unfold eval_scondition; try erewrite <- H; inversion REF; eauto.
+ destruct p as [c' l']; simpl.
+ exploit target_cbranch_expanse_correct; eauto.
+Qed.
+Local Hint Resolve cbranch_expanse_correct: wlp.
+Global Opaque cbranch_expanse.
+
+Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A :=
+ match o with
+ | Some x => RET x
+ | None => FAILWITH msg
+ end.
+
+Definition hrs_init: ?? ristate
+ := DO hm <~ hSinit ();;
+ RET {| ris_smem := hm; ris_input_init := true; ok_rsval := nil; ris_sreg := PTree.empty _ |}.
+
+Lemma hrs_init_correct:
+ WHEN hrs_init ~> hrs THEN
+ forall ctx, ris_refines ctx hrs sis_init.
+Proof.
+ unfold hrs_init, sis_init; wlp_simplify.
+ econstructor; simpl in *; eauto.
+ + split; destruct 1; econstructor; simpl in *;
+ try rewrite H; try congruence; trivial.
+ + destruct 1; simpl in *. unfold ris_sreg_get; simpl.
+ intros; rewrite PTree.gempty; eauto.
+Qed.
+Local Hint Resolve hrs_init_correct: wlp.
+Global Opaque hrs_init.
+
+Definition hrs_sreg_set r lr rsv (hrs: ristate): ?? ristate :=
+ DO ok_lsv <~
+ (if may_trap rsv lr
+ then DO sv <~ root_happly rsv lr hrs;;
+ XDEBUG sv (fun sv => DO sv_name <~ string_of_hashcode (sval_get_hid sv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr sv_name))%string);;
+ RET (sv::(ok_rsval hrs))
+ else RET (ok_rsval hrs));;
+ DO simp <~ simplify rsv lr hrs;;
+ RET {| ris_smem := hrs.(ris_smem);
+ ris_input_init := hrs.(ris_input_init);
+ ok_rsval := ok_lsv;
+ ris_sreg:= red_PTree_set r simp hrs |}.
+
+Lemma ok_hrset_sreg (rsv:root_sval) ctx (st: sistate) r lr:
+ si_ok ctx (set_sreg r (rsv lr st) st)
+ <-> (si_ok ctx st /\ eval_sval ctx (rsv lr st) <> None).
+Proof.
+ unfold set_sreg; simpl; split.
+ - intros. destruct H as [[OK_SV OK_PRE] OK_SMEM OK_SREG]; simpl in *.
+ repeat (split; try tauto).
+ + intros r0; generalize (OK_SREG r0); clear OK_SREG; destruct (Pos.eq_dec r r0); try congruence.
+ + generalize (OK_SREG r); clear OK_SREG; destruct (Pos.eq_dec r r); try congruence.
+ - intros (OK & SEVAL). inv OK.
+ repeat (split; try tauto; eauto).
+ intros r0; destruct (Pos.eq_dec r r0) eqn:Heq; simpl;
+ rewrite Heq; eauto.
+Qed.
+
+Lemma eval_list_sval_inj_not_none ctx st: forall l,
+ (forall r, List.In r l -> eval_sval ctx (si_sreg st r) = None -> False) ->
+ eval_list_sval ctx (list_sval_inj (map (si_sreg st) l)) = None -> False.
+Proof.
+ induction l.
+ - intuition discriminate.
+ - intros ALLR. simpl.
+ inversion_SOME v.
+ + intro SVAL. inversion_SOME lv; [discriminate|].
+ assert (forall r : reg, In r l -> eval_sval ctx (si_sreg st r) = None -> False).
+ { intros r INR. eapply ALLR. right. assumption. }
+ intro SVALLIST. intro. eapply IHl; eauto.
+ + intros. exploit (ALLR a); simpl; eauto.
+Qed.
+
+Lemma hrs_sreg_set_correct r lr rsv hrs:
+ WHEN hrs_sreg_set r lr rsv hrs ~> hrs' THEN forall ctx sis
+ (REF: ris_refines ctx hrs sis),
+ ris_refines ctx hrs' (set_sreg r (rsv lr sis) sis).
+Proof.
+ wlp_simplify; inversion REF.
+ - (* may_trap -> true *)
+ assert (X: si_ok ctx (set_sreg r (rsv lr sis) sis) <->
+ ris_ok ctx {| ris_smem := hrs;
+ ris_input_init := ris_input_init hrs;
+ ok_rsval := exta :: ok_rsval hrs;
+ ris_sreg := red_PTree_set r exta0 hrs |}).
+ {
+ rewrite ok_hrset_sreg, OK_EQUIV.
+ split.
+ + intros (ROK & SEVAL); inv ROK.
+ assert (ROK: ris_ok ctx hrs) by (econstructor; eauto).
+ econstructor; eauto; simpl.
+ intuition (subst; eauto).
+ erewrite H0 in *; eauto.
+ + intros (OK_RMEM & OK_RREG); simpl in *.
+ assert (ROK: ris_ok ctx hrs) by (econstructor; eauto).
+ erewrite <- H0 in *; eauto. }
+ split; auto; rewrite <- X, ok_hrset_sreg.
+ + intuition eauto.
+ + intros (SOK & SEVAL) r0.
+ rewrite ris_sreg_set_access.
+ erewrite red_PTree_set_refines; intuition eauto.
+ - (* may_trap -> false *)
+ assert (X: si_ok ctx (set_sreg r (rsv lr sis) sis) <->
+ ris_ok ctx {| ris_smem := hrs;
+ ris_input_init := ris_input_init hrs;
+ ok_rsval := ok_rsval hrs;
+ ris_sreg := red_PTree_set r exta hrs |}).
+ {
+ rewrite ok_hrset_sreg, OK_EQUIV.
+ split.
+ + intros (ROK & SEVAL); inv ROK.
+ econstructor; eauto.
+ + intros (OK_RMEM & OK_RREG).
+ assert (ROK: ris_ok ctx hrs) by (econstructor; eauto).
+ split; auto.
+ intros SNONE; exploit may_trap_correct; eauto.
+ * intros LNONE; eapply eval_list_sval_inj_not_none; eauto.
+ assert (SOK: si_ok ctx sis) by intuition.
+ inv SOK. intuition eauto.
+ * rewrite <- MEM_EQ; auto. }
+ split; auto; rewrite <- X, ok_hrset_sreg.
+ + intuition eauto.
+ + intros (SOK & SEVAL) r0.
+ rewrite ris_sreg_set_access.
+ erewrite red_PTree_set_refines; intuition eauto.
+Qed.
+
+Lemma hrs_ok_op_okpreserv ctx op args dest hrs:
+ WHEN hrs_sreg_set dest args (Rop op) hrs ~> rst THEN
+ ris_ok ctx rst -> ris_ok ctx hrs.
+Proof.
+ wlp_simplify; inv H2 || inv H1;
+ simpl in *; econstructor; eauto.
+Qed.
+
+Lemma hrs_ok_load_okpreserv ctx chunk addr args dest hrs trap:
+ WHEN hrs_sreg_set dest args (Rload trap chunk addr) hrs ~> rst THEN
+ ris_ok ctx rst -> ris_ok ctx hrs.
+Proof.
+ wlp_simplify; inv H2 || inv H1;
+ simpl in *; econstructor; eauto.
+Qed.
+
+Lemma hrs_ok_store_okpreserv ctx chunk addr args src hrs:
+ WHEN hrexec_store chunk addr args src hrs ~> rst THEN
+ ris_ok ctx rst -> ris_ok ctx hrs.
+Proof.
+ unfold hrexec_store; wlp_simplify.
+ generalize H2. erewrite ok_rset_mem; intuition.
+ specialize H1 with (sm2 := hrs).
+ erewrite H1; eauto. rewrite H3. repeat autodestruct.
+Qed.
+
+Global Opaque hrs_sreg_set hrexec_store.
+Local Hint Resolve hrs_sreg_set_correct: wlp.
+
+(* transfer *)
+
+Lemma ok_hrset_red_sreg ctx r rsv hrs:
+ ris_ok ctx (ris_sreg_set hrs (red_PTree_set r rsv hrs))
+ <-> (ris_ok ctx hrs).
+Proof.
+ split; destruct 1; econstructor; simpl in *; eauto.
+Qed.
+
+Fixpoint transfer_hrs (inputs: list reg) (hrs: ristate): ?? ristate :=
+ match inputs with
+ | nil => RET {| ris_smem := hrs.(ris_smem);
+ ris_input_init := false;
+ ok_rsval := hrs.(ok_rsval);
+ ris_sreg:= PTree.empty _
+ |}
+ | r::l =>
+ DO hrs' <~ transfer_hrs l hrs;;
+ DO sv <~ hrs_sreg_get hrs r;;
+ RET (ris_sreg_set hrs' (red_PTree_set r sv hrs'))
+ end.
+
+Lemma ok_transfer_hrs ctx inputs hrs:
+ WHEN transfer_hrs inputs hrs ~> hrs' THEN
+ ris_ok ctx hrs'
+ <-> (ris_ok ctx hrs).
+Proof.
+ induction inputs as [|r l]; simpl; wlp_simplify;
+ try (inv H; econstructor; eauto; fail).
+ + apply H0; rewrite <- ok_hrset_red_sreg; eauto.
+ + rewrite ok_hrset_red_sreg; auto.
+Qed.
+Local Hint Resolve ok_transfer_hrs: wlp.
+
+Lemma transfer_hrs_correct ctx inputs hrs sis:
+ WHEN transfer_hrs inputs hrs ~> hrs' THEN
+ ris_refines ctx hrs sis ->
+ ris_refines ctx hrs' (transfer_sis inputs sis).
+Proof.
+ unfold transfer_hrs;
+ induction inputs as [|r l]; wlp_simplify.
+ inversion H.
+ + constructor.
+ * erewrite ok_transfer_sis; split; intros X.
+ - constructor; simpl; rewrite OK_EQUIV in X; inv X; assumption.
+ - rewrite OK_EQUIV; constructor; inv X; assumption.
+ * intros X; inv X; simpl; apply MEM_EQ; constructor; auto.
+ * simpl. unfold ris_sreg_get.
+ intros;rewrite PTree.gempty. simpl; auto.
+ + constructor.
+ * erewrite ok_hrset_red_sreg; eauto.
+ inv H2. rewrite <- OK_EQUIV.
+ erewrite !ok_transfer_sis; reflexivity.
+ * intros X; inversion X; simpl in *.
+ rewrite ok_hrset_red_sreg in X.
+ inv H2; rewrite MEM_EQ; auto.
+ * simpl; intros X r0. inversion H2.
+ rewrite ok_hrset_red_sreg in X.
+ erewrite red_PTree_set_refines; eauto.
+ destruct (Pos.eq_dec); auto.
+ inv H1. rewrite ok_transfer_sis in OK_EQUIV.
+ rewrite <- OK_EQUIV, OK_EQUIV0 in X.
+ erewrite H0; eauto.
+Qed.
+
+Definition tr_hrs := poly_tr (fun f tbl or => transfer_hrs (Regset.elements (pre_inputs f tbl or))).
+
+Local Hint Resolve transfer_hrs_correct ok_transfer_hrs: core.
+Local Opaque transfer_hrs.
+
+Lemma ok_tr_hrs ctx fi hrs:
+ WHEN tr_hrs (cf ctx) fi hrs ~> hrs' THEN
+ ris_ok ctx hrs'
+ <-> (ris_ok ctx hrs).
+Proof.
+ destruct fi; simpl; eauto.
+Qed.
+
+Lemma tr_hrs_correct ctx fi hrs sis:
+ WHEN tr_hrs (cf ctx) fi hrs ~> hrs' THEN
+ ris_refines ctx hrs sis ->
+ ris_refines ctx hrs' (tr_sis (cf ctx) fi sis).
+Proof.
+ unfold tr_hrs, poly_tr; destruct fi; wlp_simplify;
+ exploit transfer_hrs_correct; eauto.
+ Unshelve. all: auto.
+Qed.
+
+Fixpoint hbuiltin_arg (hrs: ristate) (arg : builtin_arg reg): ?? builtin_arg sval :=
+ match arg with
+ | BA r =>
+ DO v <~ hrs_sreg_get hrs r;;
+ RET (BA v)
+ | BA_int n => RET (BA_int n)
+ | BA_long n => RET (BA_long n)
+ | BA_float f0 => RET (BA_float f0)
+ | BA_single s => RET (BA_single s)
+ | BA_loadstack chunk ptr => RET (BA_loadstack chunk ptr)
+ | BA_addrstack ptr => RET (BA_addrstack ptr)
+ | BA_loadglobal chunk id ptr => RET (BA_loadglobal chunk id ptr)
+ | BA_addrglobal id ptr => RET (BA_addrglobal id ptr)
+ | BA_splitlong ba1 ba2 =>
+ DO v1 <~ hbuiltin_arg hrs ba1;;
+ DO v2 <~ hbuiltin_arg hrs ba2;;
+ RET (BA_splitlong v1 v2)
+ | BA_addptr ba1 ba2 =>
+ DO v1 <~ hbuiltin_arg hrs ba1;;
+ DO v2 <~ hbuiltin_arg hrs ba2;;
+ RET (BA_addptr v1 v2)
+ end.
+
+Lemma hbuiltin_arg_correct hrs arg:
+ WHEN hbuiltin_arg hrs arg ~> hargs THEN forall ctx (f: reg -> sval)
+ (RR: forall r, sval_refines ctx (hrs r) (f r)),
+ eval_builtin_sval ctx hargs = eval_builtin_sval ctx (builtin_arg_map f arg).
+Proof.
+ induction arg; wlp_simplify.
+ + erewrite H; eauto.
+ + erewrite H; eauto.
+ erewrite H0; eauto.
+ + erewrite H; eauto.
+ erewrite H0; eauto.
+Qed.
+Global Opaque hbuiltin_arg.
+Local Hint Resolve hbuiltin_arg_correct: wlp.
+
+Fixpoint hbuiltin_args hrs (args: list (builtin_arg reg)): ?? list (builtin_arg sval) :=
+ match args with
+ | nil => RET nil
+ | a::l =>
+ DO ha <~ hbuiltin_arg hrs a;;
+ DO hl <~ hbuiltin_args hrs l;;
+ RET (ha::hl)
+ end.
+
+Lemma hbuiltin_args_correct hrs args:
+ WHEN hbuiltin_args hrs args ~> hargs THEN forall ctx (f: reg -> sval)
+ (RR: forall r, sval_refines ctx (hrs r) (f r)),
+ bargs_refines ctx hargs (List.map (builtin_arg_map f) args).
+Proof.
+ unfold bargs_refines; induction args; wlp_simplify.
+ erewrite H; eauto.
+ erewrite H0; eauto.
+Qed.
+Global Opaque hbuiltin_args.
+Local Hint Resolve hbuiltin_args_correct: wlp.
+
+Definition hsum_left hrs (ros: reg + ident): ?? (sval + ident) :=
+ match ros with
+ | inl r => DO hr <~ hrs_sreg_get hrs r;; RET (inl hr)
+ | inr s => RET (inr s)
+ end.
+
+Lemma hsum_left_correct hrs ros:
+ WHEN hsum_left hrs ros ~> hsi THEN forall ctx (f: reg -> sval)
+ (RR: forall r, sval_refines ctx (hrs r) (f r)),
+ rsvident_refines ctx hsi (sum_left_map f ros).
+Proof.
+ destruct ros; wlp_simplify;
+ econstructor; eauto.
+Qed.
+Global Opaque hsum_left.
+Local Hint Resolve hsum_left_correct: wlp.
+
+(** * Symbolic execution of final step *)
+Definition hrexec_final_sfv (i: final) hrs: ?? sfval :=
+ match i with
+ | Bgoto pc => RET (Sgoto pc)
+ | Bcall sig ros args res pc =>
+ DO svos <~ hsum_left hrs ros;;
+ DO sargs <~ hlist_args hrs args;;
+ RET (Scall sig svos sargs res pc)
+ | Btailcall sig ros args =>
+ DO svos <~ hsum_left hrs ros;;
+ DO sargs <~ hlist_args hrs args;;
+ RET (Stailcall sig svos sargs)
+ | Bbuiltin ef args res pc =>
+ DO sargs <~ hbuiltin_args hrs args;;
+ RET (Sbuiltin ef sargs res pc)
+ | Breturn or =>
+ match or with
+ | Some r => DO hr <~ hrs_sreg_get hrs r;; RET (Sreturn (Some hr))
+ | None => RET (Sreturn None)
+ end
+ | Bjumptable reg tbl =>
+ DO sv <~ hrs_sreg_get hrs reg;;
+ RET (Sjumptable sv tbl)
+ end.
+
+Lemma hrexec_final_sfv_correct ctx fi hrs sis:
+ WHEN hrexec_final_sfv fi hrs ~> hrs' THEN forall
+ (REF: ris_refines ctx hrs sis)
+ (OK: ris_ok ctx hrs),
+ rfv_refines ctx hrs' (sexec_final_sfv fi sis).
+Proof.
+ destruct fi; simpl; wlp_simplify; repeat econstructor;
+ try inversion REF; try erewrite H; eauto.
+Qed.
+
+Fixpoint hrexec_rec f ib hrs (k: ristate -> ?? rstate): ?? rstate :=
+ match ib with
+ | BF fin _ =>
+ DO sfv <~ hrexec_final_sfv fin hrs;;
+ DO hrs' <~ tr_hrs f fin hrs;;
+ RET (Rfinal hrs' sfv)
+ (* basic instructions *)
+ | Bnop _ => k hrs
+ | Bop op args dst _ =>
+ DO next <~ hrs_sreg_set dst args (Rop op) hrs;;
+ k next
+ | Bload trap chunk addr args dst _ =>
+ DO next <~ hrs_sreg_set dst args (Rload trap chunk addr) hrs;;
+ k next
+ | Bstore chunk addr args src _ =>
+ DO next <~ hrexec_store chunk addr args src hrs;;
+ k next
+ (* composed instructions *)
+ | Bseq ib1 ib2 =>
+ hrexec_rec f ib1 hrs (fun hrs2 => hrexec_rec f ib2 hrs2 k)
+ | Bcond cond args ifso ifnot _ =>
+ DO res <~ cbranch_expanse hrs cond args;;
+ let (cond, vargs) := res in
+ DO ifso <~ hrexec_rec f ifso hrs k;;
+ DO ifnot <~ hrexec_rec f ifnot hrs k;;
+ RET (Rcond cond vargs ifso ifnot)
+ end
+ .
+
+Definition hrexec f ib :=
+ DO init <~ hrs_init;;
+ hrexec_rec f ib init (fun _ => RET Rabort).
+
+Local Hint Resolve exec_final_refpreserv ok_tr_hrs: core.
+Local Hint Constructors rst_refines: core.
+
+Lemma hrexec_rec_correct1 ctx ib:
+ forall rk k
+ (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis)
+ (CONT: forall hrs sis lsis sfv st,
+ ris_refines ctx hrs sis ->
+ k sis = st ->
+ get_soutcome ctx (k sis) = Some (sout lsis sfv) ->
+ si_ok ctx lsis ->
+ WHEN rk hrs ~> rst THEN
+ rst_refines ctx rst (k sis))
+ hrs sis lsis sfv st
+ (REF: ris_refines ctx hrs sis)
+ (EXEC: sexec_rec (cf ctx) ib sis k = st)
+ (SOUT: get_soutcome ctx st = Some (sout lsis sfv))
+ (OK: si_ok ctx lsis),
+ WHEN hrexec_rec (cf ctx) ib hrs rk ~> rst THEN
+ rst_refines ctx rst st.
+Proof.
+ induction ib; try (wlp_simplify; fail).
+ - (* BF *)
+ wlp_simplify; exploit tr_hrs_correct; eauto.
+ intros; constructor; auto.
+ intros X; apply H0 in X.
+ exploit hrexec_final_sfv_correct; eauto.
+ - (* Bnop *) wlp_simplify; eapply CONT; eauto.
+ - (* Bseq *)
+ wlp_simplify.
+ eapply IHib1. 3-7: eauto.
+ + simpl. eapply sexec_rec_okpreserv; eauto.
+ + intros; eapply IHib2; eauto.
+ - (* Bcond *)
+ simpl; intros; wlp_simplify.
+ assert (rOK: ris_ok ctx hrs). {
+ erewrite <- OK_EQUIV. 2: eauto.
+ eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto.
+ }
+ exploit (Refcond ctx (fst exta) cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1
+ (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k));
+ exploit H; eauto; intros SEVAL; auto.
+ all:
+ simpl in SOUT; generalize SOUT; clear SOUT;
+ inversion_SOME b0; try_simplify_someHyps.
+ + intros; eapply IHib1; eauto.
+ + intros; eapply IHib2; eauto.
+Qed.
+
+Lemma hrexec_correct1 ctx ib sis sfv:
+ get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) ->
+ (si_ok ctx sis) ->
+ WHEN hrexec (cf ctx) ib ~> rst THEN
+ rst_refines ctx rst (sexec (cf ctx) ib).
+Proof.
+ unfold sexec; intros; wlp_simplify.
+ eapply hrexec_rec_correct1; eauto; simpl; congruence.
+Qed.
+
+Lemma hrexec_rec_okpreserv ctx ib:
+ forall k
+ (CONT: forall hrs lhrs rfv,
+ WHEN k hrs ~> rst THEN
+ get_routcome ctx rst = Some (rout lhrs rfv) ->
+ ris_ok ctx lhrs ->
+ ris_ok ctx hrs)
+ hrs lhrs rfv,
+ WHEN hrexec_rec (cf ctx) ib hrs k ~> rst THEN
+ get_routcome ctx rst = Some (rout lhrs rfv) ->
+ ris_ok ctx lhrs ->
+ ris_ok ctx hrs.
+Proof.
+ induction ib; simpl; try (wlp_simplify; try_simplify_someHyps; fail).
+ - (* Bop *)
+ wlp_simplify; exploit hrs_ok_op_okpreserv; eauto.
+ - (* Bload *)
+ destruct trap; wlp_simplify; try_simplify_someHyps;
+ exploit hrs_ok_load_okpreserv; eauto.
+ - (* Bstore *)
+ wlp_simplify; exploit hrs_ok_store_okpreserv; eauto.
+ - (* Bcond *)
+ wlp_simplify; generalize H2; clear H2; inversion_SOME b; destruct b; try_simplify_someHyps.
+Qed.
+Local Hint Resolve hrexec_rec_okpreserv: wlp.
+
+Lemma hrexec_rec_correct2 ctx ib:
+ forall rk k
+ (CONTh: forall hrs lhrs rfv,
+ WHEN rk hrs ~> rst THEN
+ get_routcome ctx rst = Some (rout lhrs rfv) ->
+ ris_ok ctx lhrs ->
+ ris_ok ctx hrs)
+ (CONT: forall hrs sis lhrs rfv,
+ ris_refines ctx hrs sis ->
+ WHEN rk hrs ~> rst THEN
+ get_routcome ctx rst = Some (rout lhrs rfv) ->
+ ris_ok ctx lhrs ->
+ rst_refines ctx rst (k sis))
+ hrs sis lhrs rfv
+ (REF: ris_refines ctx hrs sis),
+ WHEN hrexec_rec (cf ctx) ib hrs rk ~> rst THEN
+ get_routcome ctx rst = Some (rout lhrs rfv) ->
+ ris_ok ctx lhrs ->
+ rst_refines ctx rst (sexec_rec (cf ctx) ib sis k).
+Proof.
+ induction ib; try (wlp_simplify; fail).
+ - (* BF *)
+ wlp_simplify; exploit tr_hrs_correct; eauto.
+ intros; constructor; auto.
+ intros X; apply H0 in X.
+ exploit hrexec_final_sfv_correct; eauto.
+ - (* Bnop *) wlp_simplify; eapply CONT; eauto.
+ - (* Bseq *)
+ wlp_simplify.
+ eapply IHib1. 3-6: eauto.
+ + simpl; eapply hrexec_rec_okpreserv; eauto.
+ + intros; eapply IHib2; eauto.
+ - (* Bcond *)
+ simpl; intros; wlp_simplify.
+ assert (rOK: ris_ok ctx hrs). {
+ simpl in H2; generalize H2; inversion_SOME b0; destruct b0; try_simplify_someHyps.
+ }
+ exploit (Refcond ctx (fst exta) cond (snd exta) (list_sval_inj (map sis args)) exta0 exta1
+ (sexec_rec (cf ctx) ib1 sis k) (sexec_rec (cf ctx) ib2 sis k));
+ exploit H; eauto; intros SEVAL; auto.
+ all:
+ simpl in H2; generalize H2; clear H2;
+ inversion_SOME b0; try_simplify_someHyps.
+ + intros; eapply IHib1; eauto.
+ + intros; eapply IHib2; eauto.
+ Unshelve. all: auto.
+Qed.
+
+Lemma hrexec_correct2 ctx ib hrs rfv:
+ WHEN hrexec (cf ctx) ib ~> rst THEN
+ get_routcome ctx rst = Some (rout hrs rfv) ->
+ (ris_ok ctx hrs) ->
+ rst_refines ctx rst (sexec (cf ctx) ib).
+Proof.
+ unfold hrexec; intros; wlp_simplify.
+ eapply hrexec_rec_correct2; eauto; simpl;
+ wlp_simplify; congruence.
+Qed.
+
+Global Opaque hrexec.
+
+End CanonBuilding.
+
+(** * Implementing the simulation test with concrete hash-consed symbolic execution *)
+
+Definition phys_check {A} (x y:A) (msg: pstring): ?? unit :=
+ DO b <~ phys_eq x y;;
+ assert_b b msg;;
+ RET tt.
+
+Definition struct_check {A} (x y: A) (msg: pstring): ?? unit :=
+ DO b <~ struct_eq x y;;
+ assert_b b msg;;
+ RET tt.
+
+Lemma struct_check_correct {A} (a b: A) msg:
+ WHEN struct_check a b msg ~> _ THEN
+ a = b.
+Proof. wlp_simplify. Qed.
+Global Opaque struct_check.
+Hint Resolve struct_check_correct: wlp.
+
+Definition 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.
+
+Lemma option_eq_check_correct A (o1 o2: option A): WHEN option_eq_check o1 o2 ~> _ THEN o1=o2.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque option_eq_check.
+Hint Resolve option_eq_check_correct:wlp.
+
+Import PTree.
+
+Fixpoint PTree_eq_check {A} (d1 d2: PTree.t A): ?? unit :=
+ match d1, d2 with
+ | Leaf, Leaf => RET tt
+ | Node l1 o1 r1, Node l2 o2 r2 =>
+ option_eq_check o1 o2;;
+ PTree_eq_check l1 l2;;
+ PTree_eq_check r1 r2
+ | _, _ => FAILWITH "PTree_eq_check: some key is absent"
+ end.
+
+Lemma PTree_eq_check_correct A d1: forall (d2: t A),
+ WHEN PTree_eq_check d1 d2 ~> _ THEN forall x, PTree.get x d1 = PTree.get x d2.
+Proof.
+ induction d1 as [|l1 Hl1 o1 r1 Hr1]; destruct d2 as [|l2 o2 r2]; simpl;
+ wlp_simplify. destruct x; simpl; auto.
+Qed.
+Global Opaque PTree_eq_check.
+Local Hint Resolve PTree_eq_check_correct: wlp.
+
+Definition hrs_simu_check (hrs1 hrs2: ristate) : ?? unit :=
+ DEBUG("? hrs_simu_check");;
+ phys_check (ris_smem hrs1) (ris_smem hrs2) "hrs_simu_check: ris_smem sets aren't equiv";;
+ phys_check (ris_input_init hrs1) (ris_input_init hrs2) "hrs_simu_check: ris_input_init bools aren't equiv";;
+ PTree_eq_check (ris_sreg hrs1) (ris_sreg hrs2);;
+ Sets.assert_list_incl mk_hash_params (ok_rsval hrs2) (ok_rsval hrs1);;
+ DEBUG("=> hrs_simu_check: OK").
+
+Lemma setoid_in {A: Type} (a: A): forall l,
+ SetoidList.InA (fun x y => x = y) a l ->
+ In a l.
+Proof.
+ induction l; intros; inv H.
+ - constructor. reflexivity.
+ - right. auto.
+Qed.
+
+Lemma regset_elements_in r rs:
+ Regset.In r rs ->
+ In r (Regset.elements rs).
+Proof.
+ intros. exploit Regset.elements_1; eauto. intro SIN.
+ apply setoid_in. assumption.
+Qed.
+Local Hint Resolve regset_elements_in: core.
+
+Lemma hrs_simu_check_correct hrs1 hrs2:
+ WHEN hrs_simu_check hrs1 hrs2 ~> _ THEN
+ ris_simu hrs1 hrs2.
+Proof.
+ wlp_simplify; constructor; auto.
+ unfold ris_sreg_get; intros r; rewrite H, H1; reflexivity.
+Qed.
+Hint Resolve hrs_simu_check_correct: wlp.
+Global Opaque hrs_simu_check.
+
+Definition svos_simu_check (svos1 svos2: sval + ident) :=
+ match svos1, svos2 with
+ | inl sv1, inl sv2 => phys_check sv1 sv2 "svos_simu_check: sval mismatch"
+ | inr id1, inr id2 => phys_check id1 id2 "svos_simu_check: symbol mismatch"
+ | _, _ => FAILWITH "svos_simu_check: type mismatch"
+ end.
+
+Lemma svos_simu_check_correct svos1 svos2:
+ WHEN svos_simu_check svos1 svos2 ~> _ THEN
+ svos1 = svos2.
+Proof.
+ destruct svos1; destruct svos2; wlp_simplify.
+Qed.
+Global Opaque svos_simu_check.
+Hint Resolve svos_simu_check_correct: wlp.
+
+Fixpoint builtin_arg_simu_check (bs bs': builtin_arg sval) :=
+ match bs with
+ | BA sv =>
+ match bs' with
+ | BA sv' => phys_check sv sv' "builtin_arg_simu_check: sval mismatch"
+ | _ => FAILWITH "builtin_arg_simu_check: BA mismatch"
+ end
+ | BA_splitlong lo hi =>
+ match bs' with
+ | BA_splitlong lo' hi' =>
+ builtin_arg_simu_check lo lo';;
+ builtin_arg_simu_check hi hi'
+ | _ => FAILWITH "builtin_arg_simu_check: BA_splitlong mismatch"
+ end
+ | BA_addptr b1 b2 =>
+ match bs' with
+ | BA_addptr b1' b2' =>
+ builtin_arg_simu_check b1 b1';;
+ builtin_arg_simu_check b2 b2'
+ | _ => FAILWITH "builtin_arg_simu_check: BA_addptr mismatch"
+ end
+ | bs => struct_check bs bs' "builtin_arg_simu_check: basic mismatch"
+ end.
+
+Lemma builtin_arg_simu_check_correct: forall bs1 bs2,
+ WHEN builtin_arg_simu_check bs1 bs2 ~> _ THEN
+ bs1 = bs2.
+Proof.
+ induction bs1.
+ all: try (wlp_simplify; subst; reflexivity).
+ all: destruct bs2; wlp_simplify; congruence.
+Qed.
+Global Opaque builtin_arg_simu_check.
+Hint Resolve builtin_arg_simu_check_correct: wlp.
+
+Fixpoint list_builtin_arg_simu_check lbs1 lbs2 :=
+ match lbs1, lbs2 with
+ | nil, nil => RET tt
+ | bs1::lbs1, bs2::lbs2 =>
+ builtin_arg_simu_check bs1 bs2;;
+ list_builtin_arg_simu_check lbs1 lbs2
+ | _, _ => FAILWITH "list_builtin_arg_simu_check: length mismatch"
+ end.
+
+Lemma list_builtin_arg_simu_check_correct: forall lbs1 lbs2,
+ WHEN list_builtin_arg_simu_check lbs1 lbs2 ~> _ THEN
+ lbs1 = lbs2.
+Proof.
+ induction lbs1; destruct lbs2; wlp_simplify. congruence.
+Qed.
+Global Opaque list_builtin_arg_simu_check.
+Hint Resolve list_builtin_arg_simu_check_correct: wlp.
+
+Definition sfval_simu_check (sfv1 sfv2: sfval): ?? unit :=
+ match sfv1, sfv2 with
+ | Sgoto e1, Sgoto e2 =>
+ phys_check e1 e2 "sfval_simu_check: Sgoto successors do not match"
+ | Scall sig1 svos1 lsv1 res1 e1, Scall sig2 svos2 lsv2 res2 e2 =>
+ phys_check e1 e2 "sfval_simu_check: Scall successors do not match";;
+ phys_check sig1 sig2 "sfval_simu_check: Scall different signatures";;
+ phys_check res1 res2 "sfval_simu_check: Scall res do not match";;
+ svos_simu_check svos1 svos2;;
+ phys_check lsv1 lsv2 "sfval_simu_check: Scall args do not match"
+ | Stailcall sig1 svos1 lsv1, Stailcall sig2 svos2 lsv2 =>
+ phys_check sig1 sig2 "sfval_simu_check: Stailcall different signatures";;
+ svos_simu_check svos1 svos2;;
+ phys_check lsv1 lsv2 "sfval_simu_check: Stailcall args do not match"
+ | Sbuiltin ef1 lbs1 br1 e1, Sbuiltin ef2 lbs2 br2 e2 =>
+ phys_check e1 e2 "sfval_simu_check: Sbuiltin successors do not match";;
+ phys_check ef1 ef2 "sfval_simu_check: Sbuiltin ef do not match";;
+ phys_check br1 br2 "sfval_simu_check: Sbuiltin br do not match";;
+ list_builtin_arg_simu_check lbs1 lbs2
+ | Sjumptable sv1 le1, Sjumptable sv2 le2 =>
+ phys_check le1 le2 "sfval_simu_check: Sjumptable successors do not match";;
+ phys_check sv1 sv2 "sfval_simu_check: Sjumptable sval do not match"
+ | Sreturn osv1, Sreturn osv2 =>
+ option_eq_check osv1 osv2
+ | _, _ => FAILWITH "sfval_simu_check: structure mismatch"
+ end.
+
+Lemma sfval_simu_check_correct sfv1 sfv2:
+ WHEN sfval_simu_check sfv1 sfv2 ~> _ THEN
+ rfv_simu sfv1 sfv2.
+Proof.
+ destruct sfv1; destruct sfv2; simpl; wlp_simplify; try congruence.
+Qed.
+Hint Resolve sfval_simu_check_correct: wlp.
+Global Opaque sfval_simu_check.
+
+Fixpoint rst_simu_check (rst1 rst2: rstate) {struct rst1} :=
+ match rst1, rst2 with
+ | Rfinal hrs1 sfv1, Rfinal hrs2 sfv2 =>
+ hrs_simu_check hrs1 hrs2;;
+ sfval_simu_check sfv1 sfv2
+ | Rcond cond1 lsv1 rsL1 rsR1, Rcond cond2 lsv2 rsL2 rsR2 =>
+ struct_check cond1 cond2 "rst_simu_check: conditions do not match";;
+ phys_check lsv1 lsv2 "rst_simu_check: args do not match";;
+ rst_simu_check rsL1 rsL2;;
+ rst_simu_check rsR1 rsR2
+ | _, _ => FAILWITH "rst_simu_check: simu_check failed"
+ end.
+
+Lemma rst_simu_check_correct rst1: forall rst2,
+ WHEN rst_simu_check rst1 rst2 ~> _ THEN
+ rst_simu rst1 rst2.
+Proof.
+ induction rst1; destruct rst2;
+ wlp_simplify; constructor; auto.
+Qed.
+Hint Resolve rst_simu_check_correct: wlp.
+Global Opaque rst_simu_check.
+
+Definition simu_check_single (f1 f2: function) (ibf1 ibf2: iblock_info): ?? unit :=
+ (* creating the hash-consing tables *)
+ DO hC_sval <~ hCons hSVAL;;
+ DO hC_list_hsval <~ hCons hLSVAL;;
+ DO hC_hsmem <~ hCons hSMEM;;
+ let hrexec := hrexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in
+ (* performing the hash-consed executions *)
+ DO hst1 <~ hrexec f1 ibf1.(entry);;
+ DO hst2 <~ hrexec f2 ibf2.(entry);;
+ (* comparing the executions *)
+ rst_simu_check hst1 hst2.
+
+Lemma simu_check_single_correct (f1 f2: function) (ibf1 ibf2: iblock_info):
+ WHEN simu_check_single f1 f2 ibf1 ibf2 ~> _ THEN
+ symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry).
+Proof.
+ unfold symbolic_simu; wlp_simplify.
+ eapply rst_simu_correct; eauto.
+ + intros; eapply hrexec_correct1; eauto; wlp_simplify.
+ + intros; eapply hrexec_correct2; eauto; wlp_simplify.
+Qed.
+Global Opaque simu_check_single.
+Global Hint Resolve simu_check_single_correct: wlp.
+
+Program Definition aux_simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): ?? bool :=
+ DO r <~
+ (TRY
+ simu_check_single f1 f2 ibf1 ibf2;;
+ RET true
+ CATCH_FAIL s, _ =>
+ println ("simu_check_failure:" +; s);;
+ RET false
+ ENSURE (fun b => b=true -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry)));;
+ RET (`r).
+Obligation 1.
+ split; wlp_simplify. discriminate.
+Qed.
+
+Lemma aux_simu_check_correct (f1 f2: function) (ibf1 ibf2: iblock_info):
+ WHEN aux_simu_check f1 f2 ibf1 ibf2 ~> b THEN
+ b=true -> symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry).
+Proof.
+ unfold aux_simu_check; wlp_simplify.
+ destruct exta; simpl; auto.
+Qed.
+
+(* Coerce aux_simu_check into a pure function (this is a little unsafe like all oracles in CompCert). *)
+
+Import UnsafeImpure.
+
+Definition simu_check (f1 f2: function) (ibf1 ibf2: iblock_info): res unit :=
+ match unsafe_coerce (aux_simu_check f1 f2 ibf1 ibf2) with
+ | Some true => OK tt
+ | _ => Error (msg "simu_check has failed")
+ end.
+
+Lemma simu_check_correct f1 f2 ibf1 ibf2:
+ simu_check f1 f2 ibf1 ibf2 = OK tt ->
+ symbolic_simu f1 f2 ibf1.(entry) ibf2.(entry).
+Proof.
+ unfold simu_check.
+ destruct (unsafe_coerce (aux_simu_check f1 f2 ibf1 ibf2)) as [[|]|] eqn:Hres; simpl; try discriminate.
+ intros; eapply aux_simu_check_correct; eauto.
+ eapply unsafe_coerce_not_really_correct; eauto.
+Qed.
diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v
new file mode 100644
index 00000000..d47e53b8
--- /dev/null
+++ b/scheduling/BTL_SEsimuref.v
@@ -0,0 +1,852 @@
+(** Refinement of BTL_SEtheory data-structures
+ in order to introduce (and prove correct) a lower-level specification of the simulation test.
+*)
+
+Require Import Coqlib Maps Floats.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import Op Registers.
+Require Import RTL BTL OptionMonad BTL_SEtheory.
+
+
+Local Open Scope option_monad_scope.
+
+(** * Refinement of data-structures and of the specification of the simulation test *)
+
+Local Hint Resolve OK_PRE OK_SMEM OK_SREG: core.
+Local Hint Constructors si_ok: core.
+
+(* NB: refinement of (symbolic) internal state *)
+Record ristate :=
+ {
+ (** [ris_smem] represents the current smem symbolic evaluations.
+ (we also recover the history of smem in ris_smem) *)
+ ris_smem:> smem;
+ (** For the values in registers:
+ 1) we store a list of sval evaluations
+ 2) we encode the symbolic regset by a PTree + a boolean indicating the default sval
+ See [ris_sreg_get] below.
+ *)
+ ris_input_init: bool;
+ ok_rsval: list sval;
+ ris_sreg:> PTree.t sval
+ }.
+
+Definition ris_sreg_get (ris: ristate) r: sval :=
+ match PTree.get r ris with
+ | None => if ris_input_init ris then fSinput r else fSundef
+ | Some sv => sv
+ end.
+Coercion ris_sreg_get: ristate >-> Funclass.
+
+Definition ris_sreg_set (ris: ristate) (sreg: PTree.t sval): ristate :=
+ {| ris_smem := ris_smem ris;
+ ris_input_init := ris_input_init ris;
+ ok_rsval := ok_rsval ris;
+ ris_sreg := sreg |}.
+
+Lemma ris_sreg_set_access (ris: ristate) (sreg: PTree.t sval) r smem rsval:
+ {| ris_smem := smem;
+ ris_input_init := ris_input_init ris;
+ ok_rsval := rsval;
+ ris_sreg := sreg |} r =
+ ris_sreg_set ris sreg r.
+Proof.
+ unfold ris_sreg_set, ris_sreg_get; simpl.
+ reflexivity.
+Qed.
+
+Record ris_ok ctx (ris: ristate) : Prop := {
+ OK_RMEM: (eval_smem ctx (ris_smem ris)) <> None;
+ OK_RREG: forall sv, List.In sv (ok_rsval ris) -> eval_sval ctx sv <> None
+}.
+Local Hint Resolve OK_RMEM OK_RREG: core.
+Local Hint Constructors ris_ok: core.
+
+(**
+ NOTE that this refinement relation *cannot* be decomposed into a abstraction function of type
+ ristate -> sistate & an equivalence relation on istate.
+
+ Indeed, any [sis] satisfies [forall ctx r, si_ok ctx sis -> eval_sval ctx (sis r) <> None].
+ whereas this is generally not true for [ris] that [forall ctx r, ris_ok ctx ris -> eval_sval ctx (ris r) <> None],
+ except when, precisely, [ris_refines ris sis].
+
+ An alternative design enabling to define ris_refines as the composition of an equivalence on sistate
+ and a abstraction function would consist in constraining sistate with an additional [wf] field:
+
+ Record sistate :=
+ { si_pre: iblock_exec_context -> Prop;
+ si_sreg:> reg -> sval;
+ si_smem: smem;
+ si_wf: forall ctx, si_pre ctx -> si_smem <> None /\ forall r, si_sreg r <> None
+ }.
+
+ Such a constraint should also appear in [ristate]. This is not clear whether this alternative design
+ would be really simpler.
+
+*)
+Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := {
+ OK_EQUIV: si_ok ctx sis <-> ris_ok ctx ris;
+ MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem);
+ REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris r) = eval_sval ctx (sis r)
+}.
+Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ: core.
+Local Hint Constructors ris_refines: core.
+
+Record ris_simu ris1 ris2: Prop := {
+ SIMU_FAILS: forall sv, List.In sv ris2.(ok_rsval) -> List.In sv ris1.(ok_rsval);
+ SIMU_MEM: ris1.(ris_smem) = ris2.(ris_smem);
+ SIMU_REG: forall r, ris1 r = ris2 r
+}.
+Local Hint Resolve SIMU_FAILS SIMU_MEM SIMU_REG: core.
+Local Hint Constructors ris_simu: core.
+Local Hint Resolve sge_match: core.
+
+Lemma ris_simu_ok_preserv f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2):
+ ris_simu ris1 ris2 -> ris_ok (bctx1 ctx) ris1 -> ris_ok (bctx2 ctx) ris2.
+Proof.
+ intros SIMU OK; econstructor; eauto.
+ - erewrite <- SIMU_MEM; eauto.
+ erewrite <- smem_eval_preserved; eauto.
+ - intros; erewrite <- eval_sval_preserved; eauto.
+Qed.
+
+Lemma ris_simu_correct f1 f2 ris1 ris2 (ctx:simu_proof_context f1 f2) sis1 sis2:
+ ris_simu ris1 ris2 ->
+ ris_refines (bctx1 ctx) ris1 sis1 ->
+ ris_refines (bctx2 ctx) ris2 sis2 ->
+ sistate_simu ctx sis1 sis2.
+Proof.
+ intros RIS REF1 REF2 rs m SEM.
+ exploit sem_si_ok; eauto.
+ erewrite OK_EQUIV; eauto.
+ intros ROK1.
+ exploit ris_simu_ok_preserv; eauto.
+ intros ROK2. generalize ROK2; erewrite <- OK_EQUIV; eauto.
+ intros SOK2.
+ destruct SEM as (PRE & SMEM & SREG).
+ unfold sem_sistate; intuition eauto.
+ + erewrite <- MEM_EQ, <- SIMU_MEM; eauto.
+ erewrite <- smem_eval_preserved; eauto.
+ erewrite MEM_EQ; eauto.
+ + erewrite <- REG_EQ, <- SIMU_REG; eauto.
+ erewrite <- eval_sval_preserved; eauto.
+ erewrite REG_EQ; eauto.
+Qed.
+
+Inductive optrsv_refines ctx: (option sval) -> (option sval) -> Prop :=
+ | RefSome rsv sv
+ (REF:eval_sval ctx rsv = eval_sval ctx sv)
+ :optrsv_refines ctx (Some rsv) (Some sv)
+ | RefNone: optrsv_refines ctx None None
+ .
+
+Inductive rsvident_refines ctx: (sval + ident) -> (sval + ident) -> Prop :=
+ | RefLeft rsv sv
+ (REF:eval_sval ctx rsv = eval_sval ctx sv)
+ :rsvident_refines ctx (inl rsv) (inl sv)
+ | RefRight id1 id2
+ (IDSIMU: id1 = id2)
+ :rsvident_refines ctx (inr id1) (inr id2)
+ .
+
+Definition bargs_refines ctx (rargs: list (builtin_arg sval)) (args: list (builtin_arg sval)): Prop :=
+ eval_list_builtin_sval ctx rargs = eval_list_builtin_sval ctx args.
+
+Inductive rfv_refines ctx: sfval -> sfval -> Prop :=
+ | RefGoto pc: rfv_refines ctx (Sgoto pc) (Sgoto pc)
+ | RefCall sig rvos ros rargs args r pc
+ (SV:rsvident_refines ctx rvos ros)
+ (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args)
+ :rfv_refines ctx (Scall sig rvos rargs r pc) (Scall sig ros args r pc)
+ | RefTailcall sig rvos ros rargs args
+ (SV:rsvident_refines ctx rvos ros)
+ (LIST:eval_list_sval ctx rargs = eval_list_sval ctx args)
+ :rfv_refines ctx (Stailcall sig rvos rargs) (Stailcall sig ros args)
+ | RefBuiltin ef lbra lba br pc
+ (BARGS: bargs_refines ctx lbra lba)
+ :rfv_refines ctx (Sbuiltin ef lbra br pc) (Sbuiltin ef lba br pc)
+ | RefJumptable rsv sv lpc
+ (VAL: eval_sval ctx rsv = eval_sval ctx sv)
+ :rfv_refines ctx (Sjumptable rsv lpc) (Sjumptable sv lpc)
+ | RefReturn orsv osv
+ (OPT:optrsv_refines ctx orsv osv)
+ :rfv_refines ctx (Sreturn orsv) (Sreturn osv)
+.
+
+Definition rfv_simu (rfv1 rfv2: sfval): Prop := rfv1 = rfv2.
+
+Local Hint Resolve eval_sval_preserved list_sval_eval_preserved smem_eval_preserved eval_list_builtin_sval_preserved: core.
+
+Lemma rvf_simu_correct f1 f2 rfv1 rfv2 (ctx: simu_proof_context f1 f2) sfv1 sfv2:
+ rfv_simu rfv1 rfv2 ->
+ rfv_refines (bctx1 ctx) rfv1 sfv1 ->
+ rfv_refines (bctx2 ctx) rfv2 sfv2 ->
+ sfv_simu ctx sfv1 sfv2.
+Proof.
+ unfold rfv_simu; intros X REF1 REF2. subst.
+ unfold bctx2; destruct REF1; inv REF2; simpl; econstructor; eauto.
+ - (* call svid *)
+ inv SV; inv SV0; econstructor; eauto.
+ rewrite <- REF, <- REF0; eauto.
+ - (* call args *)
+ rewrite <- LIST, <- LIST0; eauto.
+ - (* taillcall svid *)
+ inv SV; inv SV0; econstructor; eauto.
+ rewrite <- REF, <- REF0; eauto.
+ - (* tailcall args *)
+ rewrite <- LIST, <- LIST0; eauto.
+ - (* builtin args *)
+ unfold bargs_refines, bargs_simu in *.
+ rewrite <- BARGS, <- BARGS0; eauto.
+ - rewrite <- VAL, <- VAL0; eauto.
+ - (* return *)
+ inv OPT; inv OPT0; econstructor; eauto.
+ rewrite <- REF, <- REF0; eauto.
+Qed.
+
+(* refinement of (symbolic) state *)
+Inductive rstate :=
+ | Rfinal (ris: ristate) (rfv: sfval)
+ | Rcond (cond: condition) (rargs: list_sval) (rifso rifnot: rstate)
+ | Rabort
+ .
+
+Record routcome := rout {
+ _ris: ristate;
+ _rfv: sfval;
+}.
+
+Fixpoint get_routcome ctx (rst:rstate): option routcome :=
+ match rst with
+ | Rfinal ris rfv => Some (rout ris rfv)
+ | Rcond cond args ifso ifnot =>
+ SOME b <- eval_scondition ctx cond args IN
+ get_routcome ctx (if b then ifso else ifnot)
+ | Rabort => None
+ end.
+
+Inductive rst_simu: rstate -> rstate -> Prop :=
+ | Rfinal_simu ris1 ris2 rfv1 rfv2
+ (RIS: ris_simu ris1 ris2)
+ (RFV: rfv_simu rfv1 rfv2)
+ : rst_simu (Rfinal ris1 rfv1) (Rfinal ris2 rfv2)
+ | Rcond_simu cond rargs rifso1 rifnot1 rifso2 rifnot2
+ (IFSO: rst_simu rifso1 rifso2)
+ (IFNOT: rst_simu rifnot1 rifnot2)
+ : rst_simu (Rcond cond rargs rifso1 rifnot1) (Rcond cond rargs rifso2 rifnot2)
+ | Rabort_simu: rst_simu Rabort Rabort
+(* TODO: extension à voir dans un second temps !
+ | Rcond_skip cond rargs rifso1 rifnot1 rst:
+ rst_simu rifso1 rst ->
+ rst_simu rifnot1 rst ->
+ rst_simu (Rcond cond rargs rifso1 rifnot1) rst
+*)
+ .
+
+Lemma rst_simu_lroutcome rst1 rst2:
+ rst_simu rst1 rst2 ->
+ forall f1 f2 (ctx: simu_proof_context f1 f2) ris1 rfv1,
+ get_routcome (bctx1 ctx) rst1 = Some (rout ris1 rfv1) ->
+ exists ris2 rfv2, get_routcome (bctx2 ctx) rst2 = Some (rout ris2 rfv2) /\ ris_simu ris1 ris2 /\ rfv_simu rfv1 rfv2.
+Proof.
+ induction 1; simpl; intros f1 f2 ctx lris1 lrfv1 ROUT; try_simplify_someHyps.
+ erewrite <- eval_scondition_preserved.
+ autodestruct.
+ destruct b; simpl; auto.
+Qed.
+
+Inductive rst_refines ctx: rstate -> sstate -> Prop :=
+ | Reffinal ris sis rfv sfv
+ (RIS: ris_refines ctx ris sis)
+ (RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv)
+ : rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv)
+ | Refcond rcond cond rargs args rifso rifnot ifso ifnot
+ (RCOND: eval_scondition ctx rcond rargs = eval_scondition ctx cond args)
+ (REFso: eval_scondition ctx rcond rargs = Some true -> rst_refines ctx rifso ifso)
+ (REFnot: eval_scondition ctx rcond rargs = Some false -> rst_refines ctx rifnot ifnot)
+ : rst_refines ctx (Rcond rcond rargs rifso rifnot) (Scond cond args ifso ifnot)
+ | Refabort
+ : rst_refines ctx Rabort Sabort
+ .
+
+Lemma rst_refines_outcome_up ctx rst st:
+ rst_refines ctx rst st ->
+ forall ris rfv,
+ get_routcome ctx rst = Some (rout ris rfv) ->
+ exists sis sfv, get_soutcome ctx st = Some (sout sis sfv) /\ ris_refines ctx ris sis /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv).
+Proof.
+ induction 1; simpl; intros lris lrfv ROUT; try_simplify_someHyps.
+ rewrite RCOND.
+ autodestruct.
+ destruct b; simpl; auto.
+Qed.
+
+Lemma rst_refines_outcome_okpreserv ctx rst st:
+ rst_refines ctx rst st ->
+ forall sis sfv,
+ get_soutcome ctx st = Some (sout sis sfv) ->
+ exists ris rfv, get_routcome ctx rst = Some (rout ris rfv) /\ ris_refines ctx ris sis /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv).
+Proof.
+ induction 1; simpl; intros lris lrfv ROUT; try_simplify_someHyps.
+ rewrite RCOND.
+ autodestruct.
+ destruct b; simpl; auto.
+Qed.
+
+Local Hint Resolve ris_simu_correct rvf_simu_correct: core.
+
+Lemma rst_simu_correct f1 f2 (ctx: simu_proof_context f1 f2) rst1 rst2 st1 st2
+ (SIMU: rst_simu rst1 rst2)
+ (REF1: forall sis sfv, get_soutcome (bctx1 ctx) st1 = Some (sout sis sfv) -> si_ok (bctx1 ctx) sis -> rst_refines (bctx1 ctx) rst1 st1)
+ (REF2: forall ris rfv, get_routcome (bctx2 ctx) rst2 = Some (rout ris rfv) -> ris_ok (bctx2 ctx) ris -> rst_refines (bctx2 ctx) rst2 st2)
+ :sstate_simu ctx st1 st2.
+Proof.
+ intros sis1 sfv1 SOUT OK1.
+ exploit REF1; eauto.
+ clear REF1; intros REF1.
+ exploit rst_refines_outcome_okpreserv; eauto. clear REF1 SOUT.
+ intros (ris1 & rfv1 & ROUT1 & REFI1 & REFF1).
+ rewrite OK_EQUIV in OK1; eauto.
+ exploit REFF1; eauto. clear REFF1; intros REFF1.
+ exploit rst_simu_lroutcome; eauto.
+ intros (ris2 & rfv2 & ROUT2 & SIMUI & SIMUF). clear ROUT1.
+ exploit ris_simu_ok_preserv; eauto.
+ clear OK1. intros OK2.
+ exploit REF2; eauto. clear REF2; intros REF2.
+ exploit rst_refines_outcome_up; eauto.
+ intros (sis2 & sfv2 & SOUT & REFI2 & REFF2).
+ do 2 eexists; split; eauto.
+Qed.
+
+
+(** * Properties of the (abstract) operators involved in symbolic execution *)
+
+Lemma ok_set_mem ctx sm sis:
+ si_ok ctx (set_smem sm sis)
+ <-> (si_ok ctx sis /\ eval_smem ctx sm <> None).
+Proof.
+ split; destruct 1; econstructor; simpl in *; eauto.
+ intuition eauto.
+Qed.
+
+Lemma ok_set_sreg ctx r sv sis:
+ si_ok ctx (set_sreg r sv sis)
+ <-> (si_ok ctx sis /\ eval_sval ctx sv <> None).
+Proof.
+ unfold set_sreg; split.
+ + intros [(SVAL & PRE) SMEM SREG]; simpl in *; split.
+ - econstructor; eauto.
+ intros r0; generalize (SREG r0); destruct (Pos.eq_dec r r0); try congruence.
+ - generalize (SREG r); destruct (Pos.eq_dec r r); try congruence.
+ + intros ([PRE SMEM SREG] & SVAL).
+ econstructor; simpl; eauto.
+ intros r0; destruct (Pos.eq_dec r r0); try congruence.
+Qed.
+
+Lemma si_ok_op_okpreserv ctx op args dest sis: si_ok ctx (sexec_op op args dest sis) -> si_ok ctx sis.
+Proof.
+ unfold sexec_op. rewrite ok_set_sreg. intuition.
+Qed.
+
+Lemma si_ok_load_okpreserv ctx chunk addr args dest sis trap: si_ok ctx (sexec_load trap chunk addr args dest sis) -> si_ok ctx sis.
+Proof.
+ unfold sexec_load. rewrite ok_set_sreg. intuition.
+Qed.
+
+Lemma si_ok_store_okpreserv ctx chunk addr args src sis: si_ok ctx (sexec_store chunk addr args src sis) -> si_ok ctx sis.
+Proof.
+ unfold sexec_store. rewrite ok_set_mem. intuition.
+Qed.
+
+Lemma si_ok_tr_okpreserv ctx fi sis: si_ok ctx (tr_sis (cf ctx) fi sis) -> si_ok ctx sis.
+Proof.
+ unfold tr_sis. intros OK; inv OK. simpl in *. intuition.
+Qed.
+
+(* These lemmas are very bad for eauto: we put them into a dedicated basis. *)
+Local Hint Resolve si_ok_tr_okpreserv si_ok_op_okpreserv si_ok_load_okpreserv si_ok_store_okpreserv: sis_ok.
+
+Lemma sexec_rec_okpreserv ctx ib:
+ forall k
+ (CONT: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis)
+ sis lsis sfv
+ (SOUT: get_soutcome ctx (sexec_rec (cf ctx) ib sis k) = Some (sout lsis sfv))
+ (OK: si_ok ctx lsis)
+ ,si_ok ctx sis.
+Proof.
+ induction ib; simpl; try (autodestruct; simpl).
+ 1-6: try_simplify_someHyps; eauto with sis_ok.
+ - intros. eapply IHib1. 2-3: eauto.
+ eapply IHib2; eauto.
+ - intros k CONT sis lsis sfv.
+ do 2 autodestruct; eauto.
+Qed.
+
+(* an alternative tr_sis *)
+
+Definition transfer_sis (inputs: list reg) (sis:sistate): sistate :=
+ {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None);
+ si_sreg := transfer_sreg inputs sis;
+ si_smem := sis.(si_smem) |}.
+
+Lemma ok_transfer_sis ctx inputs sis:
+ si_ok ctx (transfer_sis inputs sis)
+ <-> (si_ok ctx sis).
+Proof.
+ unfold transfer_sis. induction inputs as [|r l]; simpl.
+ + split; destruct 1; econstructor; simpl in *; intuition eauto. congruence.
+ + split.
+ * destruct 1; econstructor; simpl in *; intuition eauto.
+ * intros X; generalize X. rewrite <- IHl in X; clear IHl.
+ intros [PRE SMEM SREG].
+ econstructor; simpl; eauto.
+ intros r0; destruct (Pos.eq_dec r r0); try congruence.
+ intros H; eapply OK_SREG; eauto.
+Qed.
+
+Definition alt_tr_sis := poly_tr (fun f tbl or => transfer_sis (Regset.elements (pre_inputs f tbl or))).
+
+Lemma tr_sis_alt_def f fi sis:
+ alt_tr_sis f fi sis = tr_sis f fi sis.
+Proof.
+ unfold tr_sis, str_inputs. destruct fi; simpl; auto.
+Qed.
+
+
+(** * Refinement of the symbolic execution (e.g. refinement of [sexec] into [rexec]).
+
+TODO: Est-ce qu'on garde cette vision purement fonctionnelle de l'implementation ?
+=> ça pourrait être pratique quand on va compliquer encore le vérificateur !
+
+Du coup, on introduirait une version avec hash-consing qui serait en correspondance
+pour une relation d'equivalence sur les ristate ?
+
+Attention toutefois, il est possible que certaines parties de l'execution soient pénibles à formuler
+en programmation fonctionnelle pure (exemple: usage de l'égalité de pointeur comme test d'égalité rapide!)
+
+
+*)
+
+Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core.
+
+Lemma eval_list_sval_refpreserv ctx args ris sis:
+ ris_refines ctx ris sis ->
+ ris_ok ctx ris ->
+ eval_list_sval ctx (list_sval_inj (map ris args)) =
+ eval_list_sval ctx (list_sval_inj (map (si_sreg sis) args)).
+Proof.
+ intros REF OK.
+ induction args; simpl; eauto.
+ intros; erewrite REG_EQ, IHargs; eauto.
+Qed.
+
+Local Hint Resolve eval_list_sval_refpreserv: core.
+
+Lemma eval_builtin_sval_refpreserv ctx arg ris sis:
+ ris_refines ctx ris sis ->
+ ris_ok ctx ris ->
+ eval_builtin_sval ctx (builtin_arg_map ris arg) = eval_builtin_sval ctx (builtin_arg_map sis arg).
+Proof.
+ intros REF OK; induction arg; simpl; eauto.
+ + erewrite REG_EQ; eauto.
+ + erewrite IHarg1, IHarg2; eauto.
+ + erewrite IHarg1, IHarg2; eauto.
+Qed.
+
+Lemma bargs_refpreserv ctx args ris sis:
+ ris_refines ctx ris sis ->
+ ris_ok ctx ris ->
+ bargs_refines ctx (map (builtin_arg_map ris) args) (map (builtin_arg_map sis) args).
+Proof.
+ unfold bargs_refines. intros REF OK.
+ induction args; simpl; eauto.
+ erewrite eval_builtin_sval_refpreserv, IHargs; eauto.
+Qed.
+
+Local Hint Resolve bargs_refpreserv: core.
+
+Lemma exec_final_refpreserv ctx i ris sis:
+ ris_refines ctx ris sis ->
+ ris_ok ctx ris ->
+ rfv_refines ctx (sexec_final_sfv i ris) (sexec_final_sfv i sis).
+Proof.
+ destruct i; simpl; unfold sum_left_map; try autodestruct; eauto.
+Qed.
+
+Definition ris_init: ristate := {| ris_smem:= fSinit; ris_input_init:=true; ok_rsval := nil; ris_sreg := PTree.empty _ |}.
+
+Lemma ris_init_correct ctx:
+ ris_refines ctx ris_init sis_init.
+Proof.
+ unfold ris_init, sis_init; econstructor; simpl in *; eauto.
+ + split; destruct 1; econstructor; simpl in *; eauto.
+ congruence.
+ + destruct 1; simpl in *. unfold ris_sreg_get; simpl.
+ intros; rewrite PTree.gempty; eauto.
+Qed.
+
+Definition rset_smem rm (ris:ristate): ristate :=
+ {| ris_smem := rm;
+ ris_input_init := ris.(ris_input_init);
+ ok_rsval := ris.(ok_rsval);
+ ris_sreg:= ris.(ris_sreg)
+ |}.
+
+Lemma ok_rset_mem ctx rm (ris: ristate):
+ (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) ->
+ ris_ok ctx (rset_smem rm ris)
+ <-> (ris_ok ctx ris /\ eval_smem ctx rm <> None).
+Proof.
+ split; destruct 1; econstructor; simpl in *; eauto.
+Qed.
+
+Lemma rset_mem_correct ctx rm sm ris sis:
+ (eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) ->
+ ris_refines ctx ris sis ->
+ (ris_ok ctx ris -> eval_smem ctx rm = eval_smem ctx sm) ->
+ ris_refines ctx (rset_smem rm ris) (set_smem sm sis).
+Proof.
+ destruct 2; intros.
+ econstructor; eauto.
+ + rewrite ok_set_mem, ok_rset_mem; intuition congruence.
+ + rewrite ok_rset_mem; intuition eauto.
+ + rewrite ok_rset_mem; intuition eauto.
+Qed.
+
+Definition rexec_store chunk addr args src ris: ristate :=
+ let args := list_sval_inj (List.map (ris_sreg_get ris) args) in
+ let src := ris_sreg_get ris src in
+ let rm := fSstore ris.(ris_smem) chunk addr args src in
+ rset_smem rm ris.
+
+Lemma rexec_store_correct ctx chunk addr args src ris sis:
+ ris_refines ctx ris sis ->
+ ris_refines ctx (rexec_store chunk addr args src ris) (sexec_store chunk addr args src sis).
+Proof.
+ intros REF; eapply rset_mem_correct; simpl; eauto.
+ + intros X; rewrite X. repeat autodestruct; eauto.
+ + intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ, REG_EQ; eauto.
+Qed.
+
+(* TODO: reintroduire le "root_apply" ? *)
+
+Definition rset_sreg r rsv (ris:ristate): ristate :=
+ {| ris_smem := ris.(ris_smem);
+ ris_input_init := ris.(ris_input_init);
+ ok_rsval := rsv::ris.(ok_rsval); (* TODO: A CHANGER ? *)
+ ris_sreg:= PTree.set r rsv ris.(ris_sreg) (* TODO: A CHANGER *)
+ |}.
+
+Lemma ok_rset_sreg ctx r rsv ris:
+ ris_ok ctx (rset_sreg r rsv ris)
+ <-> (ris_ok ctx ris /\ eval_sval ctx rsv <> None).
+Proof.
+ split; destruct 1; econstructor; simpl in *; eauto.
+ intuition subst; eauto.
+ exploit OK_RREG; eauto.
+Qed.
+
+Lemma rset_reg_correct ctx r rsv sv ris sis:
+ ris_refines ctx ris sis ->
+ (ris_ok ctx ris -> eval_sval ctx rsv = eval_sval ctx sv) ->
+ ris_refines ctx (rset_sreg r rsv ris) (set_sreg r sv sis).
+Proof.
+ destruct 1; intros.
+ econstructor; eauto.
+ + rewrite ok_set_sreg, ok_rset_sreg; intuition congruence.
+ + rewrite ok_rset_sreg; intuition eauto.
+ + rewrite ok_rset_sreg. intros; unfold rset_sreg, set_sreg, ris_sreg_get; simpl. intuition eauto.
+ destruct (Pos.eq_dec _ _).
+ * subst; rewrite PTree.gss; eauto.
+ * rewrite PTree.gso; eauto.
+Qed.
+
+Definition rexec_op op args dst (ris:ristate): ristate :=
+ let args := list_sval_inj (List.map ris args) in
+ rset_sreg dst (fSop op args) ris.
+
+Lemma rexec_op_correct ctx op args dst ris sis:
+ ris_refines ctx ris sis ->
+ ris_refines ctx (rexec_op op args dst ris) (sexec_op op args dst sis).
+Proof.
+ intros REF; eapply rset_reg_correct; simpl; eauto.
+ intros OK; erewrite eval_list_sval_refpreserv; eauto.
+Qed.
+
+Definition rexec_load trap chunk addr args dst (ris:ristate): ristate :=
+ let args := list_sval_inj (List.map ris args) in
+ rset_sreg dst (fSload ris.(ris_smem) trap chunk addr args) ris.
+
+Lemma rexec_load_correct ctx trap chunk addr args dst ris sis:
+ ris_refines ctx ris sis ->
+ ris_refines ctx (rexec_load trap chunk addr args dst ris) (sexec_load trap chunk addr args dst sis).
+Proof.
+ intros REF; eapply rset_reg_correct; simpl; eauto.
+ intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ; eauto.
+Qed.
+
+Lemma eval_scondition_refpreserv ctx cond args ris sis:
+ ris_refines ctx ris sis ->
+ ris_ok ctx ris ->
+ eval_scondition ctx cond (list_sval_inj (map ris args)) = eval_scondition ctx cond (list_sval_inj (map sis args)).
+Proof.
+ intros; unfold eval_scondition.
+ erewrite eval_list_sval_refpreserv; eauto.
+Qed.
+
+
+(* transfer *)
+
+Definition rseto_sreg r rsv (ris:ristate): ristate :=
+ {| ris_smem := ris.(ris_smem);
+ ris_input_init := ris.(ris_input_init);
+ ok_rsval := ris.(ok_rsval);
+ ris_sreg:= PTree.set r rsv ris.(ris_sreg) (* TODO: A CHANGER *)
+ |}.
+
+Lemma ok_rseto_sreg ctx r rsv ris:
+ ris_ok ctx (rseto_sreg r rsv ris)
+ <-> (ris_ok ctx ris).
+Proof.
+ split; destruct 1; econstructor; simpl in *; eauto.
+Qed.
+
+Lemma rseto_reg_correct ctx r rsv sv ris sis:
+ ris_refines ctx ris sis ->
+ (ris_ok ctx ris -> eval_sval ctx rsv <> None) ->
+ (ris_ok ctx ris -> eval_sval ctx rsv = eval_sval ctx sv) ->
+ ris_refines ctx (rseto_sreg r rsv ris) (set_sreg r sv sis).
+Proof.
+ destruct 1; intros.
+ econstructor; eauto.
+ + rewrite ok_set_sreg, ok_rseto_sreg; intuition congruence.
+ + rewrite ok_rseto_sreg; intuition eauto.
+ + rewrite ok_rseto_sreg. intros; unfold rseto_sreg, set_sreg, ris_sreg_get; simpl. intuition eauto.
+ destruct (Pos.eq_dec _ _).
+ * subst; rewrite PTree.gss; eauto.
+ * rewrite PTree.gso; eauto.
+Qed.
+
+Fixpoint transfer_ris (inputs: list reg) (ris:ristate): ristate :=
+ match inputs with
+ | nil => {| ris_smem := ris.(ris_smem);
+ ris_input_init := false;
+ ok_rsval := ris.(ok_rsval);
+ ris_sreg:= PTree.empty _
+ |}
+ | r::l => rseto_sreg r (ris_sreg_get ris r) (transfer_ris l ris)
+ end.
+
+Lemma ok_transfer_ris ctx inputs ris:
+ ris_ok ctx (transfer_ris inputs ris)
+ <-> (ris_ok ctx ris).
+Proof.
+ induction inputs as [|r l]; simpl.
+ + split; destruct 1; econstructor; simpl in *; intuition eauto.
+ + rewrite ok_rseto_sreg. auto.
+Qed.
+
+Lemma transfer_ris_correct ctx inputs ris sis:
+ ris_refines ctx ris sis ->
+ ris_refines ctx (transfer_ris inputs ris) (transfer_sis inputs sis).
+Proof.
+ destruct 1; intros.
+ induction inputs as [|r l].
+ + econstructor; eauto.
+ * erewrite ok_transfer_sis, ok_transfer_ris; eauto.
+ * erewrite ok_transfer_ris; eauto.
+ * erewrite ok_transfer_ris; simpl; unfold ris_sreg_get; simpl; eauto.
+ intros; rewrite PTree.gempty. simpl; auto.
+ + econstructor; eauto.
+ * erewrite ok_transfer_sis, ok_transfer_ris; eauto.
+ * erewrite ok_transfer_ris; simpl.
+ intros; erewrite MEM_EQ. 2: eauto.
+ - unfold transfer_sis; simpl; eauto.
+ - rewrite ok_transfer_ris; simpl; eauto.
+ * erewrite ok_transfer_ris; simpl.
+ intros H r0.
+ erewrite REG_EQ. 2: eapply rseto_reg_correct; eauto.
+ - unfold set_sreg; simpl; auto.
+ destruct (Pos.eq_dec _ _); simpl; auto.
+ - intros. rewrite REG_EQ0; auto. apply OK_SREG; tauto.
+ - rewrite ok_rseto_sreg, ok_transfer_ris. auto.
+Qed.
+
+Definition tr_ris := poly_tr (fun f tbl or => transfer_ris (Regset.elements (pre_inputs f tbl or))).
+
+Local Hint Resolve transfer_ris_correct ok_transfer_ris: core.
+Local Opaque transfer_ris.
+
+Lemma ok_tr_ris ctx fi ris:
+ ris_ok ctx (tr_ris (cf ctx) fi ris)
+ <-> (ris_ok ctx ris).
+Proof.
+ destruct fi; simpl; eauto.
+Qed.
+
+Lemma ris_ok_tr_okpreserv ctx fi ris: ris_ok ctx (tr_ris (cf ctx) fi ris) -> ris_ok ctx ris.
+Proof.
+ rewrite ok_tr_ris; auto.
+Qed.
+
+Lemma tr_ris_correct ctx fi ris sis:
+ ris_refines ctx ris sis ->
+ ris_refines ctx (tr_ris (cf ctx) fi ris) (tr_sis (cf ctx) fi sis).
+Proof.
+ intros REF. rewrite <- tr_sis_alt_def.
+ destruct fi; simpl; eauto.
+Qed.
+
+(** RAFFINEMENT EXEC SYMBOLIQUE **)
+
+Fixpoint rexec_rec f ib ris (k: ristate -> rstate): rstate :=
+ match ib with
+ | BF fin _ => Rfinal (tr_ris f fin ris) (sexec_final_sfv fin ris)
+ (* basic instructions *)
+ | Bnop _ => k ris
+ | Bop op args res _ => k (rexec_op op args res ris)
+ | Bload trap chunk addr args dst _ => k (rexec_load trap chunk addr args dst ris)
+ | Bstore chunk addr args src _ => k (rexec_store chunk addr args src ris)
+ (* composed instructions *)
+ | Bseq ib1 ib2 =>
+ rexec_rec f ib1 ris (fun ris2 => rexec_rec f ib2 ris2 k)
+ | Bcond cond args ifso ifnot _ =>
+ let args := list_sval_inj (List.map ris args) in
+ let ifso := rexec_rec f ifso ris k in
+ let ifnot := rexec_rec f ifnot ris k in
+ Rcond cond args ifso ifnot
+ end
+ .
+
+Definition rexec f ib := rexec_rec f ib ris_init (fun _ => Rabort).
+
+
+Local Hint Resolve ris_init_correct exec_final_refpreserv tr_ris_correct ris_ok_tr_okpreserv
+ rexec_op_correct rexec_load_correct rexec_store_correct: core.
+
+Local Hint Constructors rst_refines: core.
+
+Lemma rexec_rec_correct1 ctx ib:
+ forall rk k
+ (CONTh: forall sis lsis sfv, get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> si_ok ctx sis)
+ (CONT: forall ris sis lsis sfv st, ris_refines ctx ris sis -> k sis = st -> get_soutcome ctx (k sis) = Some (sout lsis sfv) -> si_ok ctx lsis -> rst_refines ctx (rk ris) (k sis))
+ ris sis lsis sfv st
+ (REF: ris_refines ctx ris sis)
+ (EXEC: sexec_rec (cf ctx) ib sis k = st)
+ (SOUT: get_soutcome ctx st = Some (sout lsis sfv))
+ (OK: si_ok ctx lsis)
+ , rst_refines ctx (rexec_rec (cf ctx) ib ris rk) st.
+Proof.
+ induction ib; simpl; try (intros; subst; eauto; fail).
+ - (* seq *)
+ intros; subst.
+ eapply IHib1. 3-6: eauto.
+ + simpl. eapply sexec_rec_okpreserv; eauto.
+ + intros; subst. eapply IHib2; eauto.
+ - (* cond *)
+ intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst.
+ assert (rOK: ris_ok ctx ris). {
+ erewrite <- OK_EQUIV. 2: eauto.
+ eapply sexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto.
+ }
+ generalize OUT; clear OUT; simpl.
+ autodestruct.
+ intros COND; generalize COND.
+ erewrite <- eval_scondition_refpreserv; eauto.
+ econstructor; try_simplify_someHyps.
+Qed.
+
+Lemma rexec_correct1 ctx ib sis sfv:
+ get_soutcome ctx (sexec (cf ctx) ib) = Some (sout sis sfv) ->
+ (si_ok ctx sis) ->
+ rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib).
+Proof.
+ unfold sexec; intros; eapply rexec_rec_correct1; eauto; simpl; congruence.
+Qed.
+
+
+(** COPIER-COLLER ... Y a-t-il moyen de l'eviter ? **)
+
+Lemma ris_ok_op_okpreserv ctx op args dest ris: ris_ok ctx (rexec_op op args dest ris) -> ris_ok ctx ris.
+Proof.
+ unfold rexec_op. rewrite ok_rset_sreg. intuition.
+Qed.
+
+Lemma ris_ok_load_okpreserv ctx chunk addr args dest ris trap: ris_ok ctx (rexec_load trap chunk addr args dest ris) -> ris_ok ctx ris.
+Proof.
+ unfold rexec_load. rewrite ok_rset_sreg. intuition.
+Qed.
+
+Lemma ris_ok_store_okpreserv ctx chunk addr args src ris: ris_ok ctx (rexec_store chunk addr args src ris) -> ris_ok ctx ris.
+Proof.
+ unfold rexec_store. rewrite ok_rset_mem; simpl.
+ + intuition.
+ + intros X; rewrite X; simpl.
+ repeat autodestruct.
+Qed.
+
+(* These lemmas are very bad for eauto: we put them into a dedicated basis. *)
+Local Hint Resolve ris_ok_store_okpreserv ris_ok_op_okpreserv ris_ok_load_okpreserv: ris_ok.
+
+Lemma rexec_rec_okpreserv ctx ib:
+ forall k
+ (CONT: forall ris lris rfv, get_routcome ctx (k ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris)
+ ris lris rfv
+ (ROUT: get_routcome ctx (rexec_rec (cf ctx) ib ris k) = Some (rout lris rfv))
+ (OK: ris_ok ctx lris)
+ ,ris_ok ctx ris.
+Proof.
+ induction ib; simpl; try (autodestruct; simpl).
+ 1-6: try_simplify_someHyps; eauto with ris_ok.
+ - intros. eapply IHib1. 2-3: eauto.
+ eapply IHib2; eauto.
+ - intros k CONT sis lsis sfv.
+ do 2 autodestruct; eauto.
+Qed.
+
+Lemma rexec_rec_correct2 ctx ib:
+ forall rk k
+ (CONTh: forall ris lris rfv, get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> ris_ok ctx ris)
+ (CONT: forall ris sis lris rfv st, ris_refines ctx ris sis -> rk ris = st -> get_routcome ctx (rk ris) = Some (rout lris rfv) -> ris_ok ctx lris -> rst_refines ctx (rk ris) (k sis))
+ ris sis lris rfv st
+ (REF: ris_refines ctx ris sis)
+ (EXEC: rexec_rec (cf ctx) ib ris rk = st)
+ (SOUT: get_routcome ctx st = Some (rout lris rfv))
+ (OK: ris_ok ctx lris)
+ , rst_refines ctx st (sexec_rec (cf ctx) ib sis k).
+Proof.
+ induction ib; simpl; try (intros; subst; eauto; fail).
+ - (* seq *)
+ intros; subst.
+ eapply IHib1. 3-6: eauto.
+ + simpl. eapply rexec_rec_okpreserv; eauto.
+ + intros; subst. eapply IHib2; eauto.
+ - (* cond *)
+ intros rk k CONTh CONT ris sis lsis sfv st REF EXEC OUT OK. subst.
+ assert (OK0: ris_ok ctx ris). {
+ eapply rexec_rec_okpreserv with (ib:=(Bcond cond args ib1 ib2 iinfo)); simpl; eauto.
+ }
+ generalize OUT; clear OUT; simpl.
+ autodestruct.
+ intros COND; generalize COND.
+ erewrite eval_scondition_refpreserv; eauto.
+ econstructor; try_simplify_someHyps.
+Qed.
+
+Lemma rexec_correct2 ctx ib ris rfv:
+ get_routcome ctx (rexec (cf ctx) ib) = Some (rout ris rfv) ->
+ (ris_ok ctx ris) ->
+ rst_refines ctx (rexec (cf ctx) ib) (sexec (cf ctx) ib).
+Proof.
+ unfold rexec; intros; eapply rexec_rec_correct2; eauto; simpl; congruence.
+Qed.
+
+Theorem rexec_simu_correct f1 f2 ib1 ib2:
+ rst_simu (rexec f1 ib1) (rexec f2 ib2) ->
+ symbolic_simu f1 f2 ib1 ib2.
+Proof.
+ intros SIMU ctx.
+ eapply rst_simu_correct; eauto.
+ + intros; eapply rexec_correct1; eauto.
+ + intros; eapply rexec_correct2; eauto.
+Qed.
diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v
new file mode 100644
index 00000000..033ed843
--- /dev/null
+++ b/scheduling/BTL_SEtheory.v
@@ -0,0 +1,1336 @@
+(** A theory of symbolic simulation (i.e. simulation of symbolic executions) on BTL blocks.
+
+NB: an efficient implementation with hash-consing will be defined in another file (some day)
+
+The main theorem of this file is [symbolic_simu_correct] stating
+that the abstract definition of symbolic simulation of two BTL blocks
+implies the simulation for BTL.fsem block-steps.
+
+
+*)
+
+Require Import Coqlib Maps Floats.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import Op Registers.
+Require Import RTL BTL OptionMonad.
+Require Export Impure.ImpHCons.
+Import HConsing.
+
+(** * Syntax and semantics of symbolic values *)
+
+(** The semantics of symbolic execution is parametrized by the context of the execution of a block *)
+Record iblock_exec_context := Bctx {
+ cge: BTL.genv; (** usual environment for identifiers *)
+ cf: function; (** ambient function of the block *)
+ csp: val; (** stack pointer *)
+ crs0: regset; (** initial state of registers (at the block entry) *)
+ cm0: mem (** initial memory state *)
+}.
+
+
+(** symbolic value *)
+Inductive sval :=
+ | Sundef (hid: hashcode)
+ | Sinput (r: reg) (hid: hashcode)
+ | Sop (op:operation) (lsv: list_sval) (hid: hashcode)
+ | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (hid: hashcode)
+(** list of symbolic values *)
+with list_sval :=
+ | Snil (hid: hashcode)
+ | Scons (sv: sval) (lsv: list_sval) (hid: hashcode)
+(** symbolic memory *)
+with smem :=
+ | Sinit (hid: hashcode)
+ | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval) (hid: hashcode)
+.
+
+Scheme sval_mut := Induction for sval Sort Prop
+with list_sval_mut := Induction for list_sval Sort Prop
+with smem_mut := Induction for smem Sort Prop.
+
+(** "fake" smart-constructors using an [unknown_hid] instead of the one provided by hash-consing.
+These smart-constructors are those used in the abstract model of symbolic execution.
+They will also appear in the implementation of rewriting rules (in order to avoid hash-consing handling
+in proofs of rewriting rules).
+*)
+
+Definition fSundef := Sundef unknown_hid.
+Definition fSinput (r: reg) := Sinput r unknown_hid.
+Definition fSop (op:operation) (lsv: list_sval) := Sop op lsv unknown_hid.
+Definition fSload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval)
+ := Sload sm trap chunk addr lsv unknown_hid.
+
+Definition fSnil := Snil unknown_hid.
+Definition fScons (sv: sval) (lsv: list_sval) := Scons sv lsv unknown_hid.
+
+Definition fSinit := Sinit unknown_hid.
+Definition fSstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval)
+ := Sstore sm chunk addr lsv srce unknown_hid.
+
+Fixpoint list_sval_inj (l: list sval): list_sval :=
+ match l with
+ | nil => fSnil
+ | v::l => fScons v (list_sval_inj l)
+ end.
+
+Local Open Scope option_monad_scope.
+
+(** Semantics *)
+Fixpoint eval_sval ctx (sv: sval): option val :=
+ match sv with
+ | Sundef _ => Some Vundef
+ | Sinput r _ => Some ((crs0 ctx)#r)
+ | Sop op l _ =>
+ SOME args <- eval_list_sval ctx l IN
+ eval_operation (cge ctx) (csp ctx) op args (cm0 ctx)
+ | Sload sm trap chunk addr lsv _ =>
+ match trap with
+ | TRAP =>
+ SOME args <- eval_list_sval ctx lsv IN
+ SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN
+ SOME m <- eval_smem ctx sm IN
+ Mem.loadv chunk m a
+ | NOTRAP =>
+ SOME args <- eval_list_sval ctx lsv IN
+ match (eval_addressing (cge ctx) (csp ctx) addr args) with
+ | None => Some Vundef
+ | Some a =>
+ SOME m <- eval_smem ctx sm IN
+ match (Mem.loadv chunk m a) with
+ | None => Some Vundef
+ | Some val => Some val
+ end
+ end
+ end
+ end
+with eval_list_sval ctx (lsv: list_sval): option (list val) :=
+ match lsv with
+ | Snil _ => Some nil
+ | Scons sv lsv' _ =>
+ SOME v <- eval_sval ctx sv IN
+ SOME lv <- eval_list_sval ctx lsv' IN
+ Some (v::lv)
+ end
+with eval_smem ctx (sm: smem): option mem :=
+ match sm with
+ | Sinit _ => Some (cm0 ctx)
+ | Sstore sm chunk addr lsv srce _ =>
+ SOME args <- eval_list_sval ctx lsv IN
+ SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN
+ SOME m <- eval_smem ctx sm IN
+ SOME sv <- eval_sval ctx srce IN
+ Mem.storev chunk m a sv
+ end.
+
+(** The symbolic memory preserves predicate Mem.valid_pointer with respect to initial memory.
+ Hence, arithmetic operations and Boolean conditions do not depend on the current memory of the block
+ (their semantics only depends on the initial memory of the block).
+
+ The correctness of this idea is proved on lemmas [sexec_op_correct] and [eval_scondition_eq].
+*)
+Lemma valid_pointer_preserv ctx sm:
+ forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer (cm0 ctx) b ofs = Mem.valid_pointer m b ofs.
+Proof.
+ induction sm; simpl; intros; try_simplify_someHyps; auto.
+ repeat autodestruct; intros; erewrite IHsm by reflexivity.
+ eapply Mem.storev_preserv_valid; eauto.
+Qed.
+Local Hint Resolve valid_pointer_preserv: core.
+
+Lemma eval_list_sval_inj ctx l (sreg: reg -> sval) rs:
+ (forall r : reg, eval_sval ctx (sreg r) = Some (rs # r)) ->
+ eval_list_sval ctx (list_sval_inj (map sreg l)) = Some (rs ## l).
+Proof.
+ intros H; induction l as [|r l]; simpl; repeat autodestruct; auto.
+Qed.
+
+Definition eval_scondition ctx (cond: condition) (lsv: list_sval): option bool :=
+ SOME args <- eval_list_sval ctx lsv IN
+ eval_condition cond args (cm0 ctx).
+
+
+(** * Auxiliary definitions on Builtins *)
+(* TODO: clean this. Some generic stuffs could be put in [AST.v] *)
+
+Section EVAL_BUILTIN_SARG. (* adapted from Events.v *)
+
+Variable ctx: iblock_exec_context.
+Variable m: mem.
+
+Inductive eval_builtin_sarg: builtin_arg sval -> val -> Prop :=
+ | seval_BA: forall x v,
+ eval_sval ctx x = Some v ->
+ eval_builtin_sarg (BA x) v
+ | seval_BA_int: forall n,
+ eval_builtin_sarg (BA_int n) (Vint n)
+ | seval_BA_long: forall n,
+ eval_builtin_sarg (BA_long n) (Vlong n)
+ | seval_BA_float: forall n,
+ eval_builtin_sarg (BA_float n) (Vfloat n)
+ | seval_BA_single: forall n,
+ eval_builtin_sarg (BA_single n) (Vsingle n)
+ | seval_BA_loadstack: forall chunk ofs v,
+ Mem.loadv chunk m (Val.offset_ptr (csp ctx) ofs) = Some v ->
+ eval_builtin_sarg (BA_loadstack chunk ofs) v
+ | seval_BA_addrstack: forall ofs,
+ eval_builtin_sarg (BA_addrstack ofs) (Val.offset_ptr (csp ctx) ofs)
+ | seval_BA_loadglobal: forall chunk id ofs v,
+ Mem.loadv chunk m (Senv.symbol_address (cge ctx) id ofs) = Some v ->
+ eval_builtin_sarg (BA_loadglobal chunk id ofs) v
+ | seval_BA_addrglobal: forall id ofs,
+ eval_builtin_sarg (BA_addrglobal id ofs) (Senv.symbol_address (cge ctx) id ofs)
+ | seval_BA_splitlong: forall hi lo vhi vlo,
+ eval_builtin_sarg hi vhi -> eval_builtin_sarg lo vlo ->
+ eval_builtin_sarg (BA_splitlong hi lo) (Val.longofwords vhi vlo)
+ | seval_BA_addptr: forall a1 a2 v1 v2,
+ eval_builtin_sarg a1 v1 -> eval_builtin_sarg a2 v2 ->
+ eval_builtin_sarg (BA_addptr a1 a2)
+ (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2)
+.
+
+Definition eval_builtin_sargs (al: list (builtin_arg sval)) (vl: list val) : Prop :=
+ list_forall2 eval_builtin_sarg al vl.
+
+Lemma eval_builtin_sarg_determ:
+ forall a v, eval_builtin_sarg a v -> forall v', eval_builtin_sarg a v' -> v' = v.
+Proof.
+ induction 1; intros v' EV; inv EV; try congruence.
+ f_equal; eauto.
+ apply IHeval_builtin_sarg1 in H3. apply IHeval_builtin_sarg2 in H5. subst; auto.
+Qed.
+
+Lemma eval_builtin_sargs_determ:
+ forall al vl, eval_builtin_sargs al vl -> forall vl', eval_builtin_sargs al vl' -> vl' = vl.
+Proof.
+ induction 1; intros v' EV; inv EV; f_equal; eauto using eval_builtin_sarg_determ.
+Qed.
+
+End EVAL_BUILTIN_SARG.
+
+(* NB: generic function that could be put into [AST] file *)
+Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B :=
+ match arg with
+ | BA x => BA (f x)
+ | BA_int n => BA_int n
+ | BA_long n => BA_long n
+ | BA_float f => BA_float f
+ | BA_single s => BA_single s
+ | BA_loadstack chunk ptr => BA_loadstack chunk ptr
+ | BA_addrstack ptr => BA_addrstack ptr
+ | BA_loadglobal chunk id ptr => BA_loadglobal chunk id ptr
+ | BA_addrglobal id ptr => BA_addrglobal id ptr
+ | BA_splitlong ba1 ba2 => BA_splitlong (builtin_arg_map f ba1) (builtin_arg_map f ba2)
+ | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2)
+ end.
+
+Lemma eval_builtin_sarg_correct ctx rs m sreg: forall arg varg,
+ (forall r, eval_sval ctx (sreg r) = Some rs # r) ->
+ eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg ->
+ eval_builtin_sarg ctx m (builtin_arg_map sreg arg) varg.
+Proof.
+ induction arg.
+ all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence).
+ - intros varg SEVAL BARG. inv BARG. simpl. constructor.
+ eapply IHarg1; eauto. eapply IHarg2; eauto.
+ - intros varg SEVAL BARG. inv BARG. simpl. constructor.
+ eapply IHarg1; eauto. eapply IHarg2; eauto.
+Qed.
+
+Lemma eval_builtin_sargs_correct ctx rs m sreg args vargs:
+ (forall r, eval_sval ctx (sreg r) = Some rs # r) ->
+ eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs ->
+ eval_builtin_sargs ctx m (map (builtin_arg_map sreg) args) vargs.
+Proof.
+ induction 2.
+ - constructor.
+ - simpl. constructor; [| assumption].
+ eapply eval_builtin_sarg_correct; eauto.
+Qed.
+
+Lemma eval_builtin_sarg_exact ctx rs m sreg: forall arg varg,
+ (forall r, eval_sval ctx (sreg r) = Some rs # r) ->
+ eval_builtin_sarg ctx m (builtin_arg_map sreg arg) varg ->
+ eval_builtin_arg (cge ctx) (fun r => rs # r) (csp ctx) m arg varg.
+Proof.
+ induction arg.
+ 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 eval_builtin_sargs_exact ctx rs m sreg: forall args vargs,
+ (forall r, eval_sval ctx (sreg r) = Some rs # r) ->
+ eval_builtin_sargs ctx m (map (builtin_arg_map sreg) args) vargs ->
+ eval_builtin_args (cge ctx) (fun r => rs # r) (csp ctx) m args vargs.
+Proof.
+ induction args.
+ - simpl. intros. inv H0. constructor.
+ - intros vargs SEVAL BARG. simpl in BARG. inv BARG.
+ constructor; [| eapply IHargs; eauto].
+ eapply eval_builtin_sarg_exact; eauto.
+Qed.
+
+Fixpoint eval_builtin_sval ctx bsv :=
+ match bsv with
+ | BA sv => SOME v <- eval_sval ctx sv IN Some (BA v)
+ | BA_splitlong sv1 sv2 =>
+ SOME v1 <- eval_builtin_sval ctx sv1 IN
+ SOME v2 <- eval_builtin_sval ctx sv2 IN
+ Some (BA_splitlong v1 v2)
+ | BA_addptr sv1 sv2 =>
+ SOME v1 <- eval_builtin_sval ctx sv1 IN
+ SOME v2 <- eval_builtin_sval ctx sv2 IN
+ Some (BA_addptr v1 v2)
+ | BA_int i => Some (BA_int i)
+ | BA_long l => Some (BA_long l)
+ | BA_float f => Some (BA_float f)
+ | BA_single s => Some (BA_single s)
+ | BA_loadstack chk ptr => Some (BA_loadstack chk ptr)
+ | BA_addrstack ptr => Some (BA_addrstack ptr)
+ | BA_loadglobal chk id ptr => Some (BA_loadglobal chk id ptr)
+ | BA_addrglobal id ptr => Some (BA_addrglobal id ptr)
+ end.
+
+Fixpoint eval_list_builtin_sval ctx lbsv :=
+ match lbsv with
+ | nil => Some nil
+ | bsv::lbsv => SOME v <- eval_builtin_sval ctx bsv IN
+ SOME lv <- eval_list_builtin_sval ctx lbsv IN
+ Some (v::lv)
+ end.
+
+Lemma eval_list_builtin_sval_nil ctx lbs2:
+ eval_list_builtin_sval ctx lbs2 = Some nil ->
+ lbs2 = nil.
+Proof.
+ destruct lbs2; simpl; repeat autodestruct; congruence.
+Qed.
+
+Lemma eval_builtin_sval_arg ctx bs:
+ forall ba m v,
+ eval_builtin_sval ctx bs = Some ba ->
+ eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v ->
+ eval_builtin_sarg ctx m bs v.
+Proof.
+ induction bs; simpl;
+ try (intros ba m v H; inversion H; subst; clear H;
+ intros H; inversion H; subst;
+ econstructor; auto; fail).
+ - intros ba m v; destruct (eval_sval _ _) eqn: SV;
+ intros H; inversion H; subst; clear H.
+ intros H; inversion H; subst.
+ econstructor; auto.
+ - intros ba m v.
+ destruct (eval_builtin_sval _ bs1) eqn: SV1; try congruence.
+ destruct (eval_builtin_sval _ bs2) eqn: SV2; try congruence.
+ intros H; inversion H; subst; clear H.
+ intros H; inversion H; subst.
+ econstructor; eauto.
+ - intros ba m v.
+ destruct (eval_builtin_sval _ bs1) eqn: SV1; try congruence.
+ destruct (eval_builtin_sval _ bs2) eqn: SV2; try congruence.
+ intros H; inversion H; subst; clear H.
+ intros H; inversion H; subst.
+ econstructor; eauto.
+Qed.
+
+Lemma eval_builtin_sarg_sval ctx m v: forall bs,
+ eval_builtin_sarg ctx m bs v ->
+ exists ba,
+ eval_builtin_sval ctx bs = Some ba
+ /\ eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m ba v.
+Proof.
+ induction 1.
+ all: try (eexists; constructor; [simpl; reflexivity | constructor]).
+ 2-3: try assumption.
+ - eexists. constructor.
+ + simpl. rewrite H. reflexivity.
+ + constructor.
+ - destruct IHeval_builtin_sarg1 as (ba1 & A1 & B1).
+ destruct IHeval_builtin_sarg2 as (ba2 & A2 & B2).
+ eexists. constructor.
+ + simpl. rewrite A1. rewrite A2. reflexivity.
+ + constructor; assumption.
+ - destruct IHeval_builtin_sarg1 as (ba1 & A1 & B1).
+ destruct IHeval_builtin_sarg2 as (ba2 & A2 & B2).
+ eexists. constructor.
+ + simpl. rewrite A1. rewrite A2. reflexivity.
+ + constructor; assumption.
+Qed.
+
+Lemma eval_builtin_sval_args ctx lbs:
+ forall lba m v,
+ eval_list_builtin_sval ctx lbs = Some lba ->
+ list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba v ->
+ eval_builtin_sargs ctx m lbs v.
+Proof.
+ unfold eval_builtin_sargs; induction lbs; simpl; intros lba m v.
+ - intros H; inversion H; subst; clear H.
+ intros H; inversion H. econstructor.
+ - destruct (eval_builtin_sval _ _) eqn:SV; try congruence.
+ destruct (eval_list_builtin_sval _ _) eqn: SVL; try congruence.
+ intros H; inversion H; subst; clear H.
+ intros H; inversion H; subst; clear H.
+ econstructor; eauto.
+ eapply eval_builtin_sval_arg; eauto.
+Qed.
+
+Lemma eval_builtin_sargs_sval ctx m lv: forall lbs,
+ eval_builtin_sargs ctx m lbs lv ->
+ exists lba,
+ eval_list_builtin_sval ctx lbs = Some lba
+ /\ list_forall2 (eval_builtin_arg (cge ctx) (fun id => id) (csp ctx) m) lba lv.
+Proof.
+ induction 1.
+ - eexists. constructor.
+ + simpl. reflexivity.
+ + constructor.
+ - destruct IHlist_forall2 as (lba & A & B).
+ apply eval_builtin_sarg_sval in H. destruct H as (ba & A' & B').
+ eexists. constructor.
+ + simpl. rewrite A'. rewrite A. reflexivity.
+ + constructor; assumption.
+Qed.
+
+Lemma eval_builtin_sval_correct ctx m: forall bs1 v bs2,
+ eval_builtin_sarg ctx m bs1 v ->
+ (eval_builtin_sval ctx bs1) = (eval_builtin_sval ctx bs2) ->
+ eval_builtin_sarg ctx m bs2 v.
+Proof.
+ intros. exploit eval_builtin_sarg_sval; eauto.
+ intros (ba & X1 & X2).
+ eapply eval_builtin_sval_arg; eauto.
+ congruence.
+Qed.
+
+Lemma eval_list_builtin_sval_correct ctx m vargs: forall lbs1,
+ eval_builtin_sargs ctx m lbs1 vargs ->
+ forall lbs2, (eval_list_builtin_sval ctx lbs1) = (eval_list_builtin_sval ctx lbs2) ->
+ eval_builtin_sargs ctx m lbs2 vargs.
+Proof.
+ intros. exploit eval_builtin_sargs_sval; eauto.
+ intros (ba & X1 & X2).
+ eapply eval_builtin_sval_args; eauto.
+ congruence.
+Qed.
+
+(** * Symbolic (final) value of a block *)
+
+(** TODO: faut-il hash-conser les valeurs symboliques finales. Pas très utile si pas de join interne.
+Mais peut être utile dans le cas contraire. *)
+
+Inductive sfval :=
+ | Sgoto (pc: exit)
+ | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:exit)
+ | Stailcall: signature -> sval + ident -> list_sval -> sfval
+ | Sbuiltin (ef:external_function) (sargs: list (builtin_arg sval)) (res: builtin_res reg) (pc:exit)
+ | Sjumptable (sv: sval) (tbl: list exit)
+ | Sreturn: option sval -> sfval
+.
+
+Definition sfind_function ctx (svos : sval + ident): option fundef :=
+ match svos with
+ | inl sv => SOME v <- eval_sval ctx sv IN Genv.find_funct (cge ctx) v
+ | inr symb => SOME b <- Genv.find_symbol (cge ctx) symb IN Genv.find_funct_ptr (cge ctx) b
+ end
+.
+
+Import ListNotations.
+Local Open Scope list_scope.
+
+Inductive sem_sfval ctx stk: sfval -> regset -> mem -> trace -> state -> Prop :=
+ | exec_Sgoto pc rs m:
+ sem_sfval ctx stk (Sgoto pc) rs m E0 (State stk (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] None rs) m)
+ | exec_Sreturn pstk osv rs m m' v:
+ (csp ctx) = (Vptr pstk Ptrofs.zero) ->
+ Mem.free m pstk 0 (cf ctx).(fn_stacksize) = Some m' ->
+ match osv with Some sv => eval_sval ctx sv | None => Some Vundef end = Some v ->
+ sem_sfval ctx stk (Sreturn osv) rs m
+ E0 (Returnstate stk v m')
+ | exec_Scall rs m sig svos lsv args res pc fd:
+ sfind_function ctx svos = Some fd ->
+ funsig fd = sig ->
+ eval_list_sval ctx lsv = Some args ->
+ sem_sfval ctx stk (Scall sig svos lsv res pc) rs m
+ E0 (Callstate (Stackframe res (cf ctx) (csp ctx) pc (tr_inputs ctx.(cf) [pc] (Some res) rs)::stk) fd args m)
+ | exec_Stailcall pstk rs m sig svos args fd m' lsv:
+ sfind_function ctx svos = Some fd ->
+ funsig fd = sig ->
+ (csp ctx) = Vptr pstk Ptrofs.zero ->
+ Mem.free m pstk 0 (cf ctx).(fn_stacksize) = Some m' ->
+ eval_list_sval ctx lsv = Some args ->
+ sem_sfval ctx stk (Stailcall sig svos lsv) rs m
+ E0 (Callstate stk fd args m')
+ | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs:
+ eval_builtin_sargs ctx m sargs vargs ->
+ external_call ef (cge ctx) vargs m t vres m' ->
+ sem_sfval ctx stk (Sbuiltin ef sargs res pc) rs m
+ t (State stk (cf ctx) (csp ctx) pc (regmap_setres res vres (tr_inputs (cf ctx) [pc] (reg_builtin_res res) rs)) m')
+ | exec_Sjumptable sv tbl pc' n rs m:
+ eval_sval ctx sv = Some (Vint n) ->
+ list_nth_z tbl (Int.unsigned n) = Some pc' ->
+ sem_sfval ctx stk (Sjumptable sv tbl) rs m
+ E0 (State stk (cf ctx) (csp ctx) pc' (tr_inputs ctx.(cf) tbl None rs) m)
+.
+
+(* Syntax and Semantics of symbolic internal states *)
+(* [si_pre] is a precondition on initial context *)
+Record sistate := { si_pre: iblock_exec_context -> Prop; si_sreg:> reg -> sval; si_smem: smem }.
+
+(* Predicate on which (rs, m) is a possible final state after evaluating [st] on ((crs0 ctx), (cm0 ctx)) *)
+Definition sem_sistate ctx (sis: sistate) (rs: regset) (m: mem): Prop :=
+ sis.(si_pre) ctx
+ /\ eval_smem ctx sis.(si_smem) = Some m
+ /\ forall (r:reg), eval_sval ctx (sis.(si_sreg) r) = Some (rs#r).
+
+(** * Symbolic execution of final step *)
+Definition sexec_final_sfv (i: final) (sreg: reg -> sval): sfval :=
+ match i with
+ | Bgoto pc => Sgoto pc
+ | Bcall sig ros args res pc =>
+ let svos := sum_left_map sreg ros in
+ let sargs := list_sval_inj (List.map sreg args) in
+ Scall sig svos sargs res pc
+ | Btailcall sig ros args =>
+ let svos := sum_left_map sreg ros in
+ let sargs := list_sval_inj (List.map sreg args) in
+ Stailcall sig svos sargs
+ | Bbuiltin ef args res pc =>
+ let sargs := List.map (builtin_arg_map sreg) args in
+ Sbuiltin ef sargs res pc
+ | Breturn or =>
+ let sor := SOME r <- or IN Some (sreg r) in
+ Sreturn sor
+ | Bjumptable reg tbl =>
+ let sv := sreg reg in
+ Sjumptable sv tbl
+ end.
+
+Local Hint Constructors sem_sfval: core.
+
+Lemma sexec_final_sfv_correct ctx stk i sis t rs m s:
+ sem_sistate ctx sis rs m ->
+ final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s ->
+ sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s.
+Proof.
+ intros (PRE&MEM&REG).
+ destruct 1; subst; try_simplify_someHyps; simpl; intros; try autodestruct; eauto.
+ + (* Bcall *) intros; eapply exec_Scall; auto.
+ - destruct ros; simpl in * |- *; auto.
+ rewrite REG; auto.
+ - erewrite eval_list_sval_inj; simpl; auto.
+ + (* Btailcall *) intros. eapply exec_Stailcall; eauto.
+ - destruct ros; simpl in * |- *; eauto.
+ rewrite REG; eauto.
+ - erewrite eval_list_sval_inj; simpl; auto.
+ + (* Bbuiltin *) intros. eapply exec_Sbuiltin; eauto.
+ eapply eval_builtin_sargs_correct; eauto.
+ + (* Bjumptable *) intros. eapply exec_Sjumptable; eauto. congruence.
+Qed.
+
+Local Hint Constructors final_step: core.
+Local Hint Resolve eval_builtin_sargs_exact: core.
+
+Lemma sexec_final_sfv_exact ctx stk i sis t rs m s:
+ sem_sistate ctx sis rs m ->
+ sem_sfval ctx stk (sexec_final_sfv i sis) rs m t s
+ -> final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs m i t s.
+Proof.
+ intros (PRE&MEM&REG).
+ destruct i; simpl; intros LAST; inv LAST; eauto.
+ + (* Breturn *)
+ enough (v=regmap_optget res Vundef rs) as ->; eauto.
+ destruct res; simpl in *; congruence.
+ + (* Bcall *)
+ erewrite eval_list_sval_inj in *; try_simplify_someHyps.
+ intros; eapply exec_Bcall; eauto.
+ destruct fn; simpl in * |- *; auto.
+ rewrite REG in * |- ; auto.
+ + (* Btailcall *)
+ erewrite eval_list_sval_inj in *; try_simplify_someHyps.
+ intros; eapply exec_Btailcall; eauto.
+ destruct fn; simpl in * |- *; auto.
+ rewrite REG in * |- ; auto.
+ + (* Bjumptable *)
+ eapply exec_Bjumptable; eauto.
+ congruence.
+Qed.
+
+(** * symbolic execution of basic instructions *)
+
+Definition sis_init : sistate := {| si_pre:= fun _ => True; si_sreg:= fun r => fSinput r; si_smem:= fSinit |}.
+
+Lemma sis_init_correct ctx:
+ sem_sistate ctx sis_init (crs0 ctx) (cm0 ctx).
+Proof.
+ unfold sis_init, sem_sistate; simpl; intuition eauto.
+Qed.
+
+Definition set_sreg (r:reg) (sv:sval) (sis:sistate): sistate :=
+ {| si_pre:=(fun ctx => eval_sval ctx (sis.(si_sreg) r) <> None /\ (sis.(si_pre) ctx));
+ si_sreg:=fun y => if Pos.eq_dec r y then sv else sis.(si_sreg) y;
+ si_smem:= sis.(si_smem)|}.
+
+Lemma set_sreg_correct ctx dst sv sis (rs rs': regset) m:
+ sem_sistate ctx sis rs m ->
+ (eval_sval ctx sv = Some rs' # dst) ->
+ (forall r, r <> dst -> rs'#r = rs#r) ->
+ sem_sistate ctx (set_sreg dst sv sis) rs' m.
+Proof.
+ intros (PRE&MEM&REG) NEW OLD.
+ unfold sem_sistate; simpl.
+ intuition.
+ - rewrite REG in *; congruence.
+ - destruct (Pos.eq_dec dst r); simpl; subst; eauto.
+ rewrite REG in *. rewrite OLD; eauto.
+Qed.
+
+Definition set_smem (sm:smem) (sis:sistate): sistate :=
+ {| si_pre:=(fun ctx => eval_smem ctx sis.(si_smem) <> None /\ (sis.(si_pre) ctx));
+ si_sreg:= sis.(si_sreg);
+ si_smem:= sm |}.
+
+Lemma set_smem_correct ctx sm sis rs m m':
+ sem_sistate ctx sis rs m ->
+ eval_smem ctx sm = Some m' ->
+ sem_sistate ctx (set_smem sm sis) rs m'.
+Proof.
+ intros (PRE&MEM&REG) NEW.
+ unfold sem_sistate; simpl.
+ intuition.
+ rewrite MEM in *; congruence.
+Qed.
+
+Definition sexec_op op args dst sis: sistate :=
+ let args := list_sval_inj (List.map sis.(si_sreg) args) in
+ set_sreg dst (fSop op args) sis.
+
+Lemma sexec_op_correct ctx op args dst sis rs m v
+ (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = Some v)
+ (SIS: sem_sistate ctx sis rs m)
+ :(sem_sistate ctx (sexec_op op args dst sis) (rs#dst <- v) m).
+Proof.
+ eapply set_sreg_correct; eauto.
+ - simpl. destruct SIS as (PRE&MEM&REG).
+ rewrite Regmap.gss; simpl; auto.
+ erewrite eval_list_sval_inj; simpl; auto.
+ try_simplify_someHyps.
+ intros; erewrite op_valid_pointer_eq; eauto.
+ - intros; rewrite Regmap.gso; auto.
+Qed.
+
+Definition sexec_load trap chunk addr args dst sis: sistate :=
+ let args := list_sval_inj (List.map sis.(si_sreg) args) in
+ set_sreg dst (fSload sis.(si_smem) trap chunk addr args) sis.
+
+Lemma sexec_load_correct ctx chunk addr args dst sis rs m v trap
+ (HLOAD: has_loaded (cge ctx) (csp ctx) rs m chunk addr args v trap)
+ (SIS: sem_sistate ctx sis rs m)
+ :(sem_sistate ctx (sexec_load trap chunk addr args dst sis) (rs#dst <- v) m).
+Proof.
+ inv HLOAD; eapply set_sreg_correct; eauto.
+ 2,4: intros; rewrite Regmap.gso; auto.
+ - simpl. destruct SIS as (PRE&MEM&REG).
+ destruct trap; rewrite Regmap.gss; simpl; auto;
+ erewrite eval_list_sval_inj; simpl; auto;
+ try_simplify_someHyps.
+ intros. rewrite LOAD; auto.
+ - simpl. destruct SIS as (PRE&MEM&REG).
+ rewrite Regmap.gss; simpl; auto.
+ erewrite eval_list_sval_inj; simpl; auto.
+ autodestruct; rewrite MEM, LOAD; auto.
+Qed.
+
+Definition sexec_store chunk addr args src sis: sistate :=
+ let args := list_sval_inj (List.map sis.(si_sreg) args) in
+ let src := sis.(si_sreg) src in
+ let sm := fSstore sis.(si_smem) chunk addr args src in
+ set_smem sm sis.
+
+Lemma sexec_store_correct ctx chunk addr args src sis rs m m' a
+ (EVAL: eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a)
+ (STORE: Mem.storev chunk m a (rs # src) = Some m')
+ (SIS: sem_sistate ctx sis rs m)
+ :(sem_sistate ctx (sexec_store chunk addr args src sis) rs m').
+Proof.
+ eapply set_smem_correct; eauto.
+ simpl. destruct SIS as (PRE&MEM&REG).
+ erewrite eval_list_sval_inj; simpl; auto.
+ try_simplify_someHyps.
+ rewrite REG; auto.
+Qed.
+
+Lemma eval_scondition_eq ctx cond args sis rs m
+ (SIS : sem_sistate ctx sis rs m)
+ :eval_scondition ctx cond (list_sval_inj (map (si_sreg sis) args)) = eval_condition cond rs ## args m.
+Proof.
+ destruct SIS as (PRE&MEM&REG); unfold eval_scondition; simpl.
+ erewrite eval_list_sval_inj; simpl; auto.
+ eapply cond_valid_pointer_eq; eauto.
+Qed.
+
+(** * Compute sistate associated to final values *)
+Fixpoint transfer_sreg (inputs: list reg) (sreg: reg -> sval): reg -> sval :=
+ match inputs with
+ | nil => fun r => fSundef
+ | r1::l => fun r => if Pos.eq_dec r1 r then sreg r1 else transfer_sreg l sreg r
+ end.
+
+Definition str_inputs (f:function) (tbl: list exit) (or:option reg) := transfer_sreg (Regset.elements (pre_inputs f tbl or)).
+
+Lemma str_inputs_correct ctx sis rs tbl or r:
+ (forall r : reg, eval_sval ctx (si_sreg sis r) = Some rs # r) ->
+ eval_sval ctx (str_inputs (cf ctx) tbl or (si_sreg sis) r) = Some (tr_inputs (cf ctx) tbl or rs) # r.
+Proof.
+ intros H.
+ unfold str_inputs, tr_inputs, transfer_regs.
+ induction (Regset.elements _) as [|x l]; simpl.
+ + rewrite Regmap.gi; auto.
+ + autodestruct; intros; subst.
+ * rewrite Regmap.gss; auto.
+ * rewrite Regmap.gso; eauto.
+Qed.
+
+Local Hint Resolve str_inputs_correct: core.
+
+Definition tr_sis f (fi: final) (sis: sistate) :=
+ {| si_pre := fun ctx => (sis.(si_pre) ctx /\ forall r, eval_sval ctx (sis.(si_sreg) r) <> None);
+ si_sreg := poly_tr str_inputs f fi sis.(si_sreg);
+ si_smem := sis.(si_smem) |}.
+
+Lemma tr_sis_regs_correct_aux ctx fin sis rs m:
+ sem_sistate ctx sis rs m ->
+ (forall r, eval_sval ctx (tr_sis (cf ctx) fin sis r) = Some (tr_regs (cf ctx) fin rs) # r).
+Proof.
+ Local Opaque str_inputs.
+ simpl. destruct 1 as (_ & _ & REG).
+ destruct fin; simpl; eauto.
+Qed.
+
+Lemma tr_sis_regs_correct ctx fin sis rs m:
+ sem_sistate ctx sis rs m ->
+ sem_sistate ctx (tr_sis (cf ctx) fin sis) (tr_regs (cf ctx) fin rs) m.
+Proof.
+ intros H.
+ generalize (tr_sis_regs_correct_aux _ fin _ _ _ H).
+ destruct H as (PRE & MEM & REG).
+ econstructor; simpl; intuition eauto || congruence.
+Qed.
+
+Definition poly_str {A} (tr: function -> list exit -> option reg -> A) f (sfv: sfval): A :=
+ match sfv with
+ | Sgoto pc => tr f [pc] None
+ | Scall _ _ _ res pc => tr f [pc] (Some res)
+ | Stailcall _ _ args => tr f [] None
+ | Sbuiltin _ _ res pc => tr f [pc] (reg_builtin_res res)
+ | Sreturn _ => tr f [] None
+ | Sjumptable _ tbl => tr f tbl None
+ end.
+
+Definition str_regs: function -> sfval -> regset -> regset :=
+ poly_str tr_inputs.
+
+Lemma str_tr_regs_equiv f fin sis:
+ str_regs f (sexec_final_sfv fin sis) = tr_regs f fin.
+Proof.
+ destruct fin; simpl; auto.
+Qed.
+
+(** * symbolic execution of blocks *)
+
+(* symbolic state *)
+Inductive sstate :=
+ | Sfinal (sis: sistate) (sfv: sfval)
+ | Scond (cond: condition) (args: list_sval) (ifso ifnot: sstate)
+ | Sabort
+ .
+
+(* outcome of a symbolic execution path *)
+Record soutcome := sout {
+ _sis: sistate;
+ _sfv: sfval;
+}.
+
+Fixpoint get_soutcome ctx (st:sstate): option soutcome :=
+ match st with
+ | Sfinal sis sfv => Some (sout sis sfv)
+ | Scond cond args ifso ifnot =>
+ SOME b <- eval_scondition ctx cond args IN
+ get_soutcome ctx (if b then ifso else ifnot)
+ | Sabort => None
+ end.
+
+(* transition (t,s) produced by a sstate in initial context ctx *)
+Inductive sem_sstate ctx stk t s: sstate -> Prop :=
+ | sem_Sfinal sis sfv rs m
+ (SIS: sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m)
+ (SFV: sem_sfval ctx stk sfv rs m t s)
+ : sem_sstate ctx stk t s (Sfinal sis sfv)
+ | sem_Scond b cond args ifso ifnot
+ (SEVAL: eval_scondition ctx cond args = Some b)
+ (SELECT: sem_sstate ctx stk t s (if b then ifso else ifnot))
+ : sem_sstate ctx stk t s (Scond cond args ifso ifnot)
+ (* NB: Sabort: fails to produce a transition *)
+ .
+
+Lemma sem_sstate_run ctx stk st t s:
+ sem_sstate ctx stk t s st ->
+ exists sis sfv rs m,
+ get_soutcome ctx st = Some (sout sis sfv)
+ /\ sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m
+ /\ sem_sfval ctx stk sfv rs m t s
+ .
+Proof.
+ induction 1; simpl; try_simplify_someHyps; do 4 eexists; intuition eauto.
+Qed.
+
+Local Hint Resolve sem_Sfinal: core.
+
+Lemma run_sem_sstate ctx st sis sfv:
+ get_soutcome ctx st = Some (sout sis sfv) ->
+ forall rs m stk s t,
+ sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m ->
+ sem_sfval ctx stk sfv rs m t s ->
+ sem_sstate ctx stk t s st
+ .
+Proof.
+ induction st; simpl; try_simplify_someHyps.
+ autodestruct; intros; econstructor; eauto.
+ autodestruct; eauto.
+Qed.
+
+
+(** * Model of Symbolic Execution with Continuation Passing Style
+
+Parameter [k] is the continuation, i.e. the [sstate] construction that will be applied in each execution branch.
+Its input parameter is the symbolic internal state of the branch.
+
+*)
+
+Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate :=
+ match ib with
+ | BF fin _ => Sfinal (tr_sis f fin sis) (sexec_final_sfv fin sis)
+ (* basic instructions *)
+ | Bnop _ => k sis
+ | Bop op args res _ => k (sexec_op op args res sis)
+ | Bload trap chunk addr args dst _ => k (sexec_load trap chunk addr args dst sis)
+ | Bstore chunk addr args src _ => k (sexec_store chunk addr args src sis)
+ (* composed instructions *)
+ | Bseq ib1 ib2 =>
+ sexec_rec f ib1 sis (fun sis2 => sexec_rec f ib2 sis2 k)
+ | Bcond cond args ifso ifnot _ =>
+ let args := list_sval_inj (List.map sis.(si_sreg) args) in
+ let ifso := sexec_rec f ifso sis k in
+ let ifnot := sexec_rec f ifnot sis k in
+ Scond cond args ifso ifnot
+ end
+ .
+
+Definition sexec f ib := sexec_rec f ib sis_init (fun _ => Sabort).
+
+Local Hint Constructors sem_sstate: core.
+Local Hint Resolve sexec_op_correct sexec_final_sfv_correct tr_sis_regs_correct_aux tr_sis_regs_correct
+ sexec_load_correct sexec_store_correct sis_init_correct: core.
+
+Lemma sexec_rec_correct ctx stk t s ib rs m rs1 m1 ofin
+ (ISTEP: iblock_istep (cge ctx) (csp ctx) rs m ib rs1 m1 ofin): forall sis k
+ (SIS: sem_sistate ctx sis rs m)
+ (CONT: match ofin with
+ | None => forall sis', sem_sistate ctx sis' rs1 m1 -> sem_sstate ctx stk t s (k sis')
+ | Some fin => final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs1 m1 fin t s
+ end),
+ sem_sstate ctx stk t s (sexec_rec (cf ctx) ib sis k).
+Proof.
+ induction ISTEP; simpl; try autodestruct; eauto.
+ (* final value *)
+ intros; econstructor; eauto.
+ rewrite str_tr_regs_equiv; eauto.
+ (* condition *)
+ all: intros;
+ eapply sem_Scond; eauto; [
+ erewrite eval_scondition_eq; eauto |
+ replace (if b then sexec_rec (cf ctx) ifso sis k else sexec_rec (cf ctx) ifnot sis k) with (sexec_rec (cf ctx) (if b then ifso else ifnot) sis k);
+ try autodestruct; eauto ].
+Qed.
+
+
+(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec])
+ (sexec is a correct over-approximation)
+*)
+Theorem sexec_correct ctx stk ib t s:
+ iblock_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s ->
+ sem_sstate ctx stk t s (sexec (cf ctx) ib).
+Proof.
+ destruct 1 as (rs' & m' & fin & ISTEP & FSTEP).
+ eapply sexec_rec_correct; simpl; eauto.
+Qed.
+
+(* Remark that we want to reason modulo "extensionality" wrt Regmap.get about regsets.
+ And, nothing in their representation as (val * PTree.t val) enforces that
+ (forall r, rs1#r = rs2#r) -> rs1 = rs2
+*)
+Lemma sem_sistate_tr_sis_determ ctx sis rs1 m1 fi rs2 m2:
+ sem_sistate ctx sis rs1 m1 ->
+ sem_sistate ctx (tr_sis (cf ctx) fi sis) rs2 m2 ->
+ (forall r, rs2#r = (tr_regs (cf ctx) fi rs1)#r)
+ /\ m1 = m2.
+Proof.
+ intros H1 H2.
+ lapply (tr_sis_regs_correct_aux ctx fi sis rs1 m1); eauto.
+ intros X.
+ destruct H1 as (_&MEM1&REG1).
+ destruct H2 as (_&MEM2&REG2); simpl in *.
+ intuition try congruence.
+ cut (Some rs2 # r = Some (tr_regs (cf ctx) fi rs1)#r).
+ { congruence. }
+ rewrite <- REG2, X. auto.
+Qed.
+
+Local Hint Constructors equiv_stackframe list_forall2: core.
+Local Hint Resolve regmap_setres_eq equiv_stack_refl equiv_stack_refl: core.
+
+Lemma sem_sfval_equiv rs1 rs2 ctx stk sfv m t s:
+ sem_sfval ctx stk sfv rs1 m t s ->
+ (forall r, (str_regs (cf ctx) sfv rs1)#r = (str_regs (cf ctx) sfv rs2)#r) ->
+ exists s', sem_sfval ctx stk sfv rs2 m t s' /\ equiv_state s s'.
+Proof.
+ unfold str_regs.
+ destruct 1; simpl in *; intros; subst; eexists; split; econstructor; eauto; try congruence.
+Qed.
+
+Definition abort_sistate ctx (sis: sistate): Prop :=
+ ~(sis.(si_pre) ctx)
+ \/ eval_smem ctx sis.(si_smem) = None
+ \/ exists (r: reg), eval_sval ctx (sis.(si_sreg) r) = None.
+
+Lemma set_sreg_preserves_abort ctx sv dst sis:
+ abort_sistate ctx sis ->
+ abort_sistate ctx (set_sreg dst sv sis).
+Proof.
+ unfold abort_sistate; simpl; intros [PRE|[MEM|REG]]; try tauto.
+ destruct REG as [r REG].
+ destruct (Pos.eq_dec dst r) as [TEST|TEST] eqn: HTEST.
+ - subst; rewrite REG; tauto.
+ - right. right. eexists; rewrite HTEST. auto.
+Qed.
+
+Lemma sexec_op_preserves_abort ctx op args dest sis:
+ abort_sistate ctx sis
+ -> abort_sistate ctx (sexec_op op args dest sis).
+Proof.
+ intros; eapply set_sreg_preserves_abort; eauto.
+Qed.
+
+Lemma sexec_load_preserves_abort ctx chunk addr args dest sis trap:
+ abort_sistate ctx sis
+ -> abort_sistate ctx (sexec_load trap chunk addr args dest sis).
+Proof.
+ intros; eapply set_sreg_preserves_abort; eauto.
+Qed.
+
+Lemma set_smem_preserves_abort ctx sm sis:
+ abort_sistate ctx sis ->
+ abort_sistate ctx (set_smem sm sis).
+Proof.
+ unfold abort_sistate; simpl; try tauto.
+Qed.
+
+Lemma sexec_store_preserves_abort ctx chunk addr args src sis:
+ abort_sistate ctx sis
+ -> abort_sistate ctx (sexec_store chunk addr args src sis).
+Proof.
+ intros; eapply set_smem_preserves_abort; eauto.
+Qed.
+
+Lemma sem_sistate_tr_sis_exclude_abort ctx sis fi rs m:
+ sem_sistate ctx (tr_sis (cf ctx) fi sis) rs m ->
+ abort_sistate ctx sis ->
+ False.
+Proof.
+ intros ((PRE1 & PRE2) & MEM & REG); simpl in *.
+ intros [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; try congruence.
+Qed.
+
+Local Hint Resolve sexec_op_preserves_abort sexec_load_preserves_abort
+ sexec_store_preserves_abort sem_sistate_tr_sis_exclude_abort: core.
+
+Lemma sexec_exclude_abort ctx stk ib t s1: forall sis k
+ (SEXEC: sem_sstate ctx stk t s1 (sexec_rec (cf ctx) ib sis k))
+ (CONT: forall sis', sem_sstate ctx stk t s1 (k sis') -> (abort_sistate ctx sis') -> False)
+ (ABORT: abort_sistate ctx sis),
+ False.
+Proof.
+ induction ib; simpl; intros; eauto.
+ - (* final *) inversion SEXEC; subst; eauto.
+ - (* seq *)
+ eapply IHib1; eauto.
+ simpl. eauto.
+ - (* cond *)
+ inversion SEXEC; subst; eauto. clear SEXEC.
+ destruct b; eauto.
+Qed.
+
+Lemma set_sreg_abort ctx dst sv sis rs m:
+ sem_sistate ctx sis rs m ->
+ (eval_sval ctx sv = None) ->
+ abort_sistate ctx (set_sreg dst sv sis).
+Proof.
+ intros (PRE&MEM&REG) NEW.
+ unfold sem_sistate, abort_sistate; simpl.
+ right; right.
+ exists dst; destruct (Pos.eq_dec dst dst); simpl; try congruence.
+Qed.
+
+Lemma sexec_op_abort ctx sis op args dest rs m
+ (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = None)
+ (SIS: sem_sistate ctx sis rs m)
+ : abort_sistate ctx (sexec_op op args dest sis).
+Proof.
+ eapply set_sreg_abort; eauto.
+ simpl. destruct SIS as (PRE&MEM&REG).
+ erewrite eval_list_sval_inj; simpl; auto.
+ try_simplify_someHyps.
+ intros; erewrite op_valid_pointer_eq; eauto.
+Qed.
+
+Lemma sexec_load_TRAP_abort ctx chunk addr args dst sis rs m
+ (EVAL: forall a, eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a -> Mem.loadv chunk m a = None)
+ (SIS: sem_sistate ctx sis rs m)
+ : abort_sistate ctx (sexec_load TRAP chunk addr args dst sis).
+Proof.
+ eapply set_sreg_abort; eauto.
+ simpl. destruct SIS as (PRE&MEM&REG).
+ erewrite eval_list_sval_inj; simpl; auto.
+ intros; autodestruct; try_simplify_someHyps.
+Qed.
+
+Lemma set_smem_abort ctx sm sis rs m:
+ sem_sistate ctx sis rs m ->
+ eval_smem ctx sm = None ->
+ abort_sistate ctx (set_smem sm sis).
+Proof.
+ intros (PRE&MEM&REG) NEW.
+ unfold abort_sistate; simpl.
+ tauto.
+Qed.
+
+Lemma sexec_store_abort ctx chunk addr args src sis rs m
+ (EVAL: forall a, eval_addressing (cge ctx) (csp ctx) addr rs ## args = Some a -> Mem.storev chunk m a (rs # src) = None)
+ (SIS: sem_sistate ctx sis rs m)
+ :abort_sistate ctx (sexec_store chunk addr args src sis).
+Proof.
+ eapply set_smem_abort; eauto.
+ simpl. destruct SIS as (PRE&MEM&REG).
+ erewrite eval_list_sval_inj; simpl; auto.
+ try_simplify_someHyps.
+ intros; rewrite REG; autodestruct; try_simplify_someHyps.
+Qed.
+
+Local Hint Resolve sexec_op_abort sexec_load_TRAP_abort sexec_store_abort sexec_final_sfv_exact: core.
+
+Lemma sexec_rec_exact ctx stk ib t s1: forall sis k
+ (SEXEC: sem_sstate ctx stk t s1 (sexec_rec (cf ctx) ib sis k))
+ rs m
+ (SIS: sem_sistate ctx sis rs m)
+ (CONT: forall sis', sem_sstate ctx stk t s1 (k sis') -> (abort_sistate ctx sis') -> False)
+ ,
+ match iblock_istep_run (cge ctx) (csp ctx) ib rs m with
+ | Some (out rs' m' (Some fin)) =>
+ exists s2, final_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) rs' m' fin t s2 /\ equiv_state s1 s2
+ | Some (out rs' m' None) => exists sis', (sem_sstate ctx stk t s1 (k sis')) /\ (sem_sistate ctx sis' rs' m')
+ | None => False
+ end.
+Proof.
+ induction ib; simpl; intros; eauto.
+ - (* final *)
+ inv SEXEC.
+ exploit (sem_sistate_tr_sis_determ ctx sis rs m fi); eauto.
+ intros (REG&MEM); subst.
+ exploit (sem_sfval_equiv rs0 rs); eauto.
+ * intros; rewrite REG, str_tr_regs_equiv; auto.
+ * intros (s2 & EQUIV & SFV'); eauto.
+ - (* Bop *) autodestruct; eauto.
+ - destruct trap.
+ + repeat autodestruct.
+ { eexists; split; eauto.
+ eapply sexec_load_correct; eauto.
+ econstructor; eauto. }
+ all:
+ intros; eapply CONT; eauto;
+ eapply sexec_load_TRAP_abort; eauto;
+ intros; try_simplify_someHyps.
+ + repeat autodestruct;
+ eexists; split; eauto;
+ eapply sexec_load_correct; eauto;
+ try (econstructor; eauto; fail).
+ all: eapply has_loaded_default; auto; try_simplify_someHyps.
+ - repeat autodestruct; eauto.
+ all: intros; eapply CONT; eauto;
+ eapply sexec_store_abort; eauto;
+ intros; try_simplify_someHyps.
+ - (* Bseq *)
+ exploit IHib1; eauto. clear sis SEXEC SIS.
+ { simpl; intros; eapply sexec_exclude_abort; eauto. }
+ destruct (iblock_istep_run _ _ _ _ _) eqn: ISTEP; auto.
+ destruct o.
+ destruct _fin eqn: OFIN; simpl; eauto.
+ intros (sis1 & SEXEC1 & SIS1).
+ exploit IHib2; eauto.
+ - (* Bcond *)
+ inv SEXEC.
+ erewrite eval_scondition_eq in SEVAL; eauto.
+ rewrite SEVAL.
+ destruct b.
+ + exploit IHib1; eauto.
+ + exploit IHib2; eauto.
+Qed.
+
+
+(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution
+ (sexec is exact).
+*)
+Theorem sexec_exact ctx stk ib t s1:
+ sem_sstate ctx stk t s1 (sexec (cf ctx) ib) ->
+ exists s2, iblock_step tr_inputs (cge ctx) stk (cf ctx) (csp ctx) (crs0 ctx) (cm0 ctx) ib t s2
+ /\ equiv_state s1 s2.
+Proof.
+ intros; exploit sexec_rec_exact; eauto.
+ { intros sis' SEXEC; inversion SEXEC. }
+ repeat autodestruct; simpl; try tauto.
+ - intros D1 D2 ISTEP (s2 & FSTEP & EQSTEP); subst.
+ eexists; split; eauto.
+ repeat eexists; eauto.
+ erewrite iblock_istep_run_equiv; eauto.
+ - intros D1 D2 ISTEP (sis & SEXEC & _); subst.
+ inversion SEXEC.
+Qed.
+
+(** * High-Level specification of the symbolic simulation test as predicate [symbolic_simu] *)
+
+Record simu_proof_context {f1 f2: BTL.function} := Sctx {
+ sge1: BTL.genv;
+ sge2: BTL.genv;
+ sge_match: forall s, Genv.find_symbol sge1 s = Genv.find_symbol sge2 s;
+ ssp: val;
+ srs0: regset;
+ sm0: mem
+}.
+Arguments simu_proof_context: clear implicits.
+
+Definition bctx1 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge1) f1 ctx.(ssp) ctx.(srs0) ctx.(sm0).
+Definition bctx2 {f1 f2} (ctx: simu_proof_context f1 f2):= Bctx ctx.(sge2) f2 ctx.(ssp) ctx.(srs0) ctx.(sm0).
+
+(* NOTE: we need to mix semantical simulation and syntactic definition on [sfval] in order to abstract
+ the [match_states] of BTL_Schedulerproof.
+
+ Indeed, the [match_states] involves [match_function] in [match_stackframe].
+ And, here, we aim to define a notion of simulation for defining [match_function].
+
+ A syntactic definition of the simulation on [sfval] avoids the circularity issue.
+
+*)
+
+Inductive optsv_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (option sval) -> (option sval) -> Prop :=
+ | Ssome_simu sv1 sv2
+ (SIMU:eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2)
+ :optsv_simu ctx (Some sv1) (Some sv2)
+ | Snone_simu: optsv_simu ctx None None
+ .
+
+Inductive svident_simu {f1 f2: function} (ctx: simu_proof_context f1 f2): (sval + ident) -> (sval + ident) -> Prop :=
+ | Sleft_simu sv1 sv2
+ (SIMU:eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2)
+ :svident_simu ctx (inl sv1) (inl sv2)
+ | Sright_simu id1 id2
+ (IDSIMU: id1 = id2)
+ :svident_simu ctx (inr id1) (inr id2)
+ .
+
+Definition bargs_simu {f1 f2: function} (ctx: simu_proof_context f1 f2) (args1 args2: list (builtin_arg sval)): Prop :=
+ eval_list_builtin_sval (bctx1 ctx) args1 = eval_list_builtin_sval (bctx2 ctx) args2.
+
+Inductive sfv_simu {f1 f2} (ctx: simu_proof_context f1 f2): sfval -> sfval -> Prop :=
+ | Sgoto_simu pc: sfv_simu ctx (Sgoto pc) (Sgoto pc)
+ | Scall_simu sig ros1 ros2 args1 args2 r pc
+ (SVID: svident_simu ctx ros1 ros2)
+ (ARGS:eval_list_sval (bctx1 ctx) args1 = eval_list_sval (bctx2 ctx) args2)
+ :sfv_simu ctx (Scall sig ros1 args1 r pc) (Scall sig ros2 args2 r pc)
+ | Stailcall_simu sig ros1 ros2 args1 args2
+ (SVID: svident_simu ctx ros1 ros2)
+ (ARGS:eval_list_sval (bctx1 ctx) args1 = eval_list_sval (bctx2 ctx) args2)
+ :sfv_simu ctx (Stailcall sig ros1 args1) (Stailcall sig ros2 args2)
+ | Sbuiltin_simu ef lba1 lba2 br pc
+ (BARGS: bargs_simu ctx lba1 lba2)
+ :sfv_simu ctx (Sbuiltin ef lba1 br pc) (Sbuiltin ef lba2 br pc)
+ | Sjumptable_simu sv1 sv2 lpc
+ (VAL: eval_sval (bctx1 ctx) sv1 = eval_sval (bctx2 ctx) sv2)
+ :sfv_simu ctx (Sjumptable sv1 lpc) (Sjumptable sv2 lpc)
+ | simu_Sreturn osv1 osv2
+ (OPT:optsv_simu ctx osv1 osv2)
+ :sfv_simu ctx (Sreturn osv1) (Sreturn osv2)
+.
+
+Definition sistate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (sis1 sis2:sistate): Prop :=
+ forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sem_sistate (bctx2 ctx) sis2 rs m.
+
+Record si_ok ctx (sis: sistate): Prop := {
+ OK_PRE: (sis.(si_pre) ctx);
+ OK_SMEM: eval_smem ctx sis.(si_smem) <> None;
+ OK_SREG: forall (r: reg), eval_sval ctx (si_sreg sis r) <> None
+}.
+
+Lemma sem_si_ok ctx sis rs m:
+ sem_sistate ctx sis rs m -> si_ok ctx sis.
+Proof.
+ unfold sem_sistate;
+ econstructor;
+ intuition congruence.
+Qed.
+
+Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2:sstate): Prop :=
+ forall sis1 sfv1, get_soutcome (bctx1 ctx) st1 = Some (sout sis1 sfv1) -> si_ok (bctx1 ctx) sis1 ->
+ exists sis2 sfv2, get_soutcome (bctx2 ctx) st2 = Some (sout sis2 sfv2)
+ /\ sistate_simu ctx sis1 sis2
+ /\ (forall rs m, sem_sistate (bctx1 ctx) sis1 rs m -> sfv_simu ctx sfv1 sfv2)
+ .
+
+Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2).
+
+(* REM. L'approche suivie initialement ne marche pas !!!
+*)
+(*
+Definition sstate_simu {f1 f2} (ctx: simu_proof_context f1 f2) (st1 st2: sstate) :=
+ forall t s1, sem_sstate (bctx1 ctx) t s1 st1 ->
+ exists s2, sem_sstate (bctx2 ctx) t s2 st2 /\ equiv_state s1 s2.
+
+Definition symbolic_simu f1 f2 ib1 ib2: Prop := forall (ctx: simu_proof_context f1 f2), sstate_simu ctx (sexec f1 ib1) (sexec f2 ib2).
+
+Theorem symbolic_simu_correct f1 f2 ib1 ib2:
+ symbolic_simu f1 f2 ib1 ib2 ->
+ forall (ctx: simu_proof_context f1 f2) t s1, iblock_step tr_inputs (sge1 ctx) (sstk1 ctx) f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ib1 t s1 ->
+ exists s2, iblock_step tr_inputs (sge2 ctx) (sstk2 ctx) f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ib2 t s2 /\ equiv_state s1 s2.
+Proof.
+ unfold symbolic_simu, sstate_simu.
+ intros SIMU ctx t s1 STEP1.
+ exploit (sexec_correct (bctx1 ctx)); simpl; eauto.
+ intros; exploit SIMU; eauto.
+ intros (s2 & SEM1 & EQ1).
+ exploit (sexec_exact (bctx2 ctx)); simpl; eauto.
+ intros (s3 & STEP2 & EQ2).
+ clear STEP1; eexists; split; eauto.
+ eapply equiv_state_trans; eauto.
+Qed.
+*)
+
+(** * Preservation properties under a [simu_proof_context] *)
+
+Section SymbValPreserved.
+
+Variable f1 f2: function.
+
+Hypothesis ctx: simu_proof_context f1 f2.
+Local Hint Resolve sge_match: core.
+
+Lemma eval_sval_preserved sv:
+ eval_sval (bctx1 ctx) sv = eval_sval (bctx2 ctx) sv.
+Proof.
+ induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv)
+ (P1 := fun sm => eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm); simpl; auto.
+ + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto.
+ erewrite eval_operation_preserved; eauto.
+ + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto.
+ erewrite eval_addressing_preserved; eauto.
+ destruct (eval_addressing _ _ _ _); auto.
+ rewrite IHsv; auto.
+ + rewrite IHsv; clear IHsv. destruct (eval_sval _ _); auto.
+ rewrite IHsv0; auto.
+ + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto.
+ erewrite eval_addressing_preserved; eauto.
+ destruct (eval_addressing _ _ _ _); auto.
+ rewrite IHsv; clear IHsv. destruct (eval_smem _ _); auto.
+ rewrite IHsv1; auto.
+Qed.
+
+Lemma list_sval_eval_preserved lsv:
+ eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv.
+Proof.
+ induction lsv; simpl; auto.
+ rewrite eval_sval_preserved. destruct (eval_sval _ _); auto.
+ rewrite IHlsv; auto.
+Qed.
+
+Lemma smem_eval_preserved sm:
+ eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm.
+Proof.
+ induction sm; simpl; auto.
+ rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto.
+ erewrite eval_addressing_preserved; eauto.
+ destruct (eval_addressing _ _ _ _); auto.
+ rewrite IHsm; clear IHsm. destruct (eval_smem _ _); auto.
+ rewrite eval_sval_preserved; auto.
+Qed.
+
+Lemma eval_builtin_sval_preserved sv:
+ eval_builtin_sval (bctx1 ctx) sv = eval_builtin_sval (bctx2 ctx) sv.
+Proof.
+ induction sv; simpl; auto.
+ all: try (erewrite eval_sval_preserved by eauto); trivial.
+ all: erewrite IHsv1 by eauto; erewrite IHsv2 by eauto; reflexivity.
+Qed.
+
+Lemma eval_list_builtin_sval_preserved lsv:
+ eval_list_builtin_sval (bctx1 ctx) lsv = eval_list_builtin_sval (bctx2 ctx) lsv.
+Proof.
+ induction lsv; simpl; auto.
+ erewrite eval_builtin_sval_preserved by eauto.
+ erewrite IHlsv by eauto.
+ reflexivity.
+Qed.
+
+Lemma eval_scondition_preserved cond lsv:
+ eval_scondition (bctx1 ctx) cond lsv = eval_scondition (bctx2 ctx) cond lsv.
+Proof.
+ unfold eval_scondition.
+ rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto.
+Qed.
+
+(* additional preservation properties under this additional hypothesis *)
+Hypothesis senv_preserved_BTL: Senv.equiv (sge1 ctx) (sge2 ctx).
+
+Lemma senv_find_symbol_preserved id:
+ Senv.find_symbol (sge1 ctx) id = Senv.find_symbol (sge2 ctx) id.
+Proof.
+ destruct senv_preserved_BTL as (A & B & C). congruence.
+Qed.
+
+Lemma senv_symbol_address_preserved id ofs:
+ Senv.symbol_address (sge1 ctx) id ofs = Senv.symbol_address (sge2 ctx) id ofs.
+Proof.
+ unfold Senv.symbol_address. rewrite senv_find_symbol_preserved.
+ reflexivity.
+Qed.
+
+Lemma eval_builtin_sarg_preserved m: forall bs varg,
+ eval_builtin_sarg (bctx1 ctx) m bs varg ->
+ eval_builtin_sarg (bctx2 ctx) m bs varg.
+Proof.
+ induction 1; simpl.
+ all: try (constructor; auto).
+ - rewrite <- eval_sval_preserved. assumption.
+ - rewrite <- senv_symbol_address_preserved. assumption.
+ - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal.
+Qed.
+
+Lemma eval_builtin_sargs_preserved m lbs vargs:
+ eval_builtin_sargs (bctx1 ctx) m lbs vargs ->
+ eval_builtin_sargs (bctx2 ctx) m lbs vargs.
+Proof.
+ induction 1; constructor; eauto.
+ eapply eval_builtin_sarg_preserved; auto.
+Qed.
+
+End SymbValPreserved.
+
diff --git a/scheduling/BTL_Scheduler.v b/scheduling/BTL_Scheduler.v
new file mode 100644
index 00000000..78c597e0
--- /dev/null
+++ b/scheduling/BTL_Scheduler.v
@@ -0,0 +1,176 @@
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad BTL.
+
+Require Import Errors Linking BTL_SEtheory BTL_SEimpl.
+
+(** External oracle *)
+Axiom scheduler: BTL.function -> BTL.code.
+
+Extract Constant scheduler => "BTLScheduleraux.btl_scheduler".
+
+Definition equiv_input_regs (f1 f2: BTL.function): Prop :=
+ (forall pc, (fn_code f1)!pc = None <-> (fn_code f2)!pc = None)
+ /\ (forall pc ib1 ib2, (fn_code f1)!pc = Some ib1 -> (fn_code f2)!pc = Some ib2 -> ib1.(input_regs) = ib2.(input_regs)).
+
+Lemma equiv_input_regs_union f1 f2:
+ equiv_input_regs f1 f2 -> forall tbl, union_inputs f1 tbl = union_inputs f2 tbl.
+Proof.
+ intros (EQNone & EQSome). induction tbl as [|pc l']; simpl; auto.
+ generalize (EQNone pc) (EQSome pc); clear EQNone EQSome; intros EQN EQS.
+ do 2 autodestruct; intuition; try_simplify_someHyps.
+ intros; exploit EQS; eauto; clear EQS. congruence.
+Qed.
+
+Lemma equiv_input_regs_pre f1 f2 tbl or:
+ equiv_input_regs f1 f2 -> pre_inputs f1 tbl or = pre_inputs f2 tbl or.
+Proof.
+ intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto.
+Qed.
+
+Lemma equiv_input_regs_tr_inputs f1 f2 l oreg rs:
+ equiv_input_regs f1 f2 ->
+ tr_inputs f1 l oreg rs = tr_inputs f2 l oreg rs.
+Proof.
+ intros; unfold tr_inputs; erewrite equiv_input_regs_pre; eauto.
+Qed.
+
+(* a specification of the verification to do on each function *)
+Record match_function (f1 f2: BTL.function): Prop := {
+ 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: fn_entrypoint f1 = fn_entrypoint f2;
+ preserv_inputs: equiv_input_regs f1 f2;
+ symbolic_simu_ok: forall pc ib1, (fn_code f1)!pc = Some ib1 ->
+ exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2);
+}.
+
+Inductive match_fundef: fundef -> fundef -> Prop :=
+ | match_Internal f f': match_function f f' -> match_fundef (Internal f) (Internal f')
+ | match_External ef: match_fundef (External ef) (External ef).
+
+Inductive match_stackframes: stackframe -> stackframe -> Prop :=
+ | match_stackframe_intro
+ res f sp pc rs rs' f'
+ (EQREGS: forall r, rs # r = rs' # r)
+ (TRANSF: match_function f f')
+ : match_stackframes (BTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc rs').
+
+Inductive match_states: state -> state -> Prop :=
+ | match_states_intro
+ st f sp pc rs rs' m st' f'
+ (EQREGS: forall r, rs # r = rs' # r)
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_function f f')
+ : match_states (State st f sp pc rs m) (State st' f' sp pc rs' m)
+ | match_states_call
+ st st' f f' args m
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_fundef f f')
+ : match_states (Callstate st f args m) (Callstate st' f' args m)
+ | match_states_return
+ st st' v m
+ (STACKS: list_forall2 match_stackframes st st')
+ : match_states (Returnstate st v m) (Returnstate st' v m)
+ .
+
+Lemma match_stack_equiv stk1 stk2:
+ list_forall2 match_stackframes stk1 stk2 ->
+ forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
+ list_forall2 match_stackframes stk1 stk3.
+Proof.
+ induction 1; intros stk3 EQ; inv EQ; constructor; eauto.
+ inv H3; inv H; econstructor; eauto.
+ intros; rewrite <- EQUIV; auto.
+Qed.
+
+Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3.
+Proof.
+ Local Hint Resolve match_stack_equiv: core.
+ destruct 1; intros EQ; inv EQ; econstructor; eauto.
+ intros; rewrite <- EQUIV; auto.
+Qed.
+
+Local Open Scope error_monad_scope.
+
+Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res unit :=
+ match lpc with
+ | nil => OK tt
+ | pc :: lpc' =>
+ match (fn_code f1)!pc, (fn_code f2)!pc with
+ | Some ibf1, Some ibf2 =>
+ do _ <- simu_check f1 f2 ibf1 ibf2;
+ check_symbolic_simu_rec f1 f2 lpc'
+ | _, _ => Error (msg "check_symbolic_simu_rec: code tree mismatch")
+ end
+ end.
+
+Definition erase_input_regs f1 f2 :=
+ let code := PTree.map (fun pc ibf =>
+ let oibf := match (fn_code f2)!pc with
+ | None => ibf
+ | Some oibf => oibf
+ end in
+ {| entry:= oibf.(entry); input_regs := ibf.(input_regs); binfo := oibf.(binfo) |}) (fn_code f1) in
+ BTL.mkfunction (fn_sig f2) (fn_params f2) (fn_stacksize f2) code (fn_entrypoint f2).
+
+Lemma erase_input_regs_correct f1 f2 f3:
+ erase_input_regs f1 f2 = f3 ->
+ equiv_input_regs f1 f3.
+Proof.
+ unfold erase_input_regs; simpl; intros.
+ unfold equiv_input_regs; intuition auto.
+ - subst; simpl. rewrite PTree.gmap, H0; simpl; auto.
+ - subst; simpl in H0. rewrite PTree.gmap in H0.
+ destruct ((fn_code f1)!pc) eqn:EQF; simpl in H0; inv H0; auto.
+ - subst; simpl in H1. rewrite PTree.gmap in H1.
+ rewrite H0 in H1; simpl in H1; inv H1; simpl; auto.
+Qed.
+Global Opaque erase_input_regs.
+
+Definition check_symbolic_simu (f1 f2: BTL.function): res unit :=
+ check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))).
+
+Definition transf_function (f: BTL.function) :=
+ let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in
+ let tf' := erase_input_regs f tf in
+ do _ <- check_symbolic_simu f tf';
+ OK tf'.
+
+Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x,
+ check_symbolic_simu_rec f1 f2 lpc = OK x ->
+ forall pc ib1, (fn_code f1)!pc = Some ib1 /\ In pc lpc ->
+ exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2).
+Proof.
+ induction lpc; simpl; intros f1 f2 x X pc ib1 (PC & HIN); try contradiction.
+ destruct HIN; subst.
+ - rewrite PC in X; destruct ((fn_code f2)!pc); monadInv X.
+ exists i; split; auto. destruct x0. eapply simu_check_correct; eauto.
+ - destruct ((fn_code f1)!a), ((fn_code f2)!a); monadInv X.
+ eapply IHlpc; eauto.
+Qed.
+
+Lemma check_symbolic_simu_correct x f1 f2:
+ check_symbolic_simu f1 f2 = OK x ->
+ forall pc ib1, (fn_code f1)!pc = Some ib1 ->
+ exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2).
+Proof.
+ unfold check_symbolic_simu; intros X pc ib1 PC.
+ eapply check_symbolic_simu_rec_correct; intuition eauto.
+ apply PTree.elements_correct in PC. eapply (in_map fst) in PC; auto.
+Qed.
+
+Lemma check_symbolic_simu_input_equiv x f1 f2:
+ transf_function f1 = OK f2 ->
+ check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2.
+Proof.
+ unfold transf_function; intros TRANSF X; monadInv TRANSF.
+ exploit erase_input_regs_correct; eauto.
+Qed.
+
+Definition transf_fundef (f: fundef) : res fundef :=
+ transf_partial_fundef (fun f => transf_function f) f.
+
+Definition transf_program (p: program) : res program :=
+ transform_partial_program transf_fundef p.
diff --git a/scheduling/BTL_Schedulerproof.v b/scheduling/BTL_Schedulerproof.v
new file mode 100644
index 00000000..9a73d278
--- /dev/null
+++ b/scheduling/BTL_Schedulerproof.v
@@ -0,0 +1,434 @@
+Require Import AST Linking Values Maps Globalenvs Smallstep Registers.
+Require Import Coqlib Maps Events Errors Op OptionMonad.
+Require Import RTL BTL BTL_SEtheory.
+Require Import BTLmatchRTL BTL_Scheduler.
+
+Definition match_prog (p tp: program) :=
+ match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
+
+Lemma transf_program_match:
+ forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
+Proof.
+ intros. eapply match_transform_partial_program_contextual; eauto.
+Qed.
+
+Section PRESERVATION.
+
+Variable prog: program.
+Variable tprog: program.
+
+Hypothesis TRANSL: match_prog prog tprog.
+
+Let pge := Genv.globalenv prog.
+Let tpge := Genv.globalenv tprog.
+
+Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s.
+Proof.
+ rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
+Qed.
+
+Lemma senv_preserved:
+ Senv.equiv pge tpge.
+Proof.
+ eapply (Genv.senv_match TRANSL).
+Qed.
+
+Lemma functions_preserved:
+ forall (v: val) (f: fundef),
+ Genv.find_funct pge v = Some f ->
+ exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tpge v = Some tf /\ linkorder cunit prog.
+Proof.
+ intros. exploit (Genv.find_funct_match TRANSL); eauto.
+ intros (cu & tf & A & B & C).
+ repeat eexists; intuition eauto.
+ + unfold incl; auto.
+ + eapply linkorder_refl.
+Qed.
+
+Lemma transf_function_correct f f':
+ transf_function f = OK f' -> match_function f f'.
+Proof.
+ unfold transf_function. intros H.
+ assert (OH: transf_function f = OK f') by auto. monadInv H.
+ econstructor; eauto.
+ eapply check_symbolic_simu_input_equiv; eauto.
+ eapply check_symbolic_simu_correct; eauto.
+Qed.
+
+Lemma transf_fundef_correct f f':
+ transf_fundef f = OK f' -> match_fundef f f'.
+Proof.
+ intros TRANSF; destruct f; simpl; inv TRANSF.
+ + monadInv H0. exploit transf_function_correct; eauto.
+ constructor; auto.
+ + eapply match_External.
+Qed.
+
+Lemma function_ptr_preserved:
+ forall v f,
+ Genv.find_funct_ptr pge v = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf.
+Proof.
+ intros.
+ exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
+Qed.
+
+Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = funsig f.
+Proof.
+ intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto.
+ symmetry; erewrite preserv_fnsig; eauto.
+Qed.
+
+Theorem transf_initial_states:
+ forall s1, initial_state prog s1 ->
+ exists s2, initial_state tprog s2 /\ match_states s1 s2.
+Proof.
+ intros. inv H.
+ exploit function_ptr_preserved; eauto. intros (tf & FIND & TRANSF).
+ eexists. split.
+ - econstructor; eauto.
+ + intros; apply (Genv.init_mem_match TRANSL); eauto.
+ + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
+ symmetry. eapply match_program_main. eauto.
+ + erewrite function_sig_translated; eauto.
+ - constructor; eauto.
+ + constructor.
+ + apply transf_fundef_correct; auto.
+Qed.
+
+Lemma transf_final_states s1 s2 r:
+ match_states s1 s2 -> final_state s1 r -> final_state s2 r.
+Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+Qed.
+
+Section SYMBOLIC_CTX.
+
+Variables f1 f2: function.
+Variable sp0: val.
+Variable rs0: regset.
+Variable m0: Memory.Mem.mem.
+
+Lemma symbols_preserved_rev s: Genv.find_symbol pge s = Genv.find_symbol tpge s.
+Proof.
+ symmetry; apply symbols_preserved.
+Qed.
+
+Let ctx := Sctx f1 f2 pge tpge symbols_preserved_rev sp0 rs0 m0.
+
+Lemma str_regs_equiv sfv1 sfv2 rs:
+ sfv_simu ctx sfv1 sfv2 ->
+ match_function f1 f2 ->
+ str_regs (cf (bctx1 ctx)) sfv1 rs = str_regs (cf (bctx2 ctx)) sfv2 rs.
+Proof.
+ intros SFV MF; inv MF.
+ inv SFV; simpl;
+ erewrite equiv_input_regs_tr_inputs; auto.
+Qed.
+
+Lemma sfind_function_preserved ros1 ros2 fd:
+ svident_simu ctx ros1 ros2 ->
+ sfind_function (bctx1 ctx) ros1 = Some fd ->
+ exists fd',
+ sfind_function (bctx2 ctx) ros2 = Some fd'
+ /\ transf_fundef fd = OK fd'.
+Proof.
+ unfold sfind_function. intros [sv1 sv2 SIMU|]; simpl in *.
+ + rewrite !SIMU. clear SIMU.
+ destruct (eval_sval _ _); try congruence.
+ intros; exploit functions_preserved; eauto.
+ intros (fd' & cunit & (X1 & X2 & X3)). eexists; eauto.
+ + subst. rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence.
+ intros; exploit function_ptr_preserved; eauto.
+Qed.
+
+Theorem symbolic_simu_ssem stk1 stk2 st st' s t:
+ sstate_simu ctx st st' ->
+ sem_sstate (bctx1 ctx) stk1 t s st ->
+ list_forall2 match_stackframes stk1 stk2 ->
+ match_function f1 f2 ->
+ exists s',
+ sem_sstate (bctx2 ctx) stk2 t s' st'
+ /\ match_states s s'.
+Proof.
+ intros SIMU SST1 STACKS MF. inversion MF.
+ exploit sem_sstate_run; eauto.
+ intros (sis1 & sfv1 & rs' & m' & SOUT1 & SIS1 & SF).
+ exploit sem_si_ok; eauto. intros SOK.
+ exploit SIMU; simpl; eauto.
+ intros (sis2 & sfv2 & SOUT2 & SSIMU & SFVSIM).
+ remember SIS1 as EQSFV. clear HeqEQSFV.
+ eapply SFVSIM in EQSFV.
+ exploit SSIMU; eauto. intros SIS2.
+ destruct SIS1 as (PRE1 & SMEM1 & SREGS1).
+ try_simplify_someHyps; intros.
+ set (EQSFV2:=sfv2).
+ inversion EQSFV; subst; inversion SF; subst.
+ - (* goto *)
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + erewrite <- str_regs_equiv; simpl; eauto.
+ + econstructor.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ erewrite equiv_input_regs_tr_inputs; eauto.
+ - (* call *)
+ exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND).
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + erewrite <- str_regs_equiv; eauto.
+ + econstructor; eauto.
+ * apply function_sig_translated; auto.
+ * erewrite <- ARGS; eauto.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ * erewrite equiv_input_regs_tr_inputs; eauto.
+ repeat (econstructor; eauto).
+ * apply transf_fundef_correct; auto.
+ - (* tailcall *)
+ exploit sfind_function_preserved; eauto. intros (fd' & SFIND & FUND).
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + econstructor; eauto.
+ * apply function_sig_translated; auto.
+ * simpl; erewrite <- preserv_fnstacksize; eauto.
+ * erewrite <- ARGS; eauto.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ apply transf_fundef_correct; auto.
+ - (* builtin *)
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + erewrite <- str_regs_equiv; simpl; eauto.
+ + econstructor.
+ * eapply eval_builtin_sargs_preserved.
+ eapply senv_preserved.
+ eapply eval_list_builtin_sval_correct; eauto.
+ inv BARGS.
+ erewrite <- eval_list_builtin_sval_preserved in H0; auto.
+ * simpl in *; eapply external_call_symbols_preserved; eauto.
+ eapply senv_preserved.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ erewrite equiv_input_regs_tr_inputs; eauto.
+ - (* jumptable *)
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + erewrite <- str_regs_equiv; simpl; eauto.
+ + econstructor; eauto.
+ rewrite <- VAL; auto.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ erewrite equiv_input_regs_tr_inputs; eauto.
+ - (* return *)
+ exploit (run_sem_sstate (bctx2 ctx) st' sis2 EQSFV2); eauto.
+ + econstructor; eauto.
+ * simpl; erewrite <- preserv_fnstacksize; eauto.
+ * inv OPT; eauto.
+ rewrite <- SIMU0; auto.
+ + intros SST2. eexists; split; eauto.
+ econstructor; simpl; eauto.
+ Unshelve. all: auto.
+Qed.
+
+Theorem symbolic_simu_correct stk1 stk2 ibf1 ibf2 s t:
+ sstate_simu ctx (sexec f1 ibf1.(entry)) (sexec f2 ibf2.(entry)) ->
+ iblock_step tr_inputs (sge1 ctx) stk1 f1 (ssp ctx) (srs0 ctx) (sm0 ctx) ibf1.(entry) t s ->
+ list_forall2 match_stackframes stk1 stk2 ->
+ match_function f1 f2 ->
+ exists s',
+ iblock_step tr_inputs (sge2 ctx) stk2 f2 (ssp ctx) (srs0 ctx) (sm0 ctx) ibf2.(entry) t s'
+ /\ match_states s s'.
+Proof.
+ unfold sstate_simu.
+ intros SIMU STEP1 STACKS MF.
+ exploit (sexec_correct (bctx1 ctx)); simpl; eauto.
+ intros SST1. exploit symbolic_simu_ssem; eauto.
+ intros (s2 & SST2 & MS).
+ exploit (sexec_exact (bctx2 ctx)); simpl; eauto.
+ intros (s3 & STEP & EQUIV).
+ exists s3. intuition eauto.
+ eapply match_states_equiv; eauto.
+Qed.
+
+End SYMBOLIC_CTX.
+
+Lemma tr_inputs_eqregs_list f tbl: forall rs rs' r' ores,
+ (forall r : positive, rs # r = rs' # r) ->
+ (tr_inputs f tbl ores rs) # r' =
+ (tr_inputs f tbl ores rs') # r'.
+Proof.
+ induction tbl as [|pc tbl IHtbl];
+ intros; destruct ores; simpl;
+ rewrite !tr_inputs_get; autodestruct.
+Qed.
+
+Lemma find_function_eqregs rs rs' ros fd:
+ (forall r : positive, rs # r = rs' # r) ->
+ find_function tpge ros rs = Some fd ->
+ find_function tpge ros rs' = Some fd.
+Proof.
+ intros EQREGS; destruct ros; simpl.
+ rewrite EQREGS; auto.
+ autodestruct.
+Qed.
+
+Lemma args_eqregs (args: list reg): forall (rs rs': regset),
+ (forall r : positive, rs # r = rs' # r) ->
+ rs ## args = rs' ## args.
+Proof.
+ induction args; simpl; trivial.
+ intros rs rs' EQREGS; rewrite EQREGS.
+ apply f_equal; eauto.
+Qed.
+
+Lemma f_arg_eqregs (rs rs': regset):
+ (forall r : positive, rs # r = rs' # r) ->
+ (fun r : positive => rs # r) = (fun r : positive => rs' # r).
+Proof.
+ intros EQREGS; apply functional_extensionality; eauto.
+Qed.
+
+Lemma eqregs_update (rs rs': regset) v dest:
+ (forall r, rs # r = rs' # r) ->
+ (forall r, (rs # dest <- v) # r = (rs' # dest <- v) # r).
+Proof.
+ intros EQREGS r; destruct (Pos.eq_dec dest r); subst.
+ * rewrite !Regmap.gss; reflexivity.
+ * rewrite !Regmap.gso; auto.
+Qed.
+
+Local Hint Resolve equiv_stack_refl tr_inputs_eqregs_list find_function_eqregs eqregs_update: core.
+
+Lemma iblock_istep_eqregs_None sp ib: forall rs m rs' rs1 m1
+ (EQREGS: forall r, rs # r = rs' # r)
+ (ISTEP: iblock_istep tpge sp rs m ib rs1 m1 None),
+ exists rs1',
+ iblock_istep tpge sp rs' m ib rs1' m1 None
+ /\ (forall r, rs1 # r = rs1' # r).
+Proof.
+ induction ib; intros; inv ISTEP.
+ - repeat econstructor; eauto.
+ - repeat econstructor; eauto.
+ erewrite args_eqregs; eauto.
+ - inv LOAD; repeat econstructor; eauto;
+ erewrite args_eqregs; eauto.
+ - repeat econstructor; eauto.
+ erewrite args_eqregs; eauto.
+ rewrite <- EQREGS; eauto.
+ - exploit IHib1; eauto.
+ intros (rs1' & ISTEP1 & EQREGS1).
+ exploit IHib2; eauto.
+ intros (rs2' & ISTEP2 & EQREGS2).
+ eexists; split. econstructor; eauto.
+ eauto.
+ - destruct b.
+ 1: exploit IHib1; eauto.
+ 2: exploit IHib2; eauto.
+ all:
+ intros (rs1' & ISTEP & EQREGS');
+ eexists; split; try eapply exec_cond;
+ try erewrite args_eqregs; eauto.
+Qed.
+
+Lemma iblock_step_eqregs stk sp f ib: forall rs rs' m t s
+ (EQREGS: forall r, rs # r = rs' # r)
+ (STEP: iblock_step tr_inputs tpge stk f sp rs m ib t s),
+ exists s',
+ iblock_step tr_inputs tpge stk f sp rs' m ib t s'
+ /\ equiv_state s s'.
+Proof.
+ induction ib; intros;
+ destruct STEP as (rs2 & m2 & fin & ISTEP & FSTEP); inv ISTEP.
+ - inv FSTEP;
+ eexists; split; repeat econstructor; eauto.
+ + destruct (or); simpl; try rewrite EQREGS; constructor; eauto.
+ + erewrite args_eqregs; eauto. repeat econstructor; eauto.
+ + erewrite args_eqregs; eauto. econstructor; eauto.
+ + erewrite f_arg_eqregs; eauto.
+ + destruct res; simpl; eauto.
+ + rewrite <- EQREGS; auto.
+ - exploit IHib1; eauto.
+ + econstructor. do 2 eexists; split; eauto.
+ + intros (s' & STEP & EQUIV). clear FSTEP.
+ destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP).
+ repeat (econstructor; eauto).
+ - exploit iblock_istep_eqregs_None; eauto.
+ intros (rs1' & ISTEP1 & EQREGS1).
+ exploit IHib2; eauto.
+ + econstructor. do 2 eexists; split; eauto.
+ + intros (s' & STEP & EQUIV). clear FSTEP.
+ destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP).
+ eexists; split; eauto.
+ econstructor; do 2 eexists; split; eauto.
+ eapply exec_seq_continue; eauto.
+ - destruct b.
+ 1: exploit IHib1; eauto.
+ 3: exploit IHib2; eauto.
+ 1,3: econstructor; do 2 eexists; split; eauto.
+ all:
+ intros (s' & STEP & EQUIV); clear FSTEP;
+ destruct STEP as (rs3 & m3 & fin2 & ISTEP & FSTEP);
+ eexists; split; eauto; econstructor;
+ do 2 eexists; split; eauto;
+ eapply exec_cond; try erewrite args_eqregs; eauto.
+Qed.
+
+Lemma iblock_step_simulation stk1 stk2 f1 f2 sp rs rs' m ibf t s1 pc
+ (STEP: iblock_step tr_inputs pge stk1 f1 sp rs m ibf.(entry) t s1)
+ (PC: (fn_code f1) ! pc = Some ibf)
+ (EQREGS: forall r : positive, rs # r = rs' # r)
+ (STACKS: list_forall2 match_stackframes stk1 stk2)
+ (TRANSF: match_function f1 f2)
+ :exists ibf' s2,
+ (fn_code f2) ! pc = Some ibf'
+ /\ iblock_step tr_inputs tpge stk2 f2 sp rs' m ibf'.(entry) t s2
+ /\ match_states s1 s2.
+Proof.
+ exploit symbolic_simu_ok; eauto. intros (ib2 & PC' & SYMBS).
+ exploit symbolic_simu_correct; repeat (eauto; simpl).
+ intros (s2 & STEP2 & MS).
+ exploit iblock_step_eqregs; eauto. intros (s3 & STEP3 & EQUIV).
+ clear STEP2. exists ib2; exists s3. repeat split; auto.
+ eapply match_states_equiv; eauto.
+Qed.
+
+Local Hint Constructors step: core.
+
+Theorem step_simulation s1 s1' t s2
+ (STEP: step tr_inputs pge s1 t s1')
+ (MS: match_states s1 s2)
+ :exists s2',
+ step tr_inputs tpge s2 t s2'
+ /\ match_states s1' s2'.
+Proof.
+ destruct STEP as [stack ibf f sp n rs m t s PC STEP | | | ]; inv MS.
+ - (* iblock *)
+ simplify_someHyps. intros PC.
+ exploit iblock_step_simulation; eauto.
+ intros (ibf' & s2 & PC2 & STEP2 & MS2).
+ eexists; split; eauto.
+ - (* function internal *)
+ inversion TRANSF as [xf tf MF |]; subst.
+ eexists; split.
+ + econstructor. erewrite <- preserv_fnstacksize; eauto.
+ + erewrite preserv_fnparams; eauto.
+ erewrite preserv_entrypoint; eauto.
+ econstructor; eauto.
+ - (* function external *)
+ inv TRANSF. eexists; split; econstructor; eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ - (* return *)
+ inv STACKS. destruct b1 as [res' f' sp' pc' rs'].
+ eexists; split.
+ + econstructor.
+ + inv H1. econstructor; eauto.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (fsem prog) (fsem tprog).
+Proof.
+ eapply forward_simulation_step with match_states; simpl; eauto. (* lock-step with respect to match_states *)
+ - eapply senv_preserved.
+ - eapply transf_initial_states.
+ - eapply transf_final_states.
+ - intros; eapply step_simulation; eauto.
+Qed.
+
+End PRESERVATION.
diff --git a/scheduling/BTLcommonaux.ml b/scheduling/BTLcommonaux.ml
new file mode 100644
index 00000000..577e4828
--- /dev/null
+++ b/scheduling/BTLcommonaux.ml
@@ -0,0 +1,87 @@
+open Maps
+open BTL
+open Registers
+open BTLtypes
+open RTLcommonaux
+
+let undef_node = -1
+
+let mk_iinfo _inumb _pcond = { inumb = _inumb; pcond = _pcond; visited = false; liveins = Regset.empty }
+
+let def_iinfo () = { inumb = undef_node; pcond = None; visited = false; liveins = Regset.empty }
+
+let mk_binfo _bnumb _s_output_regs _typing =
+ {
+ bnumb = _bnumb;
+ visited = false;
+ s_output_regs = _s_output_regs;
+ typing = _typing;
+ }
+
+let reset_visited_ibf btl =
+ PTree.map
+ (fun n ibf ->
+ ibf.binfo.visited <- false;
+ ibf)
+ btl
+
+let reset_visited_ib btl =
+ List.iter
+ (fun (n, ibf) ->
+ let ib = ibf.entry in
+ let rec reset_visited_ib_rec ib =
+ match ib with
+ | Bseq (ib1, ib2) ->
+ reset_visited_ib_rec ib1;
+ reset_visited_ib_rec ib2
+ | Bcond (_, _, ib1, ib2, iinfo) ->
+ reset_visited_ib_rec ib1;
+ reset_visited_ib_rec ib2;
+ iinfo.visited <- false
+ | Bnop (Some iinfo)
+ | Bop (_, _, _, iinfo)
+ | Bload (_, _, _, _, _, iinfo)
+ | Bstore (_, _, _, _, iinfo)
+ | BF (_, iinfo) ->
+ iinfo.visited <- false
+ | _ -> ()
+ in
+ reset_visited_ib_rec ib)
+ (PTree.elements btl);
+ btl
+
+let jump_visit = function
+ | Bcond (_, _, _, _, iinfo)
+ | Bnop (Some iinfo)
+ | Bop (_, _, _, iinfo)
+ | Bload (_, _, _, _, _, iinfo)
+ | Bstore (_, _, _, _, iinfo)
+ | BF (_, iinfo) ->
+ if iinfo.visited then true
+ else (
+ iinfo.visited <- true;
+ false)
+ | Bseq (_, _) -> false
+ | Bnop None -> true
+
+let rec get_inumb_or_next = function
+ | BF (Bgoto s, _) -> p2i s
+ | BF (_, iinfo)
+ | Bnop (Some iinfo)
+ | Bop (_, _, _, iinfo)
+ | Bload (_, _, _, _, _, iinfo)
+ | Bstore (_, _, _, _, iinfo)
+ | Bcond (_, _, _, _, iinfo) ->
+ iinfo.inumb
+ | Bseq (ib1, _) -> get_inumb_or_next ib1
+ | _ -> failwith "get_inumb_or_next: Bnop None"
+
+let get_liveins = function
+ | BF (_, iinfo)
+ | Bnop (Some iinfo)
+ | Bop (_, _, _, iinfo)
+ | Bload (_, _, _, _, _, iinfo)
+ | Bstore (_, _, _, _, iinfo)
+ | Bcond (_, _, _, _, iinfo) ->
+ iinfo.liveins
+ | _ -> failwith "get_liveins: invalid iblock"
diff --git a/scheduling/BTLmatchRTL.v b/scheduling/BTLmatchRTL.v
new file mode 100644
index 00000000..07131893
--- /dev/null
+++ b/scheduling/BTLmatchRTL.v
@@ -0,0 +1,606 @@
+(** General notions about the bisimulation BTL <-> RTL.
+
+This bisimulation is proved on the "CFG semantics" of BTL called [cfgsem] defined below,
+such that the full state of registers is preserved when exiting blocks.
+
+*)
+
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad Lia.
+Require Export BTL.
+
+(* tid = transfer_identity *)
+Definition tid (_:function) (_:list exit) (_: option reg) (rs:regset) := rs.
+
+Definition cfgsem (p: program) :=
+ Semantics (step tid) (initial_state p) final_state (Genv.globalenv p).
+
+(* * Simulation of BTL fsem by BTL cfgsem: a lock-step less-def simulation.
+*)
+
+
+Lemma tr_inputs_lessdef f rs1 rs2 tbl or r
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ :Val.lessdef (tr_inputs f tbl or rs1)#r (tid f tbl or rs2)#r.
+Proof.
+ unfold tid; rewrite tr_inputs_get.
+ autodestruct.
+Qed.
+
+Local Hint Resolve tr_inputs_lessdef init_regs_lessdef_preserv find_function_lessdef regs_lessdef_regs
+ lessdef_state_refl: core.
+Local Hint Unfold regs_lessdef: core.
+
+Lemma fsem2cfgsem_ibistep_simu ge sp rs1 m1 rs1' m1' of ib
+ (ISTEP: iblock_istep ge sp rs1 m1 ib rs1' m1' of): forall
+ rs2 m2
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ (MEMS: Mem.extends m1 m2),
+ exists rs2' m2', iblock_istep_run ge sp ib rs2 m2 = Some (out rs2' m2' of)
+ /\ (forall r, Val.lessdef rs1'#r rs2'#r)
+ /\ (Mem.extends m1' m2').
+Proof.
+ induction ISTEP; simpl; try_simplify_someHyps; intros.
+ - (* Bop *)
+ exploit (@eval_operation_lessdef _ _ ge sp op (rs ## args)); eauto.
+ intros (v1 & EVAL' & LESSDEF).
+ do 2 eexists; rewrite EVAL'. repeat (split; eauto).
+ eapply set_reg_lessdef; eauto.
+ - (* Bload *)
+ inv LOAD.
+ + exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto.
+ intros (v2 & EVAL' & LESSDEF). exploit Mem.loadv_extends; eauto.
+ intros (v3 & LOAD' & LESSDEF'); autodestruct;
+ do 2 eexists; rewrite EVAL', LOAD';
+ repeat (split; eauto); eapply set_reg_lessdef; eauto.
+ + destruct (eval_addressing ge sp addr rs ## args) eqn:EQA;
+ repeat autodestruct; do 2 eexists;
+ repeat (split; eauto); eapply set_reg_lessdef; eauto.
+ - (* Bstore *)
+ exploit (@eval_addressing_lessdef _ _ ge sp addr (rs ## args)); eauto.
+ intros (v2 & EVAL' & LESSDEF). exploit Mem.storev_extends; eauto.
+ intros (v3 & STORE' & LESSDEF').
+ do 2 eexists; rewrite EVAL', STORE'. repeat (split; eauto).
+ - (* Bseq stop *)
+ exploit IHISTEP; eauto. intros (rs2' & m2' & IBIS & REGS' & MEMS').
+ destruct (iblock_istep_run _ _ _ _ _); try congruence.
+ destruct o, _fin; simpl in *; try congruence. eauto.
+ - (* Bseq continue *)
+ exploit IHISTEP1; eauto.
+ clear ISTEP1 REGS MEMS.
+ intros (rs3 & m3 & ISTEP3 & REGS3 & MEMS3).
+ rewrite ISTEP3; simpl. rewrite iblock_istep_run_equiv in ISTEP2.
+ exploit IHISTEP2; eauto.
+ - (* Bcond *)
+ erewrite (@eval_condition_lessdef cond (rs ## args)); eauto.
+Qed.
+
+Local Hint Constructors lessdef_stackframe lessdef_state final_step list_forall2 step: core.
+
+Lemma fsem2cfgsem_finalstep_simu ge stk1 f sp rs1 m1 fin t s1 stk2 rs2 m2
+ (FSTEP: final_step tr_inputs ge stk1 f sp rs1 m1 fin t s1)
+ (STACKS: list_forall2 lessdef_stackframe stk1 stk2)
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ (MEMS: Mem.extends m1 m2)
+ : exists s2, final_step tid ge stk2 f sp rs2 m2 fin t s2
+ /\ lessdef_state s1 s2.
+Proof.
+ destruct FSTEP; try (eexists; split; simpl; econstructor; eauto; fail).
+ - (* return *)
+ exploit Mem.free_parallel_extends; eauto.
+ intros (m2' & FREE & MEMS2).
+ eexists; split; simpl; econstructor; eauto.
+ destruct or; simpl; auto.
+ - (* tailcall *)
+ exploit Mem.free_parallel_extends; eauto.
+ intros (m2' & FREE & MEMS2).
+ eexists; split; simpl; econstructor; eauto.
+ - (* builtin *)
+ exploit (eval_builtin_args_lessdef (ge:=ge) (e1:=(fun r => rs1 # r)) (fun r => rs2 # r)); eauto.
+ intros (vl2' & BARGS & VARGS).
+ exploit external_call_mem_extends; eauto.
+ intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED).
+ eexists; split; simpl; econstructor; eauto.
+ intros; apply set_res_lessdef; eauto.
+ - (* jumptable *)
+ eexists; split; simpl; econstructor; eauto.
+ destruct (REGS arg); try congruence.
+Qed.
+
+Lemma fsem2cfgsem_ibstep_simu ge stk1 stk2 f sp rs1 m1 ib t s1 rs2 m2
+ (STEP: iblock_step tr_inputs ge stk1 f sp rs1 m1 ib t s1)
+ (STACKS : list_forall2 lessdef_stackframe stk1 stk2)
+ (REGS: forall r, Val.lessdef rs1#r rs2#r)
+ (MEMS : Mem.extends m1 m2)
+ : exists s2, iblock_step tid ge stk2 f sp rs2 m2 ib t s2
+ /\ lessdef_state s1 s2.
+Proof.
+ destruct STEP as (rs1' & m1' & fin & ISTEP & FSTEP).
+ exploit fsem2cfgsem_ibistep_simu; eauto.
+ intros (rs2' & m2' & ISTEP2 & REGS2 & MEMS2).
+ rewrite <- iblock_istep_run_equiv in ISTEP2. clear ISTEP REGS MEMS.
+ exploit fsem2cfgsem_finalstep_simu; eauto.
+ intros (s2 & FSTEP2 & LESSDEF). clear FSTEP.
+ unfold iblock_step; eexists; split; eauto.
+Qed.
+
+Local Hint Constructors step: core.
+
+Lemma fsem2cfgsem_step_simu ge s1 t s1' s2
+ (STEP: step tr_inputs ge s1 t s1')
+ (REGS: lessdef_state s1 s2)
+ :exists s2' : state,
+ step tid ge s2 t s2' /\
+ lessdef_state s1' s2'.
+Proof.
+ destruct STEP; inv REGS.
+ - (* iblock *)
+ intros; exploit fsem2cfgsem_ibstep_simu; eauto. clear STEP.
+ intros (s2 & STEP2 & REGS2).
+ eexists; split; eauto.
+ - (* internal call *)
+ exploit (Mem.alloc_extends m m2 0 (fn_stacksize f) stk m' 0 (fn_stacksize f)); eauto.
+ 1-2: lia.
+ intros (m2' & ALLOC2 & MEMS2). clear ALLOC MEMS.
+ eexists; split; econstructor; eauto.
+ - (* external call *)
+ exploit external_call_mem_extends; eauto.
+ intros (vres' & m2' & CALL2 & REGS2 & MEMS2 & UNCHANGED). clear EXTCALL MEMS.
+ eexists; split; econstructor; eauto.
+ - (* return *)
+ inversion STACKS as [|d1 d2 d3 d4 STF2 STK2]. subst.
+ inv STF2.
+ eexists; split; econstructor; eauto.
+ intros; eapply set_reg_lessdef; eauto.
+Qed.
+
+Theorem fsem2cfgsem prog:
+ forward_simulation (fsem prog) (cfgsem prog).
+Proof.
+ eapply forward_simulation_step with lessdef_state; simpl; auto.
+ - destruct 1; intros; eexists; intuition eauto. econstructor; eauto.
+ - intros s1 s2 r REGS FINAL; destruct FINAL.
+ inv REGS; inv STACKS; inv REGS0.
+ econstructor; eauto.
+ - intros; eapply fsem2cfgsem_step_simu; eauto.
+Qed.
+
+(** * Matching BTL (for cfgsem) and RTL code
+
+We define a single verifier able to prove a bisimulation between BTL and RTL code.
+
+Hence, in a sense, our verifier imitates the approach of Duplicate, where [dupmap] maps the BTL nodes to the RTL nodes.
+
+The [match_function] definition gives a "relational" specification of the verifier...
+*)
+
+Require Import Errors.
+
+Inductive match_final_inst (dupmap: PTree.t node): final -> instruction -> Prop :=
+ | mfi_return or: match_final_inst dupmap (Breturn or) (Ireturn or)
+ | mfi_call pc pc' s ri lr r:
+ dupmap!pc = (Some pc') -> match_final_inst dupmap (Bcall s ri lr r pc) (Icall s ri lr r pc')
+ | mfi_tailcall s ri lr:
+ match_final_inst dupmap (Btailcall s ri lr) (Itailcall s ri lr)
+ | mfi_builtin pc pc' ef la br:
+ dupmap!pc = (Some pc') -> match_final_inst dupmap (Bbuiltin ef la br pc) (Ibuiltin ef la br pc')
+ | mfi_jumptable ln ln' r:
+ list_forall2 (fun pc pc' => (dupmap!pc = (Some pc'))) ln ln' ->
+ match_final_inst dupmap (Bjumptable r ln) (Ijumptable r ln')
+.
+
+Inductive is_join_opt {A}: (option A) -> (option A) -> (option A) -> Prop :=
+ | ijo_None_left o: is_join_opt None o o
+ | ijo_None_right o: is_join_opt o None o
+ | ijo_Some x: is_join_opt (Some x) (Some x) (Some x)
+ .
+
+(* [match_iblock dupmap cfg isfst pc ib opc] means that [ib] match a block in a RTL code starting at [pc], with:
+ - [isfst] (in "input") indicates that no step in the surrounding block has been executed before entering [pc]
+ - if [opc] (in "ouput") is [None], this means that all branches of the block ends on a final instruction
+ - if [opc] is [Some pc'], this means that all branches of the block that do not exit, join on [pc'].
+*)
+Inductive match_iblock (dupmap: PTree.t node) (cfg: RTL.code): bool -> node -> iblock -> (option node) -> Prop :=
+ | mib_BF isfst fi pc i iinfo:
+ cfg!pc = Some i ->
+ match_final_inst dupmap fi i ->
+ match_iblock dupmap cfg isfst pc (BF fi iinfo) None
+ | mib_nop_on_rtl isfst pc pc' iinfo:
+ cfg!pc = Some (Inop pc') ->
+ match_iblock dupmap cfg isfst pc (Bnop (Some iinfo)) (Some pc')
+ | mib_nop_skip pc:
+ match_iblock dupmap cfg false pc (Bnop None) (Some pc)
+ | mib_op isfst pc pc' op lr r iinfo:
+ cfg!pc = Some (Iop op lr r pc') ->
+ match_iblock dupmap cfg isfst pc (Bop op lr r iinfo) (Some pc')
+ | mib_load isfst pc pc' trap m a lr r iinfo:
+ cfg!pc = Some (Iload trap m a lr r pc') ->
+ match_iblock dupmap cfg isfst pc (Bload trap m a lr r iinfo) (Some pc')
+ | mib_store isfst pc pc' m a lr r iinfo:
+ cfg!pc = Some (Istore m a lr r pc') ->
+ match_iblock dupmap cfg isfst pc (Bstore m a lr r iinfo) (Some pc')
+ | mib_exit pc pc' iinfo:
+ dupmap!pc = (Some pc') ->
+ match_iblock dupmap cfg false pc' (BF (Bgoto pc) iinfo) None
+ (* NB: on RTL side, we exit the block by a "basic" instruction (or Icond).
+ Thus some step should have been executed before [pc'] in the RTL code *)
+ | mib_seq_Some isfst b1 b2 pc1 pc2 opc:
+ match_iblock dupmap cfg isfst pc1 b1 (Some pc2) ->
+ match_iblock dupmap cfg false pc2 b2 opc ->
+ match_iblock dupmap cfg isfst pc1 (Bseq b1 b2) opc
+(* | mib_seq_None isfst b1 b2 pc:
+ match_iblock dupmap cfg isfst pc b1 None ->
+ match_iblock dupmap cfg isfst (Bseq b1 b2) pc None
+ (* here [b2] is dead code ! Our verifier rejects such dead code!
+ *)
+*)
+ | mib_cond isfst c lr bso bnot pcso pcnot pc opc1 opc2 opc i iinfo:
+ cfg!pc = Some (Icond c lr pcso pcnot i) ->
+ match_iblock dupmap cfg false pcso bso opc1 ->
+ match_iblock dupmap cfg false pcnot bnot opc2 ->
+ is_join_opt opc1 opc2 opc ->
+ match_iblock dupmap cfg isfst pc (Bcond c lr bso bnot iinfo) opc
+ .
+
+Definition match_cfg dupmap (cfg: BTL.code) (cfg':RTL.code): Prop :=
+ forall pc pc', dupmap!pc = Some pc' ->
+ exists ib, cfg!pc = Some ib /\ match_iblock dupmap cfg' true pc' ib.(entry) None.
+
+Record match_function dupmap f f': Prop := {
+ dupmap_correct: match_cfg dupmap (BTL.fn_code f) (RTL.fn_code f');
+ dupmap_entrypoint: dupmap!(fn_entrypoint f) = Some (RTL.fn_entrypoint f');
+ preserv_fnsig: fn_sig f = RTL.fn_sig f';
+ preserv_fnparams: fn_params f = RTL.fn_params f';
+ preserv_fnstacksize: fn_stacksize f = RTL.fn_stacksize f'
+}.
+
+(** * Shared verifier between RTL -> BTL and BTL -> RTL *)
+
+Local Open Scope error_monad_scope.
+
+Definition verify_is_copy dupmap n n' :=
+ match dupmap!n with
+ | None => Error(msg "BTL.verify_is_copy None")
+ | Some revn => match (Pos.compare n' revn) with Eq => OK tt | _ => Error(msg "BTL.verify_is_copy invalid map") end
+ end.
+
+Fixpoint verify_is_copy_list dupmap ln ln' :=
+ match ln with
+ | n::ln => match ln' with
+ | n'::ln' => do _ <- verify_is_copy dupmap n n';
+ verify_is_copy_list dupmap ln ln'
+ | nil => Error (msg "BTL.verify_is_copy_list: ln' bigger than ln") end
+ | nil => match ln' with
+ | n :: ln' => Error (msg "BTL.verify_is_copy_list: ln bigger than ln'")
+ | nil => OK tt end
+ end.
+
+Lemma verify_is_copy_correct dupmap n n' tt:
+ verify_is_copy dupmap n n' = OK tt ->
+ dupmap ! n = Some n'.
+Proof.
+ unfold verify_is_copy; repeat autodestruct.
+ intros NP H; destruct (_ ! n) eqn:REVM; [|inversion H].
+ eapply Pos.compare_eq in NP. congruence.
+Qed.
+Local Hint Resolve verify_is_copy_correct: core.
+
+Lemma verify_is_copy_list_correct dupmap ln: forall ln' tt,
+ verify_is_copy_list dupmap ln ln' = OK tt ->
+ list_forall2 (fun n n' => dupmap ! n = Some n') ln ln'.
+Proof.
+ induction ln.
+ - intros. destruct ln'; monadInv H. constructor.
+ - intros. destruct ln'; monadInv H. constructor; eauto.
+Qed.
+
+(* TODO Copied from duplicate, should we import ? *)
+Lemma product_eq {A B: Type} :
+ (forall (a b: A), {a=b} + {a<>b}) ->
+ (forall (c d: B), {c=d} + {c<>d}) ->
+ forall (x y: A+B), {x=y} + {x<>y}.
+Proof.
+ intros H H'. intros. decide equality.
+Qed.
+
+(* TODO Copied from duplicate, should we import ? *)
+(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application"
+ * error when doing so *)
+Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}.
+Proof.
+ intros.
+ apply (builtin_arg_eq Pos.eq_dec).
+Defined.
+Global Opaque builtin_arg_eq_pos.
+
+(* TODO Copied from duplicate, should we import ? *)
+Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}.
+Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed.
+Global Opaque builtin_res_eq_pos.
+
+Fixpoint verify_block (dupmap: PTree.t node) cfg isfst pc ib : res (option node) :=
+ match ib with
+ | BF fi _ =>
+ match fi with
+ | Bgoto pc1 =>
+ do u <- verify_is_copy dupmap pc1 pc;
+ if negb isfst then
+ OK None
+ else Error (msg "BTL.verify_block: isfst is true Bgoto")
+ | Breturn or =>
+ match cfg!pc with
+ | Some (Ireturn or') =>
+ if option_eq Pos.eq_dec or or' then OK None
+ else Error (msg "BTL.verify_block: different opt reg in Breturn")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Breturn")
+ end
+ | Bcall s ri lr r pc1 =>
+ match cfg!pc with
+ | Some (Icall s' ri' lr' r' pc2) =>
+ do u <- verify_is_copy dupmap pc1 pc2;
+ if (signature_eq s s') then
+ if (product_eq Pos.eq_dec ident_eq ri ri') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK None
+ else Error (msg "BTL.verify_block: different r r' in Bcall")
+ else Error (msg "BTL.verify_block: different lr in Bcall")
+ else Error (msg "BTL.verify_block: different ri in Bcall")
+ else Error (msg "BTL.verify_block: different signatures in Bcall")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bcall")
+ end
+ | Btailcall s ri lr =>
+ match cfg!pc with
+ | Some (Itailcall s' ri' lr') =>
+ if (signature_eq s s') then
+ if (product_eq Pos.eq_dec ident_eq ri ri') then
+ if (list_eq_dec Pos.eq_dec lr lr') then OK None
+ else Error (msg "BTL.verify_block: different lr in Btailcall")
+ else Error (msg "BTL.verify_block: different ri in Btailcall")
+ else Error (msg "BTL.verify_block: different signatures in Btailcall")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Btailcall")
+ end
+ | Bbuiltin ef la br pc1 =>
+ match cfg!pc with
+ | Some (Ibuiltin ef' la' br' pc2) =>
+ do u <- verify_is_copy dupmap pc1 pc2;
+ if (external_function_eq ef ef') then
+ if (list_eq_dec builtin_arg_eq_pos la la') then
+ if (builtin_res_eq_pos br br') then OK None
+ else Error (msg "BTL.verify_block: different brr in Bbuiltin")
+ else Error (msg "BTL.verify_block: different lbar in Bbuiltin")
+ else Error (msg "BTL.verify_block: different ef in Bbuiltin")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bbuiltin")
+ end
+ | Bjumptable r ln =>
+ match cfg!pc with
+ | Some (Ijumptable r' ln') =>
+ do u <- verify_is_copy_list dupmap ln ln';
+ if (Pos.eq_dec r r') then OK None
+ else Error (msg "BTL.verify_block: different r in Bjumptable")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bjumptable")
+ end
+ end
+ | Bnop oiinfo =>
+ match oiinfo with
+ | Some _ =>
+ match cfg!pc with
+ | Some (Inop pc') => OK (Some pc')
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bnop")
+ end
+ | None =>
+ if negb isfst then OK (Some pc)
+ else Error (msg "BTL.verify_block: isfst is true Bnop (on_rtl is false)")
+ end
+ | Bop op lr r _ =>
+ match cfg!pc with
+ | Some (Iop op' lr' r' pc') =>
+ if (eq_operation op op') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then
+ OK (Some pc')
+ else Error (msg "BTL.verify_block: different r in Bop")
+ else Error (msg "BTL.verify_block: different lr in Bop")
+ else Error (msg "BTL.verify_block: different operations in Bop")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bop")
+ end
+ | Bload tm m a lr r _ =>
+ match cfg!pc with
+ | Some (Iload tm' m' a' lr' r' pc') =>
+ if (trapping_mode_eq tm tm') then
+ if (chunk_eq m m') then
+ if (eq_addressing a a') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then
+ OK (Some pc')
+ else Error (msg "BTL.verify_block: different r in Bload")
+ else Error (msg "BTL.verify_block: different lr in Bload")
+ else Error (msg "BTL.verify_block: different addressing in Bload")
+ else Error (msg "BTL.verify_block: different chunk in Bload")
+ else Error (msg "BTL.verify_block: different trapping_mode in Bload")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bload")
+ end
+ | Bstore m a lr r _ =>
+ match cfg!pc with
+ | Some (Istore m' a' lr' r' pc') =>
+ if (chunk_eq m m') then
+ if (eq_addressing a a') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK (Some pc')
+ else Error (msg "BTL.verify_block: different r in Bstore")
+ else Error (msg "BTL.verify_block: different lr in Bstore")
+ else Error (msg "BTL.verify_block: different addressing in Bstore")
+ else Error (msg "BTL.verify_block: different chunk in Bstore")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bstore")
+ end
+ | Bseq b1 b2 =>
+ do opc <- verify_block dupmap cfg isfst pc b1;
+ match opc with
+ | Some pc' =>
+ verify_block dupmap cfg false pc' b2
+ | None => Error (msg "BTL.verify_block: None next pc in Bseq (deadcode)")
+ end
+ | Bcond c lr bso bnot _ =>
+ match cfg!pc with
+ | Some (Icond c' lr' pcso pcnot i') =>
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (eq_condition c c') then
+ do opc1 <- verify_block dupmap cfg false pcso bso;
+ do opc2 <- verify_block dupmap cfg false pcnot bnot;
+ match opc1, opc2 with
+ | None, o => OK o
+ | o, None => OK o
+ | Some x, Some x' =>
+ if Pos.eq_dec x x' then OK (Some x)
+ else Error (msg "BTL.verify_block: is_join_opt incorrect for Bcond")
+ end
+ else Error (msg "BTL.verify_block: incompatible conditions in Bcond")
+ else Error (msg "BTL.verify_block: different lr in Bcond")
+ | _ => Error (msg "BTL.verify_block: incorrect cfg Bcond")
+ end
+ end.
+
+(* This property expresses that "relation" [match_iblock] is a partial function (see also [iblock_istep_run_equiv] above) *)
+Lemma verify_block_correct dupmap cfg ib: forall pc isfst fin,
+ verify_block dupmap cfg isfst pc ib = (OK fin) -> match_iblock dupmap cfg isfst pc ib fin.
+Proof.
+ induction ib; intros;
+ try (unfold verify_block in H; destruct (cfg!pc) eqn:EQCFG; [ idtac | discriminate; fail ]).
+ - (* BF *)
+ destruct fi; unfold verify_block in H.
+ + (* Bgoto *)
+ monadInv H.
+ destruct (isfst); simpl in EQ0; inv EQ0.
+ eapply verify_is_copy_correct in EQ.
+ constructor; assumption.
+ + (* Breturn *)
+ destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate.
+ destruct (option_eq _ _ _); try discriminate. inv H.
+ eapply mib_BF; eauto. constructor.
+ + (* Bcall *)
+ destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate.
+ monadInv H.
+ eapply verify_is_copy_correct in EQ.
+ destruct (signature_eq _ _); try discriminate.
+ destruct (product_eq _ _ _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. subst.
+ inv EQ0. eapply mib_BF; eauto. constructor; assumption.
+ + (* Btailcall *)
+ destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate.
+ destruct (signature_eq _ _); try discriminate.
+ destruct (product_eq _ _ _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate. subst.
+ inv H. eapply mib_BF; eauto. constructor.
+ + (* Bbuiltin *)
+ destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate.
+ monadInv H.
+ eapply verify_is_copy_correct in EQ.
+ destruct (external_function_eq _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (builtin_res_eq_pos _ _); try discriminate. subst.
+ inv EQ0. eapply mib_BF; eauto. constructor; assumption.
+ + (* Bjumptable *)
+ destruct (cfg!pc) eqn:EQCFG; try destruct i; try discriminate.
+ monadInv H.
+ eapply verify_is_copy_list_correct in EQ.
+ destruct (Pos.eq_dec _ _); try discriminate. subst.
+ inv EQ0. eapply mib_BF; eauto. constructor; assumption.
+ - (* Bnop *)
+ destruct oiinfo; simpl in *.
+ + destruct (cfg!pc) eqn:EQCFG; try discriminate.
+ destruct i0; inv H. constructor; assumption.
+ + destruct isfst; try discriminate. inv H.
+ constructor; assumption.
+ - (* Bop *)
+ destruct i; try discriminate.
+ destruct (eq_operation _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. inv H.
+ constructor; assumption.
+ - (* Bload *)
+ destruct i; try discriminate.
+ destruct (trapping_mode_eq _ _); try discriminate.
+ destruct (chunk_eq _ _); try discriminate.
+ destruct (eq_addressing _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. inv H.
+ constructor; assumption.
+ - (* Bstore *)
+ destruct i; try discriminate.
+ destruct (chunk_eq _ _); try discriminate.
+ destruct (eq_addressing _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. inv H.
+ constructor; assumption.
+ - (* Bseq *)
+ monadInv H.
+ destruct x; try discriminate.
+ eapply mib_seq_Some.
+ eapply IHib1; eauto.
+ eapply IHib2; eauto.
+ - (* Bcond *)
+ destruct i; try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (eq_condition _ _); try discriminate.
+ fold (verify_block dupmap cfg false n ib1) in H.
+ fold (verify_block dupmap cfg false n0 ib2) in H.
+ monadInv H.
+ destruct x, x0; try destruct (Pos.eq_dec _ _); try discriminate.
+ all: inv EQ2; eapply mib_cond; eauto; econstructor.
+Qed.
+Local Hint Resolve verify_block_correct: core.
+
+Fixpoint verify_blocks dupmap (cfg: code) (cfg':RTL.code) l: res unit :=
+ match l with
+ | nil => OK tt
+ | (pc, pc')::l =>
+ match cfg!pc with
+ | Some ib => do o <- verify_block dupmap cfg' true pc' ib.(entry);
+ match o with
+ | None => verify_blocks dupmap cfg cfg' l
+ | _ => Error(msg "BTL.verify_blocks.end")
+ end
+ | _ => Error(msg "BTL.verify_blocks.entry")
+ end
+ end.
+
+Definition verify_cfg dupmap (cfg: code) (cfg':RTL.code): res unit :=
+ verify_blocks dupmap cfg cfg' (PTree.elements dupmap).
+
+Lemma verify_cfg_correct dupmap cfg cfg' tt:
+ verify_cfg dupmap cfg cfg' = OK tt ->
+ match_cfg dupmap cfg cfg'.
+Proof.
+ unfold verify_cfg.
+ intros X pc pc' H; generalize X; clear X.
+ exploit PTree.elements_correct; eauto.
+ generalize tt pc pc' H; clear tt pc pc' H.
+ generalize (PTree.elements dupmap).
+ induction l as [|[pc1 pc1']l]; simpl.
+ - tauto.
+ - intros pc pc' DUP u H.
+ unfold bind.
+ repeat autodestruct.
+ intros; subst.
+ destruct H as [H|H]; eauto.
+ inversion H; subst.
+ eexists; split; eauto.
+Qed.
+
+Definition verify_function dupmap f f' : res unit :=
+ do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f');
+ verify_cfg dupmap (fn_code f) (RTL.fn_code f').
+
+Lemma verify_function_correct dupmap f f' tt:
+ verify_function dupmap f f' = OK tt ->
+ fn_sig f = RTL.fn_sig f' ->
+ fn_params f = RTL.fn_params f' ->
+ fn_stacksize f = RTL.fn_stacksize f' ->
+ match_function dupmap f f'.
+Proof.
+ unfold verify_function; intro VERIF. monadInv VERIF.
+ constructor; eauto.
+ eapply verify_cfg_correct; eauto.
+Qed.
+
diff --git a/scheduling/BTLroadmap.md b/scheduling/BTLroadmap.md
new file mode 100644
index 00000000..1514c5d4
--- /dev/null
+++ b/scheduling/BTLroadmap.md
@@ -0,0 +1,454 @@
+# BTL Development Roadmap
+
+BTL aims to be an IR dedicated to defensive certification of middle-end optimizations (before register allocation).
+It provides a CFG of "loop-free" blocks, where each block is run in one step emitting at most a single observational event.
+The "local" optimizations (i.e. preserving "locally" the semantics of such blocks) would be checked by symbolic execution with rewriting rules.
+The main info required from oracles: a "dupmap" mapping block entries (and maybe strategies for controlling path explosion during symbolic execution -- see below).
+Moreover, the "global" optimizations would require some invariants annotations at block entry (provided by oracles).
+
+Examples of optimizations that we aim to support:
+
+ - instruction scheduling
+ - instruction rewritings (instruction selection)
+ - branch duplication, "if-lifting" (e.g. side-exit moved up in "traces").
+ - strength-reduction
+ - SSA optimizations
+
+We sketch below the various kinds of supported optimizations in development order...
+
+Remark that we may port most of the current optimizations from RTL to BTL (even maybe, register allocation).
+However, RTL will probably remain useful for "fine-tuned" duplication that crosses block boundaries (e.g. Duplicate module).
+
+## Introduction
+
+En gros la syntaxe d'un bloc BTL est définie par:
+
+ Inductive iblock: Type :=
+ | ... (* instructions basiques ou "finales" (call, return, goto, etc) *)
+ | Bseq (b1 b2: iblock) (* séquence de deux blocs *)
+ | Bcond (cond:condition) (args:list reg) (ifso ifnot: iblock) (* if-then-else *)
+
+Le modèle de base de l'exécution symbolique représente un tel bloc par un état symbolique de type:
+
+ Inductive sstate :=
+ | Sfinal (sis: sistate) (sfv: sfval)
+ | Scond (cond: condition) (args: list_sval) (ifso ifnot: sstate)
+ | Sabort
+ .
+
+où `sistate` est un PPA (preconditioned parallel assignment) des registres et `sfval` représente un branchement (call, return, goto, etc).
+
+Autrement dit, un état symbolique représente tous les chemins
+d'exécution possibles par une sorte de gros BDD ayant sur les feuilles
+un `Sfinal` (ou un `Sabort` s'il manque une instruction de branchement sur ce chemin).
+
+## Block boundaries, branch duplication or factorization
+
+Two possibility of branch duplications (e.g tail-duplication, loop unrolling, etc):
+
+- during RTL -> BTL (while "building" BTL blocks)
+- during BTL -> BTL. Typically, the "if-lifting" à la Justus could be performed/verified in a single pass here !
+
+Branch factorization should also be possible in BTL -> RTL pass. Example: revert "loop-unrolling".
+
+**IMPLEM NOTE:** a single verifier for RTL -> BTL and BTL -> RTL simulations, with a dupmap: BTL block-entries -> RTL nodes.
+
+
+**CURRENT STATUS**
+
+- verifier: implemented and proved w.r.t match_iblock specification.
+- Proof:
+ - BTL -> RTL: done.
+ - RTL -> BTL: done.
+- Oracles:
+ - BTL -> RTL: TODO.
+ - RTL -> BTL: started.
+
+## Simulation modulo liveness and "Functional" semantics of BTL blocks
+
+L'approche suivie pour réaliser la simulation modulo liveness dans
+RTLpath est compliquée à adapter sur BTL. En effet, un état
+symbolique RTLpath correspond à un état symbolique BTL très
+particulier: toutes les "feuilles" (les `Sfinal`) sont des `Sgoto`
+sauf éventuellement une. Or, dans RTLpath, le traitement de l'info de
+liveness sur cette feuille particulière est très adhoc et laborieux
+(cf. le traitement de `pre_output_regs` dans RTLpathScheduler, etc).
+On n'a pas envie de généraliser cette usine à gaz.
+
+On cherche donc une façon plus abstraite de faire... On a l'idée de
+coder la "simulation modulo liveness" dans une "simulation
+less_def". Ça peut rendre le cadre du test de simulation plus
+simple et plus général.
+
+**Idée_Informelle** à côté de la sémantique "à la RTL" pour BTL
+[BTLmatch.cfgsem], on définit une sémantique [BTL.fsem], où c'est
+juste "l'entrée dans un bloc" qui change de sémantique. Intuitivement,
+cette sémantique considère chaque bloc comme une sorte de
+fonction paramétrée par les `input_regs` et appelée uniquement en
+"tailcall" (via les "goto"). C'est ce qu'on va appeler la "functional
+semantics" de BTL (l'autre étant appelée qqchose comme "CFG semantics"
+?).
+
+Autrement dit, sur l'entrée dans un bloc pour un état (rs,m), on pourrait moralement
+commencer par mettre à Vundef tous les registres qui ne sont pas dans l'`input_regs`.
+
+**NOTE** cette idée de voir les blocs comme des "fonctions" correpond
+bien à la représentation "SSA" à la Appel/MLIR. Donc cette sémantique
+peut servir de base pour un support de formes SSA (partielles ou
+totales) dans BTL. Pour l'encodage de SSA, il suffira d'étendre le
+mécanisme d'initialisation du "rs0" d'un bloc avec un passage de
+paramètres.
+
+**REM** pour le test d'exécution symbolique, il semble plus adapté de
+mettre les Vundef juste à la fin de la transition (précédente) plutôt
+qu'au début (de la suivante): c'est d'ailleurs plus proche de la vision SSA.
+
+En fait, on pourrait faire les deux (dans le détail, ça ne ferait pas
+exactement deux fois la même chose: par exemple, quand on sort sur un
+call ou un builtin, le résultat du call/builtin n'est jamais dans le
+liveout puisqu'il va être écrasé de toute façon pendant la transition,
+mais il peut être -- ou pas -- dans le livein de la transition
+suivante dans ce même bloc).
+
+Avoir une initialisation à Vundef en début de bloc permettrait de
+faire une analyse des expressions mal initialisées - par exemple à
+partir du bloc d'entrée de la fonction. Ça semble complémentaire de
+l'analyse de "liveout". Mais utilisable uniquement dans le cadre d'une
+combinaison "exécution symbolique"+"value analysis" (donc pas tout de suite).
+
+Donc, dans un premier temps, la BTL.fsem encode/abstrait la notion de
+"liveout" pour le test d'exécution symbolique. Les définitions des
+deux sémantiques (symbolique et "functional") se font donc en
+simultanée.
+
+**STATUS**
+
+DONE.
+
+## "Basic" symbolic execution / symbolic simulation
+
+We will implement a "basic" (e.g without rewriting rules) simulation test for BTL.fsem.
+Dans un premier temps: pas besoin de "less_undef lockstep", on peut se contenter de tester l'égalité des registres de chaque côté: les registres hors liveout seront égaux à Vundef de chaque côté.
+
+**STATUS**
+
+- BTL_SEtheory: DONE
+ - model of symbolic execution in Continuation-Passing Style for BTL with "correctness" and "completness" thms wrt BTL.fsem
+ - high-level specification [sstate_simu] of "symbolic simulation" over iblocks (wrt BTL.fsem).
+
+- BTL_SEsimuref.v: DONE (raffinement de la notion d'état symbolique)
+
+- BTL_Schedulerproof: DONE (preuve que les simulations symboliques locales de blocks/blocks garantissent la simulation globale du programme pour BTL.fsem).
+
+**TODO**
+
+implem du hash-consing.
+
+## Port of RTLpath optimizations to BTL
+
+- Generalize superblock scheduling for a notion of "extended-blocks" such that each instruction of a block, except the block entry, has a single predecessor in the block.
+- Port rewriting rules of RTLpath.
+- Justus's "poor man SSA" + if-lifting.
+
+## Efficient comparison (guided by oracles) of "if-then-else" sequences.
+
+Le pb est complexe. Comparer des expressions booleennes contenant juste des variables booleennes est déjà NP-complet, avec "explosion exponentielle" dans le pire cas.
+
+Approche proposée: utiliser un mécanisme de vérification "simple", basée sur une comparaison par execution symbolique de "tous" les chemins d'execution (cf Intro ci-dessus).
+Ça pète exponentiellement dans le pire cas: mais on pourra contrôler ce risque d'explosion par les oracles...
+
+Ci-dessous, on liste quelques "techniques" de collaboration oracle/vérificateur pour contrôler l'explosion des chemins.
+Idée: les conditions comportent une liste d'annotations qui permet le guidage du vérificateur par l'oracle.
+
+**Remarque** déjà sur les superblocs, le test de simulation par
+exécution symbolique est quadratique dans le pire cas. Si on a "n"
+variables live sur les sorties et "m" sorties, la vérif est en
+"n*m". En pratique, le nombre de sorties est limité, notamment à cause
+des contraintes pour respecter la structure de superblocs. En RTLpath,
+malgré ce coût quadratique théorique dans le pire cas, en pratique, le
+vérificateur passe bien à l'échelle sur nos benchmarks.
+
+### Contrôle des "joins internes" dans le bloc.
+
+Si dans le bloc, toute condition a au plus un "predecesseur" (au sens
+du CFG RTL) dans le bloc: alors le nombre de "chemins sémantiques"
+(explorés par l'exécution symbolique) est identique au nombre de
+"branches syntaxiques" (présents dans le code). Une façon simple de
+contrôler l'explosion de l'exécution symbolique est de fabriquer (avec
+les oracles) des blocs avec un nombre borné (petit) de "joins
+internes".
+
+**Exemple d'application: généralisation des superblocks**
+
+On considère le bloc BTL ci-dessous (où les `i*` sont des séquences d'instructions basiques sans branchement):
+
+ i0;
+ if (c1) { i1 } else { i1'; goto pc1 }
+ if (c2) { i2; goto pc2 } else { i2' }
+ if (c3} { i3; goto pc3 } else { i3'; goto pc3' }
+
+Sur un tel bloc, il n'y a aucun "join interne" (l'exécution symbolique est donc linéaire et le test de simulation quadratique).
+Mais représenter en RTLpath un tel bloc nécessite au moins 4 blocks (1 superbloc et 3 basic-blocs):
+
+ i0; c1; i1; c2; i2'; i3'; goto pc3'
+ i1'; goto pc1
+ i2; goto pc2
+ i3; goto pc3
+
+La vérification BTL sur le gros bloc ne prendra à priori pas plus de
+temps que la vérification RTLpath cumulée des 4 "petits" blocs. Mais
+la vérification BTL sera plus *puissante*, puisque que quand on va
+vérifier les chemins d'exécutions correspondant à ceux des 3
+basic-blocs, on aura le `i0` en plus dans l'état symbolique (i.e. un
+"contexte d'exécution" plus précis).
+
+**Autre exemple d'application: le if-lifting à la Justus**
+
+Le superblock suivant:
+
+ y1 = e1(x)
+ x = e2(a)
+ y2 = e3(x)
+ if (c[x]) { goto pc } else { i4; goto pc' }
+
+peut être directement montré équivalent à
+
+ x' = e2(a) // x' un registre "frais" (pas dans les "liveout")
+ if (c[x']) {
+ y1 = e1(x);
+ x = x';
+ y2 = e3(x);
+ goto pc
+ } else {
+ y1 = e1(x);
+ x = x';
+ y2 = e3(x);
+ i4;
+ goto pc'
+ }
+
+Ici, la duplication de branche a donc lieu en BTL.
+
+L'exécution symbolique de ces deux blocs va produire deux BDD comparables (avec comparaison des feuilles modulo liveness).
+
+### Comparaison des BDD (modulo réordonnancement des conditions ?)
+
+On peut avoir envie de montrer que les deux blocs ci-dessous sont équivalents (si les dépendences sur les variables le permettent):
+
+ if (c1) { i1 } else { i1' }
+ if (c2) { i2 } else { i2' }
+
+et
+
+ if (c2) { i2 } else { i2' }
+ if (c1) { i1 } else { i1' }
+
+Ça revient (en gros) à comparer les BDD:
+
+ if (c1) { if (c2) {i1;i2} else {i1;i2'} } else { if (c2) {i1';i2} else {i1';i2'} }
+
+et
+
+ if (c2) { if (c1) {i2;i1} else {i2;i1'} } else { if (c1) {i2';i1} else {i2';i1'} }
+
+Pour gérer ça, on peut faire des "Ordered BDD": l'oracle du **bloc
+transformé** annote (certaines) conditions avec des numéros de façon à
+ce l'exécution symbolique du bloc transformé produise un "BDD" qui
+correspond à celui du bloc d'origine (cf. "Principe"
+ci-dessous). Cependant, il semble difficile d'appliquer complètement
+les techniques de mémoïsation des BDD ayant des booléens sur les
+feuilles. Car ici on veut effectuer une comparaison sur des feuilles
+2 à 2 qui n'est pas une égalité, mais une inclusion !
+
+**Principe du réordonnancement de BDD:** l'exécution symbolique du **bloc transformé** réordonne le BDD de
+façon à respecter la numérotation: un pére doit avoir un numéro inférieur à
+chacun de ses fils (en l'absence de numéro, on ignore les contraintes
+de numérotation entre ce noeud est ses voisins). Exemple ci-dessous
+avec trois conditions distinctes (pour order c1 < c2 < c3):
+
+ if (c3) { if (c1) {f1} else {f1'} } else { if (c2} {f2} else {f2'} }
+
+est réordonné en
+
+ if (c1) { if (c2) { if (c3) {f1} else {f2} } else { if (c3) {f1} else {f2'} } }
+ else { if (c2) { if (c3) {f1'} else {f2} } else { if (c3) {f1'} else {f2'} } }
+
+**rem:** on ajoute ici un undefined behavior potentiel à cause l'execution de c2 quand c3 est vrai.
+Mais si le **bloc d'origine** est simulé par cet état symbolique réordonné, c'est correct.
+Le bloc transformé a juste supprimé un test inutile...
+
+Bon, à voir si le principe ci-dessus est vraiment utile dans toute sa
+généralité. Pour simplifier, on peut aussi restreindre le
+réordonnancement du BDD aux cas où il n'y a pas de test redondant
+inséré (comme dans l'exemple initial).
+
+**Version simplifiée: comparaison des BDD sans réordonnancement des conditions**
+
+Dans un premier temps (jusqu'à ce qu'on trouve une optimisation où ça pourrait être utile): pas de réordonnacement des BDD.
+On autorise juste le BDD du bloc transformé à supprimer des conditions par rapport au bloc d'origine.
+Autrement dit, dans la comparaison récursive de `{if (c) BDD1 BDD2}` avec `{if (c') BDD1' BDD2}'`:
+
+- soit `c=c'` et on compare récursivement `BDD1` avec `BDD1'` et `BDD2` avec `BDD2'`.
+- soit `c<>c'` et on compare récursivement `BDD1` et `BDD2` avec `{if (c') BDD1' BDD2'}`
+
+Ce deuxième cas correspond au fait que le test sur `c` dans le bloc d'origine était inutile!
+
+### Propagation de valeurs symbolique de conditions (et élimination de condition) pendant l'execution symbolique
+
+L'exécution symbolique se propageant en "avant", on peut propager les valeurs des conditions symboliques, qu'on peut utiliser pour éliminer des conditions redondantes
+(et donc limiter l'explosion du nombre de chemin).
+
+Pour rendre ce mécanisme efficace et puissant, on peut guider ce mécanisme de propagation/élimination avec des annotations introduites par les oracles.
+
+- une annotation "bind_to x" qui indique de mémoriser la valeur (soit "true" soit "false" suivant la branche) de la condition symbolique avec le nom "x"
+- une annotation "eval_to b proof" qui indique que la condition s'évalue sur la constante "b", ce qui est prouvable avec la preuve "proof".
+
+Ici on peut imaginer un langage plus ou moins compliqué pour prouver l'évaluation des conditions. La version la plus simple:
+
+- "eq(x)" dit simplement que la condition symbolique est syntaxiquement égale celle appelée "x".
+- "eqnot(x)" dit que c'est la négation.
+
+Dans le cas général, on peut introduire tout un système de réécriture pour éliminer les conditions.
+
+En fait, il serait sans doute intéressant de mettre en place un
+"système de réécriture guidé par oracle" pour toutes les instructions
+BTL. Ça permet de concilier "puissance" de l'exécution symbolique et
+"efficacité". L'exécution symbolique va pouvoir éventuellement faire
+des réécritures compliquées, mais uniquement quand c'est nécessaire.
+
+**Exemple: une "if-conversion" généralisée**
+On aimerait montrer que le bloc d'origine:
+
+ if (c) {
+ x1 = e1
+ x2 = e2
+ y = e
+ x3 = e3
+ } else {
+ x3 = e3'
+ z = e'
+ x1 = e1'
+ x2 = e2'
+ }
+
+est simulable par le bloc transformé:
+
+ x1 = (c?e1:e1')
+ x2 = (c?e2:e2')
+ x3 = (c?e3:e3')
+ if (c) { y = e } else { z = e' }
+
+une solution: ajouter une régle de réécriture `x = (c?e:e')` en `if (c) { x=e } else {x=e'}`
+(attention, ce n'est pas une règle de réécriture sur les valeurs
+symboliques, mais sur du code BTL lui-même, avant l'exécution
+symbolique de ce code).
+
+L'exécution symbolique ouvre alors deux branches `c=true` et
+`c=false` sur `x1 = (c?e1:e1')`, puis la propagation/élimination de la
+condition symbolique `c` simplifie les conditionnelles sur `x2`, `x3` et `y`/`z`.
+Au final, on obtient deux BDD identiques sur cet exemple (sans explosion combinatoire).
+
+**Rem** les mécanismes de propagation/réécritures décrits ci-dessus peuvent être aussi très utile pour la simulation symbolique modulo invariants (cf. ci-dessous) !
+
+## Peut-on coder l'égalité modulo liveness comme un cas simple d'invariant symbolique (et faire la vérif de liveneness par execution symbolique) ?
+
+On peut peut-être simplifier le cadre actuel pour rendre l'intégration des invariants symboliques plus facile !! Une piste à creuser...
+
+Si la sémantique `BTL.fsem` a permis de rendre l'architecture un peu plus modulaire (mais sans vraiment simplifier sur le fond).
+Elle ne simplifie pas vraiment l'intégration de la liveness avec les invariants.
+Une approche qui pourrait réellement simplifier: voir le raisonnement modulo liveness comme un cas particulier d'invariant symbolique.
+
+**Idée** pour tout chemin d'exécution passant de `pc` et `pc'`, on ramène la simulation modulo liveness de `ib1` par `ib2` au test:
+
+ [ib1;str_inputs(pc')] lessdef [str_inputs(pc);ib2]
+
+**Avantages possibles**
+
+- plus de verif de liveness spécifique.
+- plus de sémantique `BTL.fsem`
+- les `str_inputs` ne seraient plus dans l'exécution symbolique, mais dans le test de simulation: ça évite les difficultés sur les jumptables.
+
+Par contre, il faut réintroduire une preuve `lessdef` modulo liveness dans `BTL_Schedulerproof`.
+
+### Plan d'implem. incrémentale (dans une branche)
+
+1. Repartir d'une version de BTL sans info de liveness pour simplifier au maximum!! (ie on supprime la semantique fsem et on repart sur cfgsem).
+2. Introduire la preuve `lessdef` modulo liveness dans `BTL_Schedulerproof` en *supposant* qu'on a le test de simulation symbolique qui nous arrange.
+ L'idée est de dégager des conditions suffisantes sur le test de simulation symbolique.
+3. Réaliser le test de simulation symbolique.
+
+## Symbolic Invariants at block entry
+
+Extends the symbolic simulation test to support invariants at block entry, including global renaming of registers. Case-study: support of strength-reduction.
+
+
+### EXAMPLE OF STRENGTH REDUCTION
+
+On veut passer du code C1:
+
+ init: // inputs: int *t, int n, int s
+ int i=0;
+ loop: // inputs: int *t, int n, int s, int i
+ if (i >= n) goto exit;
+ s += t[i];
+ i += 1;
+ goto loop;
+ exit: // inputs: int *t, int s
+
+au code C2:
+
+ init: // inputs: int *t, int n, int s
+ int *ti = t;
+ int *tn = t+n;
+ loop: // inputs: int *t, int s, int *ti, int *tn
+ if (ti >= tn) goto exit;
+ s += *ti;
+ ti += 1; // i.e. sizeof(int)
+ goto loop;
+ exit; // inputs: int *t, int s
+
+Pour donner la correspondance entre les variables des 2 blocs, pour chaque entrée "pc", on introduit une "fonction" de C1@pc.(inputs) -> C2@pc.(inputs).
+
+Typiquement, cette fonction est codable comme un map associant une valeur symbolique sur les C1@pc.(inputs) aux variables de C2@pc.(inputs). Exemple:
+
+ init: // map vide (identité)
+ loop:
+ ti = t+i
+ tn = t+n
+ exit: // map vide (identité)
+
+Si on note `TRANSFER` cette fonction, alors la vérification par
+exécution symbolique que le bloc `ib1` est simulé par `ib2` modulo
+`TRANSFER` sur l'entrée `pc` se ramène à montrer `ib1[TRANSFER]` avec
+`TRANSFER(pc); ib2` pour la simulation symbolique usuelle.
+
+Ci-dessus `ib1[TRANSFER]` dit qu'on a appliqué `TRANSFER(pc')` sur chacune des sorties `pc'` de `ib1`.
+
+**REM** pour que ce soit correct, il faut sans doute vérifier une condition du style `ok(ib1) => ok(ib1[TRANSFER])`...
+
+### Comment gérer le cas des registres résultat de call et builtin ?
+
+Rem: c'est la gestion de ces cas particuliers qui conduit à la notion de `pre_output_regs` (ou `BTL.pre_input`) dans la verif de la simulation modulo liveness.
+On va retrouver ça dans la verif de la simulation modulo invariants symboliques.
+Mais pas évident à éviter.
+
+Concrètement, on ne pourra pas mettre d'invariant qui porte sur le résultat d'un call ou d'un builtin (à part un pur renommage de registre).
+
+### GENERALISATION (What comes next ?)
+
+Est-ce qu'il ne faut pas envisager une combinaison "execution symbolique + value analysis" ? La value analysis pourrait se faire directement sur la sémantique symbolique (donc en sortie de l'exécution symbolique), avec un mécanisme de mémoïsation (pour éviter les calculs redondants dus à la duplication de termes dans l'état symbolique). Intérêts: la value analysis ne se ferait que sur les registres live. Elle serait aussi plus sans doute un peu simple (par exemple inutile d'avoir un "join": on peut se contenter d'un test d'inclusion sur chacun des chemins).
+
+En gros, on pourrait passer d'invariants purement symboliques à des invariants combinant valeurs symboliques et valeurs abstraites.
+
+## Support of SSA-optimizations
+
+Minimum feature: extends BTL with "register renamings" at exits. This should enable to represent SSA-forms in BTL IR, more or less like in MLIR.
+
+Maximum feature: add also a basic instruction that performs parallel renaming of registers.
+This simple feature would help to represent a very general notion of "partial SSA forms": since they could appear in the middle of a block.
+
+In both case, such register renamings would be forbidden in BTL<->RTL matching (SSA-optimizations could only happen within BTL passes).
+
+## Alias analysis in the symbolic simulation
+
+A REGARDER [papier pointé par Justus](https://vbpf.github.io/assets/prevail-paper.pdf)
diff --git a/scheduling/BTLtoRTL.v b/scheduling/BTLtoRTL.v
new file mode 100644
index 00000000..fc58533d
--- /dev/null
+++ b/scheduling/BTLtoRTL.v
@@ -0,0 +1,26 @@
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad BTL.
+Require Export BTLmatchRTL.
+
+Require Import Errors Linking.
+
+(** External oracle *)
+Axiom btl2rtl: BTL.function -> RTL.code * node * (PTree.t node).
+
+Extract Constant btl2rtl => "BTLtoRTLaux.btl2rtl".
+
+Local Open Scope error_monad_scope.
+
+Definition transf_function (f: BTL.function) : res RTL.function :=
+ let (tcte, dupmap) := btl2rtl f in
+ let (tc, te) := tcte in
+ let f' := RTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in
+ do u <- verify_function dupmap f f';
+ OK f'.
+
+Definition transf_fundef (f: fundef) : res RTL.fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: program) : res RTL.program :=
+ transform_partial_program transf_fundef p.
diff --git a/scheduling/BTLtoRTLaux.ml b/scheduling/BTLtoRTLaux.ml
new file mode 100644
index 00000000..ddec991d
--- /dev/null
+++ b/scheduling/BTLtoRTLaux.ml
@@ -0,0 +1,88 @@
+open Maps
+open BTL
+open RTL
+open Camlcoq
+open RTLcommonaux
+open DebugPrint
+open PrintBTL
+open BTLcommonaux
+open BTLtypes
+open BTLRenumber
+
+let get_inumb iinfo = P.of_int iinfo.inumb
+
+let get_ib_num ib = P.of_int (get_inumb_or_next ib)
+
+let translate_function btl entry =
+ let rtl_code = ref PTree.empty in
+ let rec translate_btl_block ib k =
+ debug "Entering translate_btl_block...\n";
+ print_btl_inst stderr ib;
+ let rtli =
+ match ib with
+ | Bcond (cond, lr, BF (Bgoto s1, _), Bnop None, iinfo) ->
+ Some
+ ( Icond (cond, lr, s1, get_ib_num (get_some k), iinfo.pcond),
+ get_inumb iinfo )
+ | Bcond (_, _, _, _, _) ->
+ failwith "translate_function: unsupported Bcond"
+ | Bseq (ib1, ib2) ->
+ translate_btl_block ib1 (Some ib2);
+ translate_btl_block ib2 None;
+ None
+ | Bnop (Some iinfo) ->
+ Some (Inop (get_ib_num (get_some k)), get_inumb iinfo)
+ | Bnop None ->
+ failwith
+ "translate_function: Bnop None can only be in the right child of \
+ Bcond"
+ | Bop (op, lr, rd, iinfo) ->
+ Some (Iop (op, lr, rd, get_ib_num (get_some k)), get_inumb iinfo)
+ | Bload (trap, chk, addr, lr, rd, iinfo) ->
+ Some
+ ( Iload (trap, chk, addr, lr, rd, get_ib_num (get_some k)),
+ get_inumb iinfo )
+ | Bstore (chk, addr, lr, src, iinfo) ->
+ Some
+ ( Istore (chk, addr, lr, src, get_ib_num (get_some k)),
+ get_inumb iinfo )
+ | BF (Bcall (sign, fn, lr, rd, s), iinfo) ->
+ Some (Icall (sign, fn, lr, rd, s), get_inumb iinfo)
+ | BF (Btailcall (sign, fn, lr), iinfo) ->
+ Some (Itailcall (sign, fn, lr), get_inumb iinfo)
+ | BF (Bbuiltin (ef, lr, rd, s), iinfo) ->
+ Some (Ibuiltin (ef, lr, rd, s), get_inumb iinfo)
+ | BF (Bjumptable (arg, tbl), iinfo) ->
+ Some (Ijumptable (arg, tbl), get_inumb iinfo)
+ | BF (Breturn oreg, iinfo) -> Some (Ireturn oreg, get_inumb iinfo)
+ | BF (Bgoto s, iinfo) -> None
+ in
+ match rtli with
+ | Some (inst, inumb) -> rtl_code := PTree.set inumb inst !rtl_code
+ | None -> ()
+ in
+ List.iter
+ (fun (n, ibf) ->
+ ibf.binfo.visited <- true;
+ let ib = ibf.entry in
+ translate_btl_block ib None)
+ (PTree.elements btl);
+ !rtl_code
+
+let btl2rtl (f : BTL.coq_function) =
+ (*debug_flag := true;*)
+ let btl = f.fn_code in
+ let entry = f.fn_entrypoint in
+ let obne, dm = renumber btl entry in
+ let ordered_btl, new_entry = obne in
+ let rtl = translate_function ordered_btl new_entry in
+ debug "Entry %d\n" (p2i new_entry);
+ debug "BTL Code:\n";
+ print_btl_code stderr ordered_btl;
+ debug "RTL Code: ";
+ print_code rtl;
+ debug "Dupmap:\n";
+ print_ptree print_pint dm;
+ debug "\n";
+ (*debug_flag := false;*)
+ ((rtl, new_entry), dm)
diff --git a/scheduling/BTLtoRTLproof.v b/scheduling/BTLtoRTLproof.v
new file mode 100644
index 00000000..cbdc81bd
--- /dev/null
+++ b/scheduling/BTLtoRTLproof.v
@@ -0,0 +1,405 @@
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad.
+
+Require Import Errors Linking BTLtoRTL.
+
+Require Import Linking.
+
+
+Inductive match_fundef: BTL.fundef -> RTL.fundef -> Prop :=
+ | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f')
+ | match_External ef: match_fundef (External ef) (External ef).
+
+Inductive match_stackframes: BTL.stackframe -> RTL.stackframe -> Prop :=
+ | match_stackframe_intro
+ dupmap res f sp pc rs f' pc'
+ (TRANSF: match_function dupmap f f')
+ (DUPLIC: dupmap!pc = Some pc')
+ : match_stackframes (BTL.Stackframe res f sp pc rs) (RTL.Stackframe res f' sp pc' rs).
+
+Inductive match_states: BTL.state -> RTL.state -> Prop :=
+ | match_states_intro
+ dupmap st f sp pc rs m st' f' pc'
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_function dupmap f f')
+ (DUPLIC: dupmap!pc = Some pc')
+ : match_states (State st f sp pc rs m) (RTL.State st' f' sp pc' rs m)
+ | match_states_call
+ st st' f f' args m
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_fundef f f')
+ : match_states (Callstate st f args m) (RTL.Callstate st' f' args m)
+ | match_states_return
+ st st' v m
+ (STACKS: list_forall2 match_stackframes st st')
+ : match_states (Returnstate st v m) (RTL.Returnstate st' v m)
+ .
+
+Lemma transf_function_correct f f':
+ transf_function f = OK f' -> exists dupmap, match_function dupmap f f'.
+Proof.
+ unfold transf_function; unfold bind. repeat autodestruct.
+ intros H _ _ X. inversion X; subst; clear X.
+ eexists; eapply verify_function_correct; simpl; eauto.
+Qed.
+
+Lemma transf_fundef_correct f f':
+ transf_fundef f = OK f' -> match_fundef f f'.
+Proof.
+ intros TRANSF; destruct f; simpl; monadInv TRANSF.
+ + exploit transf_function_correct; eauto.
+ intros (dupmap & MATCH_F).
+ eapply match_Internal; eauto.
+ + eapply match_External.
+Qed.
+
+Definition match_prog (p: program) (tp: RTL.program) :=
+ match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
+
+Lemma transf_program_match:
+ forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
+Proof.
+ intros. eapply match_transform_partial_program_contextual; eauto.
+Qed.
+
+Section RTL_SIMULATES_BTL.
+
+Variable prog: program.
+Variable tprog: RTL.program.
+
+Hypothesis TRANSL: match_prog prog tprog.
+
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof.
+ rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
+Qed.
+
+Lemma senv_preserved: Senv.equiv ge tge.
+Proof.
+ eapply (Genv.senv_match TRANSL).
+Qed.
+
+Lemma functions_translated (v: val) (f: fundef):
+ Genv.find_funct ge v = Some f ->
+ exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tge v = Some tf /\ linkorder cunit prog.
+Proof.
+ intros. exploit (Genv.find_funct_match TRANSL); eauto.
+ intros (cu & tf & A & B & C).
+ repeat eexists; intuition eauto.
+ + unfold incl; auto.
+ + eapply linkorder_refl.
+Qed.
+
+Lemma function_ptr_translated v f:
+ Genv.find_funct_ptr ge v = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = OK tf.
+Proof.
+ intros.
+ exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
+Qed.
+
+Lemma function_sig_translated f tf: transf_fundef f = OK tf -> RTL.funsig tf = funsig f.
+Proof.
+ intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto.
+ erewrite preserv_fnsig; eauto.
+Qed.
+
+Lemma transf_initial_states s1:
+ initial_state prog s1 ->
+ exists s2, RTL.initial_state tprog s2 /\ match_states s1 s2.
+Proof.
+ intros. inv H.
+ exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF).
+ eexists. split.
+ - econstructor; eauto.
+ + eapply (Genv.init_mem_transf_partial TRANSL); eauto.
+ + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
+ symmetry. eapply match_program_main. eauto.
+ + erewrite function_sig_translated; eauto.
+ - constructor; eauto.
+ constructor.
+ apply transf_fundef_correct; auto.
+Qed.
+
+Lemma transf_final_states s1 s2 r:
+ match_states s1 s2 -> final_state s1 r -> RTL.final_state s2 r.
+Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+Qed.
+
+Lemma find_function_preserved ri rs0 fd
+ (FIND : find_function ge ri rs0 = Some fd)
+ : exists fd', RTL.find_function tge ri rs0 = Some fd'
+ /\ transf_fundef fd = OK fd'.
+Proof.
+ pose symbols_preserved as SYMPRES.
+ destruct ri.
+ + simpl in FIND; apply functions_translated in FIND.
+ destruct FIND as (tf & cunit & TFUN & GFIND & LO).
+ eexists; split. eauto. assumption.
+ + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate.
+ apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF).
+ eexists; split. simpl. rewrite symbols_preserved.
+ rewrite GFS. eassumption. assumption.
+Qed.
+
+(* Inspired from Duplicateproof.v *)
+Lemma list_nth_z_dupmap:
+ forall dupmap ln ln' (pc pc': node) val,
+ list_nth_z ln val = Some pc ->
+ list_forall2 (fun n n' => dupmap!n = Some n') ln ln' ->
+ exists (pc': node),
+ list_nth_z ln' val = Some pc'
+ /\ dupmap!pc = Some pc'.
+Proof.
+ induction ln; intros until val; intros LNZ LFA.
+ - inv LNZ.
+ - inv LNZ. destruct (zeq val 0) eqn:ZEQ.
+ + inv H0. destruct ln'; inv LFA.
+ simpl. exists n. split; auto.
+ + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto.
+Qed.
+
+(* variant of [star RTL.step] but requiring proposition [P] on the [refl] (stutttering) case. *)
+Inductive cond_star_step (P: Prop): RTL.state -> trace -> RTL.state -> Prop :=
+ | css_refl s: P -> cond_star_step P s E0 s
+ | css_plus s1 t s2: plus RTL.step tge s1 t s2 -> cond_star_step P s1 t s2
+ .
+
+Lemma css_plus_trans P Q s0 s1 s2 t:
+ plus RTL.step tge s0 E0 s1 ->
+ cond_star_step P s1 t s2 ->
+ cond_star_step Q s0 t s2.
+Proof.
+ intros PLUS STAR.
+ eapply css_plus.
+ inv STAR; auto.
+ eapply plus_trans; eauto.
+Qed.
+
+Lemma css_E0_trans isfst isfst' s0 s1 s2:
+ cond_star_step (isfst=false) s0 E0 s1 ->
+ cond_star_step (false=isfst') s1 E0 s2 ->
+ cond_star_step (isfst=isfst') s0 E0 s2.
+Proof.
+ intros S1 S2.
+ inversion S1; subst; eauto.
+ inversion S2; subst; eauto.
+ eapply css_plus_trans; eauto.
+Qed.
+
+Lemma css_star P s0 s1 t:
+ cond_star_step P s0 t s1 ->
+ star RTL.step tge s0 t s1.
+Proof.
+ destruct 1.
+ - eapply star_refl; eauto.
+ - eapply plus_star; eauto.
+Qed.
+
+Local Hint Constructors RTL.step match_states: core.
+Local Hint Resolve css_refl plus_one transf_fundef_correct: core.
+
+Lemma iblock_istep_simulation sp dupmap stack' f' rs0 m0 ib rs1 m1 ofin
+ (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 ofin):
+ forall pc0 opc isfst
+ (MIB: match_iblock dupmap (RTL.fn_code f') isfst pc0 ib opc),
+ match ofin with
+ | None => exists pc1,(opc = Some pc1) /\ cond_star_step (isfst = false) (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1)
+ | Some fin =>
+ exists isfst' pc1 iinfo, match_iblock dupmap (RTL.fn_code f') isfst' pc1 (BF fin iinfo) None
+ /\ cond_star_step (isfst = isfst') (RTL.State stack' f' sp pc0 rs0 m0) E0 (RTL.State stack' f' sp pc1 rs1 m1)
+ end.
+Proof.
+ induction IBIS; simpl; intros.
+ - (* exec_final *)
+ assert (X: opc = None). { inv MIB; auto. }
+ subst.
+ repeat eexists; eauto.
+ - (* exec_nop *)
+ inv MIB; eexists; split; econstructor; eauto.
+ - (* exec_op *)
+ inv MIB. exists pc'; split; auto; constructor.
+ apply plus_one. eapply exec_Iop; eauto.
+ erewrite <- eval_operation_preserved; eauto.
+ intros; rewrite symbols_preserved; trivial.
+ - (* exec_load *)
+ inv MIB. exists pc'; split; auto; constructor.
+ apply plus_one. inv LOAD.
+ eapply exec_Iload; eauto.
+ all: try (destruct (eval_addressing _ _ _ _) eqn:EVAL;
+ [ eapply exec_Iload_notrap2 | eapply exec_Iload_notrap1]; eauto).
+ all: erewrite <- eval_addressing_preserved; eauto;
+ intros; rewrite symbols_preserved; trivial.
+ - (* exec_store *)
+ inv MIB. exists pc'; split; auto; constructor.
+ apply plus_one. eapply exec_Istore; eauto.
+ all: erewrite <- eval_addressing_preserved; eauto;
+ intros; rewrite symbols_preserved; trivial.
+ - (* exec_seq_stop *)
+ inv MIB; eauto.
+ - (* exec_seq_continue *)
+ inv MIB.
+ exploit IHIBIS1; eauto.
+ intros (pc1 & EQpc1 & STEP1); inv EQpc1.
+ exploit IHIBIS2; eauto.
+ destruct ofin; simpl.
+ + intros (ifst2 & pc2 & iinfo & M2 & STEP2).
+ repeat eexists; eauto.
+ eapply css_E0_trans; eauto.
+ + intros (pc2 & EQpc2 & STEP2); inv EQpc2.
+ eexists; split; auto.
+ eapply css_E0_trans; eauto.
+ - (* exec_cond *)
+ inv MIB.
+ rename H10 into JOIN. (* is_join_opt opc1 opc2 opc *)
+ destruct b; exploit IHIBIS; eauto.
+ + (* taking ifso *)
+ destruct ofin; simpl.
+ * (* ofin is Some final *)
+ intros (isfst' & pc1 & iinfo' & MI & STAR).
+ repeat eexists; eauto.
+ eapply css_plus_trans; eauto.
+ * (* ofin is None *)
+ intros (pc1 & OPC & PLUS). inv OPC.
+ inv JOIN; eexists; split; eauto.
+ all:
+ eapply css_plus_trans; eauto.
+ + (* taking ifnot *)
+ destruct ofin; simpl.
+ * (* ofin is Some final *)
+ intros (isfst' & pc1 & iinfo' & MI & STAR).
+ repeat eexists; eauto.
+ eapply css_plus_trans; eauto.
+ * (* ofin is None *)
+ intros (pc1 & OPC & PLUS). subst.
+ inv JOIN; eexists; split; eauto.
+ all:
+ eapply css_plus_trans; eauto.
+Qed.
+
+Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 pc1 fin t s
+ (STACKS : list_forall2 match_stackframes stack stack')
+ (TRANSF : match_function dupmap f f')
+ (FS : final_step tid ge stack f sp rs1 m1 fin t s)
+ (i : instruction)
+ (ATpc1 : (RTL.fn_code f') ! pc1 = Some i)
+ (MF : match_final_inst dupmap fin i)
+ : exists s', RTL.step tge (RTL.State stack' f' sp pc1 rs1 m1) t s' /\ match_states s s'.
+Proof.
+ inv MF; inv FS.
+ - (* return *)
+ eexists; split.
+ eapply exec_Ireturn; eauto.
+ erewrite <- preserv_fnstacksize; eauto.
+ econstructor; eauto.
+ - (* call *)
+ rename H7 into FIND.
+ exploit find_function_preserved; eauto.
+ intros (fd' & FIND' & TRANSFU).
+ eexists; split. eapply exec_Icall; eauto.
+ apply function_sig_translated. assumption.
+ repeat (econstructor; eauto).
+ - (* tailcall *)
+ rename H2 into FIND.
+ exploit find_function_preserved; eauto.
+ intros (fd' & FIND' & TRANSFU).
+ eexists; split. eapply exec_Itailcall; eauto.
+ apply function_sig_translated. assumption.
+ erewrite <- preserv_fnstacksize; eauto.
+ repeat (econstructor; eauto).
+ - (* builtin *)
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ eapply exec_Ibuiltin; eauto. eapply eval_builtin_args_preserved; eauto.
+ eapply external_call_symbols_preserved; eauto. eapply senv_preserved.
+ econstructor; eauto.
+ - (* jumptable *)
+ pose symbols_preserved as SYMPRES.
+ exploit list_nth_z_dupmap; eauto. intros (pc'1 & LNZ & REVM).
+ eexists. split.
+ eapply exec_Ijumptable; eauto.
+ econstructor; eauto.
+Qed.
+
+Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 pc0 fin t s
+ (STACKS: list_forall2 match_stackframes stack stack')
+ (TRANSF: match_function dupmap f f')
+ (IBIS: iblock_istep ge sp rs0 m0 ib rs1 m1 (Some fin))
+ (MIB : match_iblock dupmap (RTL.fn_code f') true pc0 ib None)
+ (FS : final_step tid ge stack f sp rs1 m1 fin t s)
+ : exists s', plus RTL.step tge (RTL.State stack' f' sp pc0 rs0 m0) t s' /\ match_states s s'.
+Proof.
+ intros; exploit iblock_istep_simulation; eauto.
+ simpl. intros (isfst' & pc1 & iinfo & MI & STAR). clear IBIS MIB.
+ inv MI.
+ - (* final inst except goto *)
+ exploit final_simu_except_goto; eauto.
+ intros (s' & STEP & MS). eexists; split.
+ + eapply plus_right.
+ eapply css_star; eauto.
+ eapply STEP. econstructor.
+ + eapply MS.
+ - (* goto *)
+ inv FS.
+ inv STAR; try congruence.
+ eexists; split. eauto.
+ econstructor; eauto.
+Qed.
+
+Theorem plus_simulation s1 t s1':
+ step tid ge s1 t s1' ->
+ forall s2 (MS: match_states s1 s2),
+ exists s2',
+ plus RTL.step tge s2 t s2'
+ /\ match_states s1' s2'.
+Proof.
+ destruct 1; intros; inv MS.
+ - eapply dupmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (ib' & FNC & MIB).
+ try_simplify_someHyps. destruct STEP as (rs' & m' & fin & IBIS & FS).
+ intros; exploit iblock_step_simulation; eauto.
+ (* exec_function_internal *)
+ - inversion TRANSF as [dupmap f0 f0' MATCHF|]; subst. eexists. split.
+ + eapply plus_one. apply RTL.exec_function_internal.
+ erewrite <- preserv_fnstacksize; eauto.
+ + erewrite <- preserv_fnparams; eauto.
+ eapply match_states_intro; eauto.
+ apply dupmap_entrypoint. assumption.
+ (* exec_function_external *)
+ - inversion TRANSF as [|]; subst. eexists. split.
+ + eapply plus_one. econstructor.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ + constructor. assumption.
+ (* exec_return *)
+ - inversion STACKS as [|a1 al b1 bl H1 HL]; subst.
+ destruct b1 as [res' f' sp' pc' rs'].
+ eexists. split.
+ + eapply plus_one. constructor.
+ + inv H1. econstructor; eauto.
+Qed.
+
+Theorem transf_program_correct_cfg:
+ forward_simulation (BTLmatchRTL.cfgsem prog) (RTL.semantics tprog).
+Proof.
+ eapply forward_simulation_plus with match_states.
+ - eapply senv_preserved.
+ - eapply transf_initial_states.
+ - eapply transf_final_states.
+ - eapply plus_simulation.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (BTL.fsem prog) (RTL.semantics tprog).
+Proof.
+ eapply compose_forward_simulations.
+ - eapply fsem2cfgsem.
+ - eapply transf_program_correct_cfg.
+Qed.
+
+End RTL_SIMULATES_BTL.
diff --git a/scheduling/BTLtypes.ml b/scheduling/BTLtypes.ml
new file mode 100644
index 00000000..12ca30e8
--- /dev/null
+++ b/scheduling/BTLtypes.ml
@@ -0,0 +1,15 @@
+open Registers
+
+type inst_info = {
+ mutable inumb : int;
+ mutable pcond : bool option;
+ mutable visited : bool;
+ mutable liveins: Regset.t;
+}
+
+type block_info = {
+ mutable bnumb : int;
+ mutable visited : bool;
+ s_output_regs : Regset.t;
+ typing : RTLtyping.regenv;
+}
diff --git a/scheduling/InstructionScheduler.mli b/scheduling/InstructionScheduler.mli
index 48c7bc09..29b05b6c 100644
--- a/scheduling/InstructionScheduler.mli
+++ b/scheduling/InstructionScheduler.mli
@@ -31,7 +31,7 @@ type problem = {
reference_counting : ((Registers.reg, int * int) Hashtbl.t
* ((Registers.reg * bool) list array)) option;
- (** See RTLpathScheduleraux.reference_counting. *)
+ (** See BTLScheduleraux.reference_counting. *)
instruction_usages: int array array;
(** At index {i i} the vector of resources used by instruction number {i i}. It must be the same length as [resource_bounds] *)
diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml
new file mode 100644
index 00000000..89254301
--- /dev/null
+++ b/scheduling/PrintBTL.ml
@@ -0,0 +1,118 @@
+open Printf
+open Camlcoq
+open Datatypes
+open Maps
+open BTL
+open PrintAST
+open DebugPrint
+open BTLtypes
+
+(* Printing of BTL code *)
+
+let reg pp r = fprintf pp "x%d" (P.to_int r)
+
+let rec regs pp = function
+ | [] -> ()
+ | [ r ] -> reg pp r
+ | r1 :: rl -> fprintf pp "%a, %a" reg r1 regs rl
+
+let ros pp = function
+ | Coq_inl r -> reg pp r
+ | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s)
+
+let print_succ pp s = fprintf pp "\tsucc %d\n" (P.to_int s)
+
+let print_pref pp pref = fprintf pp "%s" pref
+
+let rec print_iblock pp is_rec pref ib =
+ match ib with
+ | Bnop None ->
+ print_pref pp pref;
+ fprintf pp "??: Bnop None\n"
+ | Bnop (Some iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bnop\n" iinfo.inumb
+ | Bop (op, args, res, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bop: %a = %a\n" iinfo.inumb reg res
+ (PrintOp.print_operation reg)
+ (op, args)
+ | Bload (trap, chunk, addr, args, dst, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bload: %a = %s[%a]%a\n" iinfo.inumb reg dst
+ (name_of_chunk chunk)
+ (PrintOp.print_addressing reg)
+ (addr, args) print_trapping_mode trap
+ | Bstore (chunk, addr, args, src, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bstore: %s[%a] = %a\n" iinfo.inumb (name_of_chunk chunk)
+ (PrintOp.print_addressing reg)
+ (addr, args) reg src
+ | BF (Bcall (sg, fn, args, res, s), iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bcall: %a = %a(%a)\n" iinfo.inumb reg res ros fn regs args;
+ print_succ pp s
+ | BF (Btailcall (sg, fn, args), iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Btailcall: %a(%a)\n" iinfo.inumb ros fn regs args
+ | BF (Bbuiltin (ef, args, res, s), iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bbuiltin: %a = %s(%a)\n" iinfo.inumb
+ (print_builtin_res reg) res (name_of_external ef)
+ (print_builtin_args reg) args;
+ print_succ pp s
+ | Bcond (cond, args, ib1, ib2, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bcond: (%a) (prediction: %s)\n" iinfo.inumb
+ (PrintOp.print_condition reg)
+ (cond, args)
+ (match iinfo.pcond with
+ | None -> "none"
+ | Some true -> "branch"
+ | Some false -> "fallthrough");
+ let pref' = pref ^ " " in
+ fprintf pp "%sifso = [\n" pref;
+ if is_rec then print_iblock pp is_rec pref' ib1 else fprintf pp "...\n";
+ fprintf pp "%s]\n" pref;
+ fprintf pp "%sifnot = [\n" pref;
+ if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n";
+ fprintf pp "%s]\n" pref
+ | BF (Bjumptable (arg, tbl), iinfo) ->
+ let tbl = Array.of_list tbl in
+ print_pref pp pref;
+ fprintf pp "%d: Bjumptable: (%a)\n" iinfo.inumb reg arg;
+ for i = 0 to Array.length tbl - 1 do
+ fprintf pp "\t\tcase %d: goto %d\n" i (P.to_int tbl.(i))
+ done
+ | BF (Breturn None, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Breturn\n" iinfo.inumb
+ | BF (Breturn (Some arg), iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Breturn: %a\n" iinfo.inumb reg arg
+ | BF (Bgoto s, iinfo) ->
+ print_pref pp pref;
+ fprintf pp "%d: Bgoto: %d\n" iinfo.inumb (P.to_int s)
+ | Bseq (ib1, ib2) ->
+ if is_rec then (
+ print_iblock pp is_rec pref ib1;
+ print_iblock pp is_rec pref ib2)
+ else fprintf pp "Bseq...\n"
+
+let print_btl_inst pp ib =
+ if !debug_flag then print_iblock pp false " " ib else ()
+
+let print_btl_code pp btl =
+ if !debug_flag then (
+ fprintf pp "\n";
+ List.iter
+ (fun (n, ibf) ->
+ fprintf pp "[BTL Liveness] ";
+ print_regset ibf.input_regs;
+ fprintf pp "\n";
+ fprintf pp "[BTL block %d with inumb %d]\n" (P.to_int n) ibf.binfo.bnumb;
+ print_iblock pp true " " ibf.entry;
+ fprintf pp "\n")
+ (PTree.elements btl);
+ fprintf pp "\n")
+ else ()
diff --git a/scheduling/RTLpath.v b/scheduling/RTLpath.v
deleted file mode 100644
index b29a7759..00000000
--- a/scheduling/RTLpath.v
+++ /dev/null
@@ -1,1066 +0,0 @@
-(** 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).
-
- 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.
-Require Import Lia.
-
-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 *)
-
-(** 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] 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].
-
- Remark: a path can loop several times in the CFG.
-
-*)
-
-Record path_info := {
- 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 *)
- pre_output_regs: Regset.t;
- (** 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.
-
-Definition path_entry (pm: path_map) (n: node): Prop := pm!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.(psize) 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_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 transf_program (p: program) : RTL.program := transform_program fundef_RTL p.
-Coercion transf_program: program >-> RTL.program.
-
-(** * Path-step semantics of RTLpath programs *)
-
-(* Semantics of internal instructions (mimicking RTL semantics) *)
-
-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 (mk_istate true pc' rs m)
- | Iop op args res pc' =>
- SOME v <- eval_operation ge sp op rs##args m IN
- 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 (mk_istate true pc' (rs#dst <- v) m)
- | Iload NOTRAP chunk addr args dst pc' =>
- let default_state := mk_istate true pc' rs#dst <- Vundef 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 (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 (mk_istate true pc' rs m')
- | Icond cond args ifso ifnot _ =>
- SOME b <- eval_condition cond rs##args m IN
- Some (mk_istate (negb b) (if b then ifso else ifnot) rs m)
- | _ => None (* TODO jumptable ? *)
- end.
-
-(** 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 (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 (icontinue st) then
- isteps ge p f sp (irs st) (imem st) (ipc st)
- 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 *)
- (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.
-
-(* 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 ->
- istep ge i sp rs m = Some st ->
- path_last_step ge pge stack f sp pc rs m
- 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 ->
- 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': (* 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').
-
-(* 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 ->
- (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 ->
- (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 *)
-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.(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) ->
- 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).
-
-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.
-
-Definition semantics (p: program) :=
- Semantics (step (Genv.globalenv (transf_program 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.
-
-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
- | 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
- | H: _ = None |- _ => (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 * |- *).
-
-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.
-
-(** ** 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 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 ->
- 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.
-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.(ipc) st.(irs) st.(imem)).
-Proof.
- induction path; simpl; try_simplify_someHyps.
- inversion_SOME i; intros Hi.
- inversion_SOME st0; intros Hst0.
- destruct (icontinue st0) eqn:cont.
- + intros; eapply star_step.
- - eapply istep_correct; eauto.
- - simpl; eauto.
- - auto.
- + intros; simplify_someHyp; eapply star_step.
- - eapply istep_correct; eauto.
- - simpl; eauto.
- - auto.
-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.(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 (icontinue st0) eqn:cont.
- + intros; eapply plus_left.
- - eapply istep_correct; eauto.
- - eapply isteps_correct; eauto.
- - auto.
- + intros X; inversion X; subst.
- eapply plus_one.
- eapply istep_correct; eauto.
-Qed.
-
-Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro: core.
-
-Section CORRECTNESS.
-
-Variable p: program.
-
-Lemma match_prog_RTL: match_program (fun _ f tf => tf = fundef_RTL f) eq p (transf_program p).
-Proof.
- eapply match_transform_program; eauto.
-Qed.
-
-Let pge := Genv.globalenv p.
-Let ge := Genv.globalenv (transf_program 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 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 ->
- RTL.step ge (State stack f sp pc rs m) t s.
-Proof.
- destruct 1; try (eapply istep_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 isteps_correct_early_exit; eauto.
- + eapply plus_right.
- eapply isteps_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: core.
-
-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.
- - 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.
-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, transf_program (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. *)
-
-
-(** 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 (transf_program p) p.
-Proof.
- unfold match_program, match_program_gen; intuition.
- unfold transf_program 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. enough (path'-path'=0)%nat as ->; [simpl;eauto|lia].
- + intros pc WF; enough (S m-path'=S (m-path'))%nat as ->; [simpl;eauto|lia].
- 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.(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.(psize)).
- exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps O); lia || eauto.
- destruct 1 as (pc' & NTH_SUCC & WF'); auto.
- assert (ps - 0 = ps)%nat as HH by lia. rewrite HH in NTH_SUCC. clear HH.
- 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.(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.(psize)).
- intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps (ps-path)); eauto. { lia. }
- destruct 1 as (pc' & NTH_SUCC & WF').
- assert (ps - (ps - path) = path)%nat as HH by lia. rewrite HH in NTH_SUCC. clear HH.
- unfold nth_default_succ_inst.
- inversion WF'; clear WF'; subst. { lia. }
- simplify_someHyps; eauto.
-Qed.
-
-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.
-Local Hint Resolve fn_entry_point_wf: core.
-Local Opaque path_entry.
-
-Lemma istep_successors ge i sp rs m st:
- istep ge i sp rs m = Some st ->
- In (ipc st) (successors_instr i).
-Proof.
- destruct i; simpl; try congruence; simplify_SOME x.
- all: explore_destruct; simplify_SOME x.
-Qed.
-
-Lemma istep_normal_exit ge i sp rs m st:
- istep ge i sp rs m = Some st ->
- 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.(icontinue) = true ->
- isteps ge path f sp rs m pc = Some st ->
- 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 (icontinue st0) eqn:X; try congruence.
- try_simplify_someHyps.
- intros; erewrite istep_normal_exit; eauto.
-Qed.
-
-
-(* TODO: the three following lemmas could maybe simplified by introducing an auxiliary
- left-recursive definition equivalent to isteps ?
-*)
-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.(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.
- 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 (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 ->
- (icontinue st)=false ->
- exists st0 i path0,
- (path > path0)%nat /\
- isteps ge path0 f sp rs m pc = Some st0 /\
- 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 (icontinue 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.
- lia.
-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 ->
- (icontinue st)=false ->
- 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.
- 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 (icontinue 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)
- . (* TODO add jumptable here ? *)
-
-Lemma istep_early_exit ge i sp rs m st :
- istep ge i sp rs m = Some st ->
- 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.
- all: explore_destruct; simplify_SOME b; try discriminate.
-Qed.
-
-Section COMPLETENESS.
-
-Variable p: program.
-
-Let pge := Genv.globalenv p.
-Let ge := Genv.globalenv (transf_program 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 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
- 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: core.
-
-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 => true
- | _ => false
- 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.(psize))%nat ->
- isteps ge (path.(psize)-idx) f sp rs m pc = Some s2 ->
- 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 :=
- 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
- /\ wf_stackframe (stack_of s2).
-
-(** *** Auxiliary lemmas of completeness *)
-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.(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.
- 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 ->
- (fn_path f)!pc = Some path ->
- (S idx <= path.(psize))%nat ->
- 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))); lia || eauto.
- intros (i & pc' & Hi & Hpc & DUM).
- unfold nth_default_succ_inst in Hi.
- 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; lia || eauto.
- set (ps:=path.(psize)). enough (ps - idx = S (ps - (S idx)))%nat as ->; try lia.
- erewrite <- isteps_step_right; eauto.
-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.(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.(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 *)
- 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 isteps_normal_exit in Hi; eauto.
- destruct (default_succ i) eqn:Hn0.
- + (* exec_istate *)
- exploit istep_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 (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 || lia.
- * 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 (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto.
- exists path0.(psize); intuition eauto.
- econstructor; eauto.
- * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || lia.
- * 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.(psize); intuition eauto.
- econstructor; eauto.
- * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || lia.
- * 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:
- 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.(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.(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.(icontinue)) eqn: CONT.
- destruct idx as [ | idx].
- + (* path_step on normal_exit *)
- assert (path.(psize)-0=path.(psize))%nat as HH by lia. rewrite HH in PSTEP. clear HH.
- exploit normal_exit; eauto.
- intros (s2' & LSTEP & (idx' & MATCH)).
- exists idx'; exists s2'; intuition eauto.
- + (* stuttering step *)
- 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 (isteps_resize ge (path.(psize) - idx)%nat path.(psize)); eauto; try lia.
- clear PSTEP; intros PSTEP.
- (* TODO for clarification: move the assert below into a separate lemma *)
- 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).
- 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); lia || eauto.
- intros (i' & pc' & Hi' & Hpc' & ENTRY).
- unfold nth_default_succ_inst in Hi'.
- 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 HPATH0 as (path1 & Hpath1).
- 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)).
- exists idx'. exists s2'. constructor; auto.
- right. right. eexists; intuition eauto.
- (* now, prove the last step *)
- rewrite Hpath1size; exploit exec_normal_exit. 4:{ eauto. }
- - simpl; eauto.
- - simpl; eauto.
- - simpl; eauto.
- * (* single step case *)
- 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 lia. simpl; eauto. }
- - lia.
- - simpl; eauto.
- - simpl; eauto.
- - intuition subst.
- repeat eexists; intuition eauto.
-Qed.
-
-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 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 (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)). enough (ps-ps=O)%nat as ->; simpl; eauto.
- lia.
- - 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 (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto.
- exists path.(psize). constructor; auto.
- econstructor; eauto.
- - set (ps:=path.(psize)). enough (ps-ps=O)%nat as ->; simpl; eauto.
- lia.
- - 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' ->
- 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: 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.
- 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.
-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 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.
- - Local Hint Resolve star_refl: core.
- 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.
diff --git a/scheduling/RTLpathCommon.ml b/scheduling/RTLpathCommon.ml
deleted file mode 100644
index 3d123ba8..00000000
--- a/scheduling/RTLpathCommon.ml
+++ /dev/null
@@ -1,14 +0,0 @@
-open Maps
-open Registers
-open Camlcoq
-
-type superblock = {
- mutable instructions: P.t array; (* pointers to code instructions *)
- (* each predicted Pcb has its attached liveins *)
- (* This is indexed by the pc value *)
- mutable liveins: Regset.t PTree.t;
- (* Union of the input_regs of the last successors *)
- s_output_regs: Regset.t;
- typing: RTLtyping.regenv
-}
-
diff --git a/scheduling/RTLpathLivegen.v b/scheduling/RTLpathLivegen.v
deleted file mode 100644
index 9f646ad0..00000000
--- a/scheduling/RTLpathLivegen.v
+++ /dev/null
@@ -1,325 +0,0 @@
-(** 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 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
- | 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
- end.
-
-
-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 ->
- 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.
-
-
-Lemma ipath_checker_default_succ (f: RTLpath.function) path: forall alive pc res,
- ipath_checker path f (fn_path f) alive pc = Some res
- -> nth_default_succ (fn_code f) path pc = Some (snd res).
-Proof.
- induction path; simpl.
- + try_simplify_someHyps.
- + intros alive pc res.
- inversion_SOME i; intros INST.
- inversion_SOME res0; intros ICHK IPCHK.
- rewrite INST.
- erewrite iinst_checker_default_succ; 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.
-
-(* 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 :=
- 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 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_tt_true in EXIT.
- firstorder (subst; eauto).
-Qed.
-
-Local Hint Resolve exit_list_checker_correct: core.
-
-Definition final_inst_checker (pm: path_map) (alive por: 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 por) 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_builtin_res res por) pc' tt
- | Ijumptable arg tbl =>
- ASSERT Regset.mem arg alive IN
- ASSERT exit_list_checker pm por tbl IN
- Some tt
- | Ireturn optarg =>
- ASSERT (reg_option_mem optarg) alive IN
- Some tt
- | _ => None
- end.
-
-Lemma final_inst_checker_wellformed (c:code) pc (pm: path_map) (alive por: Regset.t) (i: instruction):
- final_inst_checker pm alive por 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).
-Qed.
-
-Definition inst_checker (pm: path_map) (alive por: Regset.t) (i: instruction): option unit :=
- match iinst_checker pm alive i with
- | Some res =>
- ASSERT Regset.subset por (fst res) IN
- exit_checker pm por (snd res) tt
- | _ =>
- ASSERT Regset.subset por alive IN
- final_inst_checker pm alive por i
- end.
-
-Lemma inst_checker_wellformed (c:code) pc (pm: path_map) (alive por: Regset.t) (i: instruction):
- inst_checker pm alive por i = Some tt ->
- c!pc = Some i -> wellformed_path c pm 0 pc.
-Proof.
- unfold inst_checker.
- destruct (iinst_checker pm alive i) as [[alive0 pc0]|] eqn: CHECK1; simpl.
- - simpl; intros CHECK2 PC. eapply wf_last_node; eauto.
- destruct i; simpl in * |- *; intuition (subst; eauto);
- try (generalize CHECK2 CHECK1; clear CHECK1 CHECK2; try (inversion_SOME path); repeat inversion_ASSERT; try_simplify_someHyps).
- intros PC CHECK1 CHECK2.
- intros; exploit exit_checker_res; eauto.
- intros X; inversion X. intros; subst; eauto.
- - simpl; intros CHECK2 PC. eapply final_inst_checker_wellformed; eauto.
- generalize CHECK2. clear CHECK2. inversion_ASSERT. try_simplify_someHyps.
-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) (path.(pre_output_regs)) 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.
-
-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.
-
diff --git a/scheduling/RTLpathLivegenaux.ml b/scheduling/RTLpathLivegenaux.ml
deleted file mode 100644
index 2a20a15d..00000000
--- a/scheduling/RTLpathLivegenaux.ml
+++ /dev/null
@@ -1,290 +0,0 @@
-open RTL
-open RTLpath
-open Registers
-open Maps
-open Camlcoq
-open Datatypes
-open Kildall
-open Lattice
-open DebugPrint
-
-let get_some = function
-| None -> failwith "Got None instead of Some _"
-| Some thing -> thing
-
-let successors_inst = function
-| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n]
-| Icond (_,_,n1,n2,_) -> [n1; n2]
-| Ijumptable (_,l) -> l
-| Itailcall _ | Ireturn _ -> []
-
-let predicted_successor = function
-| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> Some n
-| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> None
-| Icond (_,_,n1,n2,p) -> (
- match p with
- | Some true -> Some n1
- | Some false -> Some n2
- | None -> None )
-| Ijumptable _ | Itailcall _ | Ireturn _ -> None
-
-let non_predicted_successors i =
- match predicted_successor i with
- | None -> successors_inst i
- | Some n -> List.filter (fun n' -> n != n') (successors_inst i)
-
-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
-
-(* 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
- let rec dig_path e =
- if (get_some @@ PTree.get e !visited) then
- ()
- else begin
- visited := PTree.set e true !visited;
- let psize = ref (-1) in
- let path_successors = ref [] in
- let rec dig_path_rec n : (path_info * node list) option =
- let inst = get_some @@ PTree.get n code in
- begin
- psize := !psize + 1;
- 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 = Regset.empty; pre_output_regs = Regset.empty; output_regs = Regset.empty },
- !path_successors @ successors_inst inst)
- end
- 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
- end
- in begin
- dig_path entry;
- !path_map
- end
-
-let transfer f pc after = let open Liveness in
- match PTree.get pc f.fn_code with
- | Some i ->
- (match i with
- | Inop _ -> after
- | Iop (_, args, res, _) ->
- reg_list_live args (Regset.remove res after)
- | Iload (_, _, _, args, dst, _) ->
- reg_list_live args (Regset.remove dst after)
- | Istore (_, _, args, src, _) ->
- reg_list_live args (Regset.add src after)
- | Icall (_, ros, args, res, _) ->
- reg_list_live args (reg_sum_live ros (Regset.remove res after))
- | Itailcall (_, ros, args) ->
- reg_list_live args (reg_sum_live ros Regset.empty)
- | Ibuiltin (_, args, res, _) ->
- reg_list_live (AST.params_of_builtin_args args)
- (reg_list_dead (AST.params_of_builtin_res res) after)
- | Icond (_, args, _, _, _) ->
- reg_list_live args after
- | Ijumptable (arg, _) ->
- Regset.add arg after
- | Ireturn optarg ->
- reg_option_live optarg Regset.empty)
- | None -> Regset.empty
-
-module RegsetLat = LFSet(Regset)
-
-module DS = Backward_Dataflow_Solver(RegsetLat)(NodeSetBackward)
-
-let analyze f =
- let liveouts = get_some @@ DS.fixpoint f.fn_code successors_instr (transfer f) in
- PTree.map (fun n _ -> let lo = PMap.get n liveouts in transfer f n lo) f.fn_code
-
-(** OLD CODE - If needed to have our own kildall
-
-let transfer after = let open Liveness in function
- | 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
-
-let analyze 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 rec traverse code n size =
- let inst = get_some @@ PTree.get n code in
- if (size == 0) then (inst, n)
- else
- let n' = get_some @@ predicted_successor inst in
- traverse code n' (size-1)
-
-let get_outputs liveness f n pi =
- let (last_instruction, pc_last) = traverse f.fn_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
- let outputs = List.fold_left Regset.union Regset.empty list_input_regs in
- let por = match last_instruction with (* see RTLpathLivegen.final_inst_checker *)
- | Icall (_, _, _, res, _) -> Regset.remove res outputs
- | Ibuiltin (_, _, res, _) -> Liveness.reg_list_dead (AST.params_of_builtin_res res) outputs
- | Itailcall (_, _, _) | Ireturn _ ->
- assert (outputs = Regset.empty); (* defensive check for performance *)
- outputs
- | _ -> outputs
- in (por, outputs)
-
-let set_pathmap_liveness f pm =
- let liveness = analyze f in
- let new_pm = ref PTree.empty in
- begin
- debug "Liveness: "; print_ptree_regset liveness; debug "\n";
- List.iter (fun (n, pi) ->
- let inputs = get_some @@ PTree.get n liveness in
- let (por, outputs) = get_outputs liveness f n pi in
- new_pm := PTree.set n
- {psize=pi.psize; input_regs=inputs; pre_output_regs=por; output_regs=outputs} !new_pm
- ) (PTree.elements pm);
- !new_pm
- end
-
-let print_path_info pi = begin
- debug "(psize=%d; " (Camlcoq.Nat.to_int pi.psize);
- debug "\ninput_regs=";
- print_regset pi.input_regs;
- debug "\n; pre_output_regs=";
- print_regset pi.pre_output_regs;
- debug "\n; output_regs=";
- print_regset pi.output_regs;
- debug ")\n"
-end
-
-let print_path_map path_map = begin
- debug "[";
- List.iter (fun (n,pi) ->
- debug "\n\t";
- debug "%d: " (P.to_int n);
- print_path_info pi
- ) (PTree.elements path_map);
- debug "]"
-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 = set_pathmap_liveness f @@ get_path_map code entry join_points in
- begin
- debug "Join points: ";
- print_true_nodes join_points;
- debug "\nPath map: ";
- print_path_map path_map;
- debug "\n";
- path_map
- end
diff --git a/scheduling/RTLpathLivegenproof.v b/scheduling/RTLpathLivegenproof.v
deleted file mode 100644
index b02400bf..00000000
--- a/scheduling/RTLpathLivegenproof.v
+++ /dev/null
@@ -1,760 +0,0 @@
-(** Proofs of the liveness properties from the liveness checker of RTLpathLivengen.
-*)
-
-
-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 Linking Values Events.
-Require Import Program.
-
-Definition match_prog (p: RTL.program) (tp: program) :=
- match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
-
-Lemma transf_program_match:
- forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
-Proof.
- intros. eapply match_transform_partial_program_contextual; eauto.
-Qed.
-
-Section PRESERVATION.
-
-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 (RTLpath.transf_program tprog).
-
-Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof.
- rewrite <- (Genv.find_symbol_match TRANSL).
- apply (Genv.find_symbol_match (match_prog_RTL tprog)).
-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 ge tge.
-Proof.
- eapply senv_transitivity. { eapply (Genv.senv_match TRANSL). }
- eapply RTLpath.senv_preserved.
-Qed.
-
-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.
- intros; apply (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
-Qed.
-
-
-Lemma function_ptr_RTL_preserved v f: Genv.find_funct_ptr ge v = Some f -> Genv.find_funct_ptr tge v = Some f.
-Proof.
- 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 find_function_preserved ros rs fd:
- RTL.find_function ge ros rs = Some fd -> RTL.find_function tge ros rs = Some fd.
-Proof.
- 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: core.
-
-Lemma transf_program_RTL_correct:
- forward_simulation (RTL.semantics prog) (RTL.semantics (RTLpath.transf_program tprog)).
-Proof.
- 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 (RTLpath.transf_program 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.
- all: erewrite eval_addressing_preserved; eauto.
- + (* Iload notrap1 *) eapply exec_Iload_notrap1; eauto.
- all: erewrite eval_addressing_preserved; eauto.
- + (* Iload notrap2 *) eapply exec_Iload_notrap2; eauto.
- all: erewrite eval_addressing_preserved; eauto.
- + (* Istore *) eapply exec_Istore; eauto.
- all: 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.
- eapply compose_forward_simulations.
- + eapply transf_program_RTL_correct.
- + eapply RTLpath_complete.
-Qed.
-
-
-(* Properties used in hypothesis of [RTLpathLiveproofs.step_eqlive] theorem *)
-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.
- 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.
-
-Local Open Scope lazy_bool_scope.
-Local Open Scope option_monad_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.
-
-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_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_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: core.
-
-Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true.
-Proof.
- destruct b1; simpl; intuition.
-Qed.
-
-Lemma list_mem_correct (rl: list reg) (alive: Regset.t):
- list_mem rl alive = true -> forall r, List.In r rl -> ext alive r.
-Proof.
- induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto.
-Qed.
-
-Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args.
-Proof.
- induction args; simpl; auto.
- intros EQLIVE ALIVE; rewrite IHargs; auto.
- unfold eqlive_reg in EQLIVE.
- rewrite EQLIVE; auto.
-Qed.
-
-Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args.
-Proof.
- intros; eapply eqlive_reg_list; eauto.
- intros; eapply list_mem_correct; eauto.
-Qed.
-
-Record eqlive_istate alive (st1 st2: istate): Prop :=
- { 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 ->
- iinst_checker pm alive i = Some res ->
- istep ge i sp rs1 m = Some st1 ->
- 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.
- - (* Inop *)
- repeat (econstructor; eauto).
- - (* Iop *)
- inversion_ASSERT; try_simplify_someHyps.
- inversion_SOME v. intros EVAL.
- erewrite <- eqlive_reg_listmem; eauto.
- 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.
- destruct t.
- inversion_SOME a0. intros EVAL.
- erewrite <- eqlive_reg_listmem; eauto.
- try_simplify_someHyps.
- inversion_SOME v; try_simplify_someHyps.
- repeat (econstructor; simpl; eauto).
- 2:
- erewrite <- (eqlive_reg_listmem _ _ rs1 rs2); eauto;
- destruct (eval_addressing _ _ _ _);
- try destruct (Memory.Mem.loadv _ _ _);
- try (intros; inv H1; repeat (econstructor; simpl; eauto)).
- all:
- 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.
- 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.
- 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 ->
- icontinue st = true ->
- (snd res)=(ipc st).
-Proof.
- intros; exploit iinst_checker_default_succ; eauto.
- erewrite istep_normal_exit; eauto.
- congruence.
-Qed.
-
-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 ->
- 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 ->
- 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).
- 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.
- 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 ->
- 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.
- - repeat (econstructor; simpl; eauto).
- - inversion_SOME i; try_simplify_someHyps.
- inversion_SOME res0.
- inversion_SOME st0.
- intros.
- exploit iinst_checker_eqlive; eauto.
- destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]).
- try_simplify_someHyps.
- rewrite <- CONT, <- MEM, <- PC.
- destruct (icontinue st0) eqn:CONT'.
- * 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 ->
- 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 (icontinue 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 ->
- 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 (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.
- 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_stackframes_intro path res f sp pc rs1 rs2
- (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)):
- 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)
- (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
- (LIVE: liveness_ok_fundef f)
- (STACKS: list_forall2 eqlive_stackframes st1 st2):
- eqlive_states (Callstate st1 f args m) (Callstate st2 f args m)
- | eqlive_states_return st1 st2 v m
- (STACKS: list_forall2 eqlive_stackframes st1 st2):
- eqlive_states (Returnstate st1 v m) (Returnstate st2 v m).
-
-
-Section LivenessProperties.
-
-Variable prog: program.
-
-Let pge := Genv.globalenv prog.
-Let ge := Genv.globalenv (RTLpath.transf_program prog).
-
-Hypothesis all_fundef_liveness_ok: forall b f,
- Genv.find_funct_ptr pge b = Some f ->
- liveness_ok_fundef 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.
- - intros; eapply find_funct_liveness_ok; eauto.
- - destruct (Genv.find_symbol pge i); try congruence.
- eapply all_fundef_liveness_ok; 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 final_inst_checker_from_iinst_checker i sp rs m st pm alive por:
- istep ge i sp rs m = Some st ->
- final_inst_checker pm alive por i = None.
-Proof.
- destruct i; simpl; try congruence.
-Qed.
-
-(* is it useful ?
-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.
- unfold inst_checker.
- destruct (iinst_checker pm alive i); simpl; auto.
- 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.
-
-Local Hint Resolve in_or_app: local.
-Lemma eqlive_eval_builtin_args alive rs1 rs2 sp m args vargs:
- eqlive_reg alive rs1 rs2 ->
- Events.eval_builtin_args ge (fun r => rs1 # r) sp m args vargs ->
- (forall r, List.In r (params_of_builtin_args args) -> alive r) ->
- Events.eval_builtin_args ge (fun r => rs2 # r) sp m args vargs.
-Proof.
- unfold Events.eval_builtin_args.
- intros EQLIVE; induction 1 as [|a1 al b1 bl EVAL1 EVALL]; simpl.
- { econstructor; eauto. }
- intro X.
- assert (X1: eqlive_reg (fun r => In r (params_of_builtin_arg a1)) rs1 rs2).
- { eapply eqlive_reg_monotonic; eauto with local. }
- lapply IHEVALL; eauto with local.
- clear X IHEVALL; intro X. econstructor; eauto.
- generalize X1; clear EVALL X1 X.
- induction EVAL1; simpl; try (econstructor; eauto; fail).
- - intros X1; erewrite X1; [ econstructor; eauto | eauto ].
- - intros; econstructor.
- + eapply IHEVAL1_1; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- + eapply IHEVAL1_2; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- - intros; econstructor.
- + eapply IHEVAL1_1; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- + eapply IHEVAL1_2; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
-Qed.
-
-Lemma exit_checker_eqlive_builtin_res (pm: path_map) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg):
- 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)).
-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 ->
- 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 final_inst_checker_eqlive (f: function) sp alive por pc i rs1 rs2 m stk1 stk2 t s1:
- list_forall2 eqlive_stackframes stk1 stk2 ->
- eqlive_reg (ext alive) rs1 rs2 ->
- Regset.Subset por alive ->
- liveness_ok_function f ->
- (fn_code f) ! pc = Some i ->
- path_last_step ge pge stk1 f sp pc rs1 m t s1 ->
- final_inst_checker (fn_path f) alive por 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 SUB 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'|
- sp pc rs1 m arg tbl n pc' |
- st1 pc rs1 m optr m'];
- try_simplify_someHyps.
- + (* istate *)
- intros PC ISTEP. erewrite final_inst_checker_from_iinst_checker; eauto.
- congruence.
- + (* Icall *)
- repeat inversion_ASSERT. intros.
- exploit exit_checker_eqlive_ext1; eauto.
- eapply eqlive_reg_monotonic; 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_ok; eauto.
- repeat (econstructor; eauto).
- + (* Itailcall *)
- 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_ok; eauto.
- + (* Ibuiltin *)
- repeat inversion_ASSERT. intros.
- exploit exit_checker_eqlive_builtin_res; eauto.
- eapply eqlive_reg_monotonic; 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 *)
- repeat inversion_ASSERT. intros.
- exploit exit_list_checker_eqlive; eauto.
- eapply eqlive_reg_monotonic; 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.
- - eapply exec_Ireturn; eauto.
- - destruct optr; simpl in * |- *.
- * erewrite (EQLIVE r); eauto.
- eapply eqlive_states_return; eauto.
- * eapply eqlive_states_return; eauto.
-Qed.
-
-Lemma inst_checker_eqlive (f: function) sp alive por 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 por i = Some tt ->
- exists s2, path_last_step ge pge stk2 f sp pc rs2 m t s2 /\ eqlive_states s1 s2.
-Proof.
- unfold inst_checker;
- intros STACKS EQLIVE LIVENESS PC.
- destruct (iinst_checker (fn_path f) alive i) as [res|] eqn: IICHECKER.
- + destruct 1 as [i' sp pc rs1 m st1| | | | | ];
- try_simplify_someHyps.
- intros IICHECKER PC ISTEP. inversion_ASSERT.
- intros.
- destruct (icontinue 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.
- apply Regset.subset_2 in H.
- exploit exit_checker_eqlive; eauto.
- eapply eqlive_reg_monotonic; 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.
- + inversion_ASSERT.
- intros; exploit final_inst_checker_eqlive; eauto.
-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 ->
- 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) 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.
- exploit LIVE; eauto.
- 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|].
- - (* 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' ->
- 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.
- 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.
-Qed.
-
-End LivenessProperties.
diff --git a/scheduling/RTLpathSE_impl.v b/scheduling/RTLpathSE_impl.v
deleted file mode 100644
index e21d7cd1..00000000
--- a/scheduling/RTLpathSE_impl.v
+++ /dev/null
@@ -1,1664 +0,0 @@
-(** Implementation and refinement 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.
-Require Import RTLpathSE_theory RTLpathLivegenproof.
-Require Import Axioms RTLpathSE_simu_specs.
-Require Import RTLpathSE_simplify.
-
-Local Open Scope error_monad_scope.
-Local Open Scope option_monad_scope.
-
-Require Import Impure.ImpHCons.
-Import Notations.
-Import HConsing.
-
-Local Open Scope impure.
-Local Open Scope hse.
-
-Import ListNotations.
-Local Open Scope list_scope.
-
-Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := RET tt. (* TO REMOVE DEBUG INFO *)
-(*Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := DO s <~ k x;; println ("DEBUG simu_check:" +; s). (* TO INSERT DEBUG INFO *)*)
-
-Definition DEBUG (s: pstring): ?? unit := XDEBUG tt (fun _ => RET s).
-
-(** * Implementation of Data-structure use in Hash-consing *)
-
-Definition hsval_get_hid (hsv: hsval): hashcode :=
- match hsv with
- | HSinput _ hid => hid
- | HSop _ _ hid => hid
- | HSload _ _ _ _ _ hid => hid
- end.
-
-Definition list_hsval_get_hid (lhsv: list_hsval): hashcode :=
- match lhsv 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 lhsv _ => HSop o lhsv hid
- | HSload hsm trap chunk addr lhsv _ => HSload hsm trap chunk addr lhsv hid
- end.
-
-Definition list_hsval_set_hid (lhsv: list_hsval) (hid: hashcode): list_hsval :=
- match lhsv with
- | HSnil _ => HSnil hid
- | HScons hsv lhsv _ => HScons hsv lhsv hid
- end.
-
-Definition hsmem_set_hid (hsm: hsmem) (hid: hashcode): hsmem :=
- match hsm with
- | HSinit _ => HSinit hid
- | HSstore hsm chunk addr lhsv srce _ => HSstore hsm chunk addr lhsv srce hid
- end.
-
-
-Lemma hsval_set_hid_correct x y ge sp rs0 m0:
- hsval_set_hid x unknown_hid = hsval_set_hid y unknown_hid ->
- seval_hsval ge sp x rs0 m0 = seval_hsval ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve hsval_set_hid_correct: core.
-
-Lemma list_hsval_set_hid_correct x y ge sp rs0 m0:
- list_hsval_set_hid x unknown_hid = list_hsval_set_hid y unknown_hid ->
- seval_list_hsval ge sp x rs0 m0 = seval_list_hsval ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve list_hsval_set_hid_correct: core.
-
-Lemma hsmem_set_hid_correct x y ge sp rs0 m0:
- hsmem_set_hid x unknown_hid = hsmem_set_hid y unknown_hid ->
- seval_hsmem ge sp x rs0 m0 = seval_hsmem ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve hsmem_set_hid_correct: core.
-
-(** 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 _, HSop op2 lsv2 _ =>
- DO b1 <~ phys_eq lsv1 lsv2;;
- if b1
- then struct_eq op1 op2 (* NB: really need a struct_eq here ? *)
- else RET false
- | 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.
-
-
-Lemma and_true_split a b: a && b = true <-> a = true /\ b = true.
-Proof.
- destruct a; simpl; intuition.
-Qed.
-
-Lemma hsval_hash_eq_correct x y:
- WHEN hsval_hash_eq x y ~> b THEN
- b = true -> hsval_set_hid x unknown_hid = hsval_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque hsval_hash_eq.
-Local Hint Resolve hsval_hash_eq_correct: wlp.
-
-Definition list_hsval_hash_eq (lsv1 lsv2: list_hsval): ?? 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 list_hsval_hash_eq_correct x y:
- WHEN list_hsval_hash_eq x y ~> b THEN
- b = true -> list_hsval_set_hid x unknown_hid = list_hsval_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque list_hsval_hash_eq.
-Local Hint Resolve list_hsval_hash_eq_correct: wlp.
-
-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 hsmem_hash_eq_correct x y:
- WHEN hsmem_hash_eq x y ~> b THEN
- b = true -> hsmem_set_hid x unknown_hid = hsmem_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque hsmem_hash_eq.
-Local Hint Resolve hsmem_hash_eq_correct: wlp.
-
-
-Definition hSVAL: hashP hsval := {| hash_eq := hsval_hash_eq; get_hid:=hsval_get_hid; set_hid:=hsval_set_hid |}.
-Definition hLSVAL: hashP list_hsval := {| hash_eq := list_hsval_hash_eq; get_hid:= list_hsval_get_hid; set_hid:= list_hsval_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 hv =>
- DO hv_name <~ string_of_hashcode (hsval_get_hid hv);;
- println ("unexpected undef behavior of hashcode:" +; (CamlStr hv_name)) |}.
-Obligation 1.
- wlp_simplify.
-Qed.
-
-(** ** various auxiliary (trivial lemmas) *)
-Lemma hsilocal_refines_sreg ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> 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.
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_sreg: core.
-
-Lemma hsilocal_refines_valid_pointer ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> forall m b ofs, seval_smem ge sp st.(si_smem) rs0 m0 = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs.
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_valid_pointer: core.
-
-Lemma hsilocal_refines_smem_refines ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst -> smem_refines ge sp rs0 m0 (hsi_smem hst) (st.(si_smem)).
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_smem_refines: core.
-
-Lemma hsistate_refines_dyn_exits ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst st -> hsiexits_refines_dyn ge sp rs0 m0 (hsi_exits hst) (si_exits st).
-Proof.
- unfold hsistate_refines_dyn; intuition.
-Qed.
-Local Hint Resolve hsistate_refines_dyn_exits: core.
-
-Lemma hsistate_refines_dyn_local ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst 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: core.
-
-Lemma hsistate_refines_dyn_nested ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst st -> nested_sok ge sp rs0 m0 (si_local st) (si_exits st).
-Proof.
- unfold hsistate_refines_dyn; intuition.
-Qed.
-Local Hint Resolve hsistate_refines_dyn_nested: core.
-
-(** * Implementation of symbolic execution *)
-Section CanonBuilding.
-
-Variable hC_hsval: hashinfo hsval -> ?? hsval.
-
-Hypothesis hC_hsval_correct: forall hs,
- WHEN hC_hsval hs ~> hs' THEN forall ge sp rs0 m0,
- seval_hsval ge sp (hdata hs) rs0 m0 = seval_hsval ge sp hs' rs0 m0.
-
-Variable hC_list_hsval: hashinfo list_hsval -> ?? list_hsval.
-Hypothesis hC_list_hsval_correct: forall lh,
- WHEN hC_list_hsval lh ~> lh' THEN forall ge sp rs0 m0,
- seval_list_hsval ge sp (hdata lh) rs0 m0 = seval_list_hsval ge sp lh' rs0 m0.
-
-Variable hC_hsmem: hashinfo hsmem -> ?? hsmem.
-Hypothesis hC_hsmem_correct: forall hm,
- WHEN hC_hsmem hm ~> hm' THEN forall ge sp rs0 m0,
- seval_hsmem ge sp (hdata hm) rs0 m0 = seval_hsmem ge sp hm' rs0 m0.
-
-(* First, we wrap constructors for hashed values !*)
-
-Definition reg_hcode := 1.
-Definition op_hcode := 2.
-Definition load_hcode := 3.
-
-Definition hSinput_hcodes (r: reg) :=
- DO hc <~ hash reg_hcode;;
- DO hv <~ hash r;;
- RET [hc;hv].
-Extraction Inline hSinput_hcodes.
-
-Definition hSinput (r:reg): ?? hsval :=
- DO hv <~ hSinput_hcodes r;;
- hC_hsval {| hdata:=HSinput r unknown_hid; hcodes :=hv; |}.
-
-Lemma hSinput_correct r:
- WHEN hSinput r ~> hv THEN forall ge sp rs0 m0,
- sval_refines ge sp rs0 m0 hv (Sinput r).
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSinput.
-Local Hint Resolve hSinput_correct: wlp.
-
-Definition hSop_hcodes (op:operation) (lhsv: list_hsval) :=
- DO hc <~ hash op_hcode;;
- DO hv <~ hash op;;
- RET [hc;hv;list_hsval_get_hid lhsv].
-Extraction Inline hSop_hcodes.
-
-Definition hSop (op:operation) (lhsv: list_hsval): ?? hsval :=
- DO hv <~ hSop_hcodes op lhsv;;
- hC_hsval {| hdata:=HSop op lhsv unknown_hid; hcodes :=hv |}.
-
-Lemma hSop_fSop_correct op lhsv:
- WHEN hSop op lhsv ~> hv THEN forall ge sp rs0 m0,
- seval_hsval ge sp hv rs0 m0 = seval_hsval ge sp (fSop op lhsv) rs0 m0.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSop.
-Local Hint Resolve hSop_fSop_correct: wlp_raw.
-
-Lemma hSop_correct op lhsv:
- WHEN hSop op lhsv ~> hv THEN forall ge sp rs0 m0 lsv sm m
- (MEM: seval_smem ge sp sm rs0 m0 = Some m)
- (MVALID: forall b ofs, Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv),
- sval_refines ge sp rs0 m0 hv (Sop op lsv sm).
-Proof.
- generalize fSop_correct; simpl.
- intros X.
- wlp_xsimplify ltac:(intuition eauto with wlp wlp_raw).
- erewrite H, X; eauto.
-Qed.
-Local Hint Resolve hSop_correct: wlp.
-
-Definition hSload_hcodes (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval):=
- DO hc <~ hash load_hcode;;
- DO hv1 <~ hash trap;;
- DO hv2 <~ hash chunk;;
- DO hv3 <~ hash addr;;
- RET [hc; hsmem_get_hid hsm; hv1; hv2; hv3; list_hsval_get_hid lhsv].
-Extraction Inline hSload_hcodes.
-
-Definition hSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval): ?? hsval :=
- DO hv <~ hSload_hcodes hsm trap chunk addr lhsv;;
- hC_hsval {| hdata := HSload hsm trap chunk addr lhsv unknown_hid; hcodes := hv |}.
-
-Lemma hSload_correct hsm trap chunk addr lhsv:
- WHEN hSload hsm trap chunk addr lhsv ~> hv THEN forall ge sp rs0 m0 lsv sm
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv)
- (MR: smem_refines ge sp rs0 m0 hsm sm),
- sval_refines ge sp rs0 m0 hv (Sload sm trap chunk addr lsv).
-Proof.
- wlp_simplify.
- rewrite <- LR, <- MR.
- auto.
-Qed.
-Global Opaque hSload.
-Local Hint Resolve hSload_correct: wlp.
-
-Definition hSnil (_: unit): ?? list_hsval :=
- hC_list_hsval {| hdata := HSnil unknown_hid; hcodes := nil |}.
-
-Lemma hSnil_correct:
- WHEN hSnil() ~> hv THEN forall ge sp rs0 m0,
- list_sval_refines ge sp rs0 m0 hv Snil.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSnil.
-Local Hint Resolve hSnil_correct: wlp.
-
-Definition hScons (hsv: hsval) (lhsv: list_hsval): ?? list_hsval :=
- hC_list_hsval {| hdata := HScons hsv lhsv unknown_hid; hcodes := [hsval_get_hid hsv; list_hsval_get_hid lhsv] |}.
-
-Lemma hScons_correct hsv lhsv:
- WHEN hScons hsv lhsv ~> lhsv' THEN forall ge sp rs0 m0 sv lsv
- (VR: sval_refines ge sp rs0 m0 hsv sv)
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv),
- list_sval_refines ge sp rs0 m0 lhsv' (Scons sv lsv).
-Proof.
- wlp_simplify.
- rewrite <- VR, <- LR.
- auto.
-Qed.
-Global Opaque hScons.
-Local Hint Resolve hScons_correct: wlp.
-
-Definition hSinit (_: unit): ?? hsmem :=
- hC_hsmem {| hdata := HSinit unknown_hid; hcodes := nil |}.
-
-Lemma hSinit_correct:
- WHEN hSinit() ~> hm THEN forall ge sp rs0 m0,
- smem_refines ge sp rs0 m0 hm Sinit.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSinit.
-Local Hint Resolve hSinit_correct: wlp.
-
-Definition hSstore_hcodes (hsm: hsmem) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (srce: hsval):=
- DO hv1 <~ hash chunk;;
- DO hv2 <~ hash addr;;
- RET [hsmem_get_hid hsm; hv1; hv2; list_hsval_get_hid lhsv; hsval_get_hid srce].
-Extraction Inline hSstore_hcodes.
-
-Definition hSstore (hsm: hsmem) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (srce: hsval): ?? hsmem :=
- DO hv <~ hSstore_hcodes hsm chunk addr lhsv srce;;
- hC_hsmem {| hdata := HSstore hsm chunk addr lhsv srce unknown_hid; hcodes := hv |}.
-
-Lemma hSstore_correct hsm chunk addr lhsv hsv:
- WHEN hSstore hsm chunk addr lhsv hsv ~> hsm' THEN forall ge sp rs0 m0 lsv sm sv
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv)
- (MR: smem_refines ge sp rs0 m0 hsm sm)
- (VR: sval_refines ge sp rs0 m0 hsv sv),
- smem_refines ge sp rs0 m0 hsm' (Sstore sm chunk addr lsv sv).
-Proof.
- wlp_simplify.
- rewrite <- LR, <- MR, <- VR.
- auto.
-Qed.
-Global Opaque hSstore.
-Local Hint Resolve hSstore_correct: wlp.
-
-Definition hsi_sreg_get (hst: PTree.t hsval) r: ?? hsval :=
- match PTree.get r hst with
- | None => hSinput r
- | Some sv => RET sv
- end.
-
-Lemma hsi_sreg_get_correct hst r:
- WHEN hsi_sreg_get hst r ~> hsv THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- sval_refines ge sp rs0 m0 hsv (f r).
-Proof.
- unfold hsi_sreg_eval, hsi_sreg_proj; wlp_simplify; rewrite <- RR; try_simplify_someHyps.
-Qed.
-Global Opaque hsi_sreg_get.
-Local Hint Resolve hsi_sreg_get_correct: wlp.
-
-Fixpoint hlist_args (hst: PTree.t hsval) (l: list reg): ?? list_hsval :=
- match l with
- | nil => hSnil()
- | r::l =>
- DO v <~ hsi_sreg_get hst r;;
- DO lhsv <~ hlist_args hst l;;
- hScons v lhsv
- end.
-
-Lemma hlist_args_correct hst l:
- WHEN hlist_args hst l ~> lhsv THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- list_sval_refines ge sp rs0 m0 lhsv (list_sval_inj (List.map f l)).
-Proof.
- induction l; wlp_simplify.
-Qed.
-Global Opaque hlist_args.
-Local Hint Resolve hlist_args_correct: wlp.
-
-(** Convert a "fake" hash-consed term into a "real" hash-consed term *)
-
-Fixpoint fsval_proj hsv: ?? hsval :=
- match hsv with
- | HSinput r hc =>
- DO b <~ phys_eq hc unknown_hid;;
- if b
- then hSinput r (* was not yet really hash-consed *)
- else RET hsv (* already hash-consed *)
- | HSop op hl hc =>
- DO b <~ phys_eq hc unknown_hid;;
- if b
- then (* was not yet really hash-consed *)
- DO hl' <~ fsval_list_proj hl;;
- hSop op hl'
- else RET hsv (* already hash-consed *)
- | HSload hm t chk addr hl _ => RET hsv (* FIXME ? *)
- end
-with fsval_list_proj hsl: ?? list_hsval :=
- match hsl with
- | HSnil hc =>
- DO b <~ phys_eq hc unknown_hid;;
- if b
- then hSnil() (* was not yet really hash-consed *)
- else RET hsl (* already hash-consed *)
- | HScons hv hl hc =>
- DO b <~ phys_eq hc unknown_hid;;
- if b
- then (* was not yet really hash-consed *)
- DO hv' <~ fsval_proj hv;;
- DO hl' <~ fsval_list_proj hl;;
- hScons hv' hl'
- else RET hsl (* already hash-consed *)
- end.
-
-Lemma fsval_proj_correct hsv:
- WHEN fsval_proj hsv ~> hsv' THEN forall ge sp rs0 m0,
- seval_hsval ge sp hsv rs0 m0 = seval_hsval ge sp hsv' rs0 m0.
-Proof.
- induction hsv using hsval_mut
- with (P0 := fun lhsv =>
- WHEN fsval_list_proj lhsv ~> lhsv' THEN forall ge sp rs0 m0,
- seval_list_hsval ge sp lhsv rs0 m0 = seval_list_hsval ge sp lhsv' rs0 m0)
- (P1 := fun sm => True); try (wlp_simplify; tauto).
- - wlp_xsimplify ltac:(intuition eauto with wlp_raw wlp).
- rewrite H, H0; auto.
- - wlp_simplify; erewrite H0, H1; eauto.
-Qed.
-Global Opaque fsval_proj.
-Local Hint Resolve fsval_proj_correct: wlp.
-
-Lemma fsval_list_proj_correct lhsv:
- WHEN fsval_list_proj lhsv ~> lhsv' THEN forall ge sp rs0 m0,
- seval_list_hsval ge sp lhsv rs0 m0 = seval_list_hsval ge sp lhsv' rs0 m0.
-Proof.
- induction lhsv; wlp_simplify.
- erewrite H0, H1; eauto.
-Qed.
-Global Opaque fsval_list_proj.
-Local Hint Resolve fsval_list_proj_correct: wlp.
-
-
-(** ** Assignment of memory *)
-Definition hslocal_set_smem (hst:hsistate_local) hm :=
- {| hsi_smem := hm;
- hsi_ok_lsval := hsi_ok_lsval hst;
- hsi_sreg:= hsi_sreg hst
- |}.
-
-Lemma sok_local_set_mem ge sp rs0 m0 st sm:
- sok_local ge sp rs0 m0 (slocal_set_smem st sm)
- <-> (sok_local ge sp rs0 m0 st /\ seval_smem ge sp sm rs0 m0 <> None).
-Proof.
- unfold slocal_set_smem, sok_local; simpl; intuition (subst; eauto).
-Qed.
-
-Lemma hsok_local_set_mem ge sp rs0 m0 hst hsm:
- (seval_hsmem ge sp (hsi_smem hst) rs0 m0 = None -> seval_hsmem ge sp hsm rs0 m0 = None) ->
- hsok_local ge sp rs0 m0 (hslocal_set_smem hst hsm)
- <-> (hsok_local ge sp rs0 m0 hst /\ seval_hsmem ge sp hsm rs0 m0 <> None).
-Proof.
- unfold hslocal_set_smem, hsok_local; simpl; intuition.
-Qed.
-
-Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st hsm sm:
- (seval_hsmem ge sp (hsi_smem hst) rs0 m0 = None -> seval_hsmem ge sp hsm rs0 m0 = None) ->
- (forall m b ofs, seval_smem ge sp sm rs0 m0 = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs) ->
- hsilocal_refines ge sp rs0 m0 hst st ->
- (hsok_local ge sp rs0 m0 hst -> smem_refines ge sp rs0 m0 hsm sm) ->
- hsilocal_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm).
-Proof.
- intros PRESERV SMVALID (OKEQ & SMEMEQ' & REGEQ & MVALID) SMEMEQ.
- split; rewrite! hsok_local_set_mem; simpl; eauto; try tauto.
- rewrite sok_local_set_mem.
- intuition congruence.
-Qed.
-
-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 (hslocal_set_smem hst hm).
-
-Lemma hslocal_store_correct hst chunk addr args src:
- WHEN hslocal_store hst chunk addr args src ~> hst' THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st),
- hsilocal_refines ge sp rs0 m0 hst' (slocal_store st chunk addr args src).
-Proof.
- wlp_simplify.
- eapply hslocal_set_mem_correct; simpl; eauto.
- + intros X; erewrite H1; eauto.
- rewrite X. simplify_SOME z.
- + unfold hsilocal_refines in *;
- simplify_SOME z; intuition.
- erewrite <- Mem.storev_preserv_valid; [| eauto].
- eauto.
- + unfold hsilocal_refines in *; intuition eauto.
-Qed.
-Global Opaque hslocal_store.
-Local Hint Resolve hslocal_store_correct: wlp.
-
-(** ** 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_stat hst st pc hnxt nxt:
- hsistate_refines_stat hst st ->
- hsistate_refines_stat (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt).
-Proof.
- unfold hsistate_refines_stat; simpl; intuition.
-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 ->
- hsilocal_refines ge sp rs0 m0 hnxt nxt ->
- (sok_local ge sp rs0 m0 nxt -> sok_local ge sp rs0 m0 (si_local st)) ->
- 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.
- intros (EREF & LREF & NESTED) LREFN SOK; intuition.
- destruct NESTED as [|st0 se lse TOP NEST]; econstructor; simpl; auto.
-Qed.
-
-(** ** Assignment of registers *)
-
-(** locally new symbolic values during symbolic execution *)
-Inductive root_sval: Type :=
-| Rop (op: operation)
-| Rload (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing)
-.
-
-Definition root_apply (rsv: root_sval) (lr: list reg) (st: sistate_local): sval :=
- let lsv := list_sval_inj (List.map (si_sreg st) lr) in
- let sm := si_smem st in
- match rsv with
- | Rop op => Sop op lsv sm
- | Rload trap chunk addr => Sload sm trap chunk addr lsv
- end.
-Coercion root_apply: root_sval >-> Funclass.
-
-Definition root_happly (rsv: root_sval) (lr: list reg) (hst: hsistate_local) : ?? hsval :=
- DO lhsv <~ hlist_args hst lr;;
- match rsv with
- | Rop op => hSop op lhsv
- | Rload trap chunk addr => hSload hst trap chunk addr lhsv
- end.
-
-Lemma root_happly_correct (rsv: root_sval) lr hst:
- WHEN root_happly rsv lr hst ~> hv' THEN forall ge sp rs0 m0 st
- (REF:hsilocal_refines ge sp rs0 m0 hst st)
- (OK:hsok_local ge sp rs0 m0 hst),
- sval_refines ge sp rs0 m0 hv' (rsv lr st).
-Proof.
- unfold hsilocal_refines, root_apply, root_happly; destruct rsv; wlp_simplify.
- unfold sok_local in *.
- generalize (H0 ge sp rs0 m0 (list_sval_inj (map (si_sreg st) lr)) (si_smem st)); clear H0.
- destruct (seval_smem ge sp (si_smem st) rs0 m0) as [m|] eqn:X; eauto.
- intuition congruence.
-Qed.
-Global Opaque root_happly.
-Hint Resolve root_happly_correct: wlp.
-
-Local Open Scope lazy_bool_scope.
-
-(* NB: return [false] if the rsv cannot fail *)
-Definition may_trap (rsv: root_sval) (lr: list reg): bool :=
- match rsv with
- | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lr) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *)
- | Rload TRAP _ _ => true
- | _ => false
- end.
-
-Lemma lazy_orb_negb_false (b1 b2:bool):
- (b1 ||| negb b2) = false <-> (b1 = false /\ b2 = true).
-Proof.
- unfold negb; explore; simpl; intuition (try congruence).
-Qed.
-
-Lemma seval_list_sval_length ge sp rs0 m0 (f: reg -> sval) (l:list reg):
- forall l', seval_list_sval ge sp (list_sval_inj (List.map f l)) rs0 m0 = Some l' ->
- Datatypes.length l = Datatypes.length l'.
-Proof.
- induction l.
- - simpl. intros. inv H. reflexivity.
- - simpl. intros. destruct (seval_sval _ _ _ _ _); [|discriminate].
- destruct (seval_list_sval _ _ _ _ _) eqn:SLS; [|discriminate]. inv H. simpl.
- erewrite IHl; eauto.
-Qed.
-
-Lemma may_trap_correct (ge: RTL.genv) (sp:val) (rsv: root_sval) (rs0: regset) (m0: mem) (lr: list reg) st:
- may_trap rsv lr = false ->
- seval_list_sval ge sp (list_sval_inj (List.map (si_sreg st) lr)) rs0 m0 <> None ->
- seval_smem ge sp (si_smem st) rs0 m0 <> None ->
- seval_sval ge sp (rsv lr st) rs0 m0 <> None.
-Proof.
- destruct rsv; simpl; try congruence.
- - rewrite lazy_orb_negb_false. intros (TRAP1 & TRAP2) OK1 OK2.
- explore; try congruence.
- eapply is_trapping_op_sound; eauto.
- erewrite <- seval_list_sval_length; eauto.
- apply Nat.eqb_eq in TRAP2.
- assumption.
- - intros X OK1 OK2.
- explore; try congruence.
-Qed.
-
-(** simplify a symbolic value before assignment to a register *)
-Definition simplify (rsv: root_sval) (lr: list reg) (hst: hsistate_local): ?? hsval :=
- match rsv with
- | Rop op =>
- match is_move_operation op lr with
- | Some arg => hsi_sreg_get hst arg (* optimization of Omove *)
- | None =>
- match target_op_simplify op lr hst with
- | Some fhv => fsval_proj fhv
- | None =>
- DO lhsv <~ hlist_args hst lr;;
- hSop op lhsv
- end
- end
- | Rload _ chunk addr =>
- DO lhsv <~ hlist_args hst lr;;
- hSload hst NOTRAP chunk addr lhsv
- end.
-
-Lemma simplify_correct rsv lr hst:
- WHEN simplify rsv lr hst ~> hv THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st)
- (OK0: hsok_local ge sp rs0 m0 hst)
- (OK1: seval_sval ge sp (rsv lr st) rs0 m0 <> None),
- sval_refines ge sp rs0 m0 hv (rsv lr st).
-Proof.
- destruct rsv; simpl; auto.
- - (* Rop *)
- destruct (is_move_operation _ _) eqn: Hmove.
- { wlp_simplify; exploit is_move_operation_correct; eauto.
- intros (Hop & Hlsv); subst; simpl in *.
- simplify_SOME z.
- * erewrite H; eauto.
- * try_simplify_someHyps; congruence.
- * congruence. }
- destruct (target_op_simplify _ _ _) eqn: Htarget_op_simp; wlp_simplify.
- { destruct (seval_list_sval _ _ _) eqn: OKlist; try congruence.
- destruct (seval_smem _ _ _ _ _) eqn: OKmem; try congruence.
- rewrite <- H; exploit target_op_simplify_correct; eauto. }
- clear Htarget_op_simp.
- generalize (H0 ge sp rs0 m0 (list_sval_inj (map (si_sreg st) lr)) (si_smem st)); clear H0.
- destruct (seval_smem ge sp (si_smem st) rs0 m0) as [m|] eqn:X; eauto.
- intro H0; clear H0; simplify_SOME z; congruence. (* absurd case *)
- - (* Rload *)
- destruct trap; wlp_simplify.
- erewrite H0; eauto.
- erewrite H; eauto.
- erewrite hsilocal_refines_smem_refines; eauto.
- 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.
- destruct (Mem.loadv _ _ _); try congruence.
-Qed.
-Global Opaque simplify.
-Local Hint Resolve simplify_correct: wlp.
-
-Definition red_PTree_set (r: reg) (hsv: hsval) (hst: PTree.t hsval): PTree.t hsval :=
- match hsv with
- | HSinput r' _ =>
- if Pos.eq_dec r r'
- then PTree.remove r' hst
- else PTree.set r hsv hst
- | _ => PTree.set r hsv hst
- end.
-
-Lemma red_PTree_set_correct (r r0:reg) hsv hst ge sp rs0 m0:
- hsi_sreg_eval ge sp (red_PTree_set r hsv hst) r0 rs0 m0 = hsi_sreg_eval ge sp (PTree.set r hsv hst) r0 rs0 m0.
-Proof.
- destruct hsv; simpl; auto.
- destruct (Pos.eq_dec r r1); auto.
- subst; unfold hsi_sreg_eval, hsi_sreg_proj.
- destruct (Pos.eq_dec r0 r1); auto.
- - subst; rewrite PTree.grs, PTree.gss; simpl; auto.
- - rewrite PTree.gro, PTree.gso; simpl; auto.
-Qed.
-
-Lemma red_PTree_set_refines (r r0:reg) hsv hst sv st ge sp rs0 m0:
- hsilocal_refines ge sp rs0 m0 hst st ->
- sval_refines ge sp rs0 m0 hsv sv ->
- hsok_local ge sp rs0 m0 hst ->
- hsi_sreg_eval ge sp (red_PTree_set r hsv hst) r0 rs0 m0 = seval_sval ge sp (if Pos.eq_dec r r0 then sv else si_sreg st r0) rs0 m0.
-Proof.
- intros; rewrite red_PTree_set_correct.
- exploit hsilocal_refines_sreg; eauto.
- unfold hsi_sreg_eval, hsi_sreg_proj.
- destruct (Pos.eq_dec r r0); auto.
- - subst. rewrite PTree.gss; simpl; auto.
- - rewrite PTree.gso; simpl; eauto.
-Qed.
-
-Lemma sok_local_set_sreg (rsv:root_sval) ge sp rs0 m0 st r lr:
- sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st))
- <-> (sok_local ge sp rs0 m0 st /\ seval_sval ge sp (rsv lr st) rs0 m0 <> None).
-Proof.
- unfold slocal_set_sreg, sok_local; simpl; split.
- + intros ((SVAL0 & PRE) & SMEM & SVAL).
- repeat (split; try tauto).
- - intros r0; generalize (SVAL r0); clear SVAL; destruct (Pos.eq_dec r r0); try congruence.
- - generalize (SVAL r); clear SVAL; destruct (Pos.eq_dec r r); try congruence.
- + intros ((PRE & SMEM & SVAL0) & SVAL).
- repeat (split; try tauto; eauto).
- intros r0; destruct (Pos.eq_dec r r0); try congruence.
-Qed.
-
-Definition hslocal_set_sreg (hst: hsistate_local) (r: reg) (rsv: root_sval) (lr: list reg): ?? hsistate_local :=
- DO ok_lhsv <~
- (if may_trap rsv lr
- then DO hv <~ root_happly rsv lr hst;;
- XDEBUG hv (fun hv => DO hv_name <~ string_of_hashcode (hsval_get_hid hv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr hv_name))%string);;
- RET (hv::(hsi_ok_lsval hst))
- else RET (hsi_ok_lsval hst));;
- DO simp <~ simplify rsv lr hst;;
- RET {| hsi_smem := hst;
- hsi_ok_lsval := ok_lhsv;
- hsi_sreg := red_PTree_set r simp (hsi_sreg hst) |}.
-
-Lemma hslocal_set_sreg_correct hst r rsv lr:
- WHEN hslocal_set_sreg hst r rsv lr ~> hst' THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st),
- hsilocal_refines ge sp rs0 m0 hst' (slocal_set_sreg st r (rsv lr st)).
-Proof.
- wlp_simplify.
- + (* may_trap ~> true *)
- assert (X: sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st)) <->
- hsok_local ge sp rs0 m0 {| hsi_smem := hst; hsi_ok_lsval := exta :: hsi_ok_lsval hst; hsi_sreg := red_PTree_set r exta0 hst |}).
- { rewrite sok_local_set_sreg; generalize REF.
- intros (OKeq & MEM & REG & MVALID); rewrite OKeq; clear OKeq.
- unfold hsok_local; simpl; intuition (subst; eauto);
- erewrite <- H0 in *; eauto; unfold hsok_local; simpl; intuition eauto.
- }
- unfold hsilocal_refines; simpl; split; auto.
- rewrite <- X, sok_local_set_sreg. intuition eauto.
- - destruct REF; intuition eauto.
- - generalize REF; intros (OKEQ & _). rewrite OKEQ in * |-; erewrite red_PTree_set_refines; eauto.
- + (* may_trap ~> false *)
- assert (X: sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st)) <->
- hsok_local ge sp rs0 m0 {| hsi_smem := hst; hsi_ok_lsval := hsi_ok_lsval hst; hsi_sreg := red_PTree_set r exta hst |}).
- {
- rewrite sok_local_set_sreg; generalize REF.
- intros (OKeq & MEM & REG & MVALID); rewrite OKeq.
- unfold hsok_local; simpl; intuition (subst; eauto).
- assert (X0:hsok_local ge sp rs0 m0 hst). { unfold hsok_local; intuition. }
- exploit may_trap_correct; eauto.
- * intro X1; eapply seval_list_sval_inj_not_none; eauto.
- assert (X2: sok_local ge sp rs0 m0 st). { intuition. }
- unfold sok_local in X2; intuition eauto.
- * rewrite <- MEM; eauto.
- }
- unfold hsilocal_refines; simpl; split; auto.
- rewrite <- X, sok_local_set_sreg. intuition eauto.
- - destruct REF; intuition eauto.
- - generalize REF; intros (OKEQ & _). rewrite OKEQ in * |-; erewrite red_PTree_set_refines; eauto.
-Qed.
-Global Opaque hslocal_set_sreg.
-Local Hint Resolve hslocal_set_sreg_correct: wlp.
-
-(** ** Execution of one instruction *)
-
-(* TODO gourdinl
- * This is just useful for debugging fake values hashcode projection *)
-Fixpoint check_no_uhid lhsv :=
- match lhsv with
- | HSnil hc =>
- DO b <~ phys_eq hc unknown_hid;;
- assert_b (negb b) "fail no uhid";;
- RET tt
- | HScons hsv lhsv' hc =>
- DO b <~ phys_eq hc unknown_hid;;
- assert_b (negb b) "fail no uhid";;
- check_no_uhid lhsv'
- end.
-
-Definition cbranch_expanse (prev: hsistate_local) (cond: condition) (args: list reg): ?? (condition * list_hsval) :=
- match target_cbranch_expanse prev cond args with
- | Some (cond', vargs) =>
- DO vargs' <~ fsval_list_proj vargs;;
- RET (cond', vargs')
- | None =>
- DO vargs <~ hlist_args prev args ;;
- RET (cond, vargs)
- end.
-
-Lemma cbranch_expanse_correct hst c l:
- WHEN cbranch_expanse hst c l ~> r THEN forall ge sp rs0 m0 st
- (LREF : hsilocal_refines ge sp rs0 m0 hst st)
- (OK: hsok_local ge sp rs0 m0 hst),
- seval_condition ge sp (fst r) (hsval_list_proj (snd r)) (si_smem st) rs0 m0 =
- seval_condition ge sp c (list_sval_inj (map (si_sreg st) l)) (si_smem st) rs0 m0.
-Proof.
- unfold cbranch_expanse.
- destruct (target_cbranch_expanse _ _ _) eqn: TARGET; wlp_simplify;
- unfold seval_condition; erewrite <- H; eauto.
- destruct p as [c' l']; simpl.
- exploit target_cbranch_expanse_correct; eauto.
-Qed.
-Local Hint Resolve cbranch_expanse_correct: wlp.
-Global Opaque cbranch_expanse.
-
-Definition hsiexec_inst (i: instruction) (hst: hsistate): ?? (option hsistate) :=
- match i with
- | Inop pc' =>
- RET (Some (hsist_set_local hst pc' hst.(hsi_local)))
- | Iop op args dst pc' =>
- 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' =>
- 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' =>
- 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
- DO res <~ cbranch_expanse prev cond args;;
- let (cond, vargs) := res in
- let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in
- RET (Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |})
- | _ => RET None
- end.
-
-Remark hsiexec_inst_None_correct i hst:
- WHEN hsiexec_inst i hst ~> o THEN forall st, o = None -> siexec_inst i st = None.
-Proof.
- destruct i; wlp_simplify; congruence.
-Qed.
-
-Lemma seval_condition_refines hst st ge sp cond hargs args rs m:
- hsok_local ge sp rs m hst ->
- hsilocal_refines ge sp rs m hst st ->
- list_sval_refines ge sp rs m hargs args ->
- hseval_condition ge sp cond hargs (hsi_smem hst) rs m
- = seval_condition ge sp cond args (si_smem st) rs m.
- Proof.
- intros HOK (_ & MEMEQ & _) LR. unfold hseval_condition, seval_condition.
- rewrite LR, <- MEMEQ; auto.
-Qed.
-
-Lemma sok_local_set_sreg_simp (rsv:root_sval) ge sp rs0 m0 st r lr:
- sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st))
- -> sok_local ge sp rs0 m0 st.
-Proof.
- rewrite sok_local_set_sreg; intuition.
-Qed.
-
-Local Hint Resolve hsist_set_local_correct_stat: core.
-
-Lemma hsiexec_cond_noexp (hst: hsistate): forall l c0 n n0,
- WHEN DO res <~
- (DO vargs <~ hlist_args (hsi_local hst) l;; RET ((c0, vargs)));;
- (let (cond, vargs) := res in
- RET (Some
- {|
- hsi_pc := n0;
- hsi_exits := {|
- hsi_cond := cond;
- hsi_scondargs := vargs;
- hsi_elocal := hsi_local hst;
- hsi_ifso := n |} :: hsi_exits hst;
- hsi_local := hsi_local hst |})) ~> o0
- THEN (forall (hst' : hsistate) (st : sistate),
- o0 = Some hst' ->
- exists st' : sistate,
- Some
- {|
- si_pc := n0;
- si_exits := {|
- si_cond := c0;
- si_scondargs := list_sval_inj
- (map (si_sreg (si_local st)) l);
- si_elocal := si_local st;
- si_ifso := n |} :: si_exits st;
- si_local := si_local st |} = Some st' /\
- (hsistate_refines_stat hst st -> hsistate_refines_stat hst' st') /\
- (forall (ge : RTL.genv) (sp : val) (rs0 : regset) (m0 : mem),
- hsistate_refines_dyn ge sp rs0 m0 hst st ->
- hsistate_refines_dyn ge sp rs0 m0 hst' st')).
-Proof.
- intros.
- wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- - unfold hsistate_refines_stat, hsiexits_refines_stat in *; simpl; intuition.
- constructor; simpl; eauto.
- constructor.
- - destruct H0 as (EXREF & LREF & NEST).
- split.
- + constructor; simpl; auto.
- constructor; simpl; auto.
- intros; erewrite seval_condition_refines; eauto.
- + split; simpl; auto.
- destruct NEST as [|st0 se lse TOP NEST];
- econstructor; simpl; auto; constructor; auto.
-Qed.
-
-Lemma hsiexec_inst_correct i hst:
- WHEN hsiexec_inst i hst ~> o THEN forall hst' st,
- o = Some hst' ->
- exists st', siexec_inst i st = Some st'
- /\ (forall (REF:hsistate_refines_stat hst st), hsistate_refines_stat hst' st')
- /\ (forall ge sp rs0 m0 (REF:hsistate_refines_dyn ge sp rs0 m0 hst st), hsistate_refines_dyn ge sp rs0 m0 hst' st').
-Proof.
- destruct i; simpl;
- try (wlp_simplify; try_simplify_someHyps; eexists; intuition eauto; fail).
- - (* refines_dyn Iop *)
- wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- eapply hsist_set_local_correct_dyn; eauto.
- generalize (sok_local_set_sreg_simp (Rop o)); simpl; eauto.
- - (* refines_dyn Iload *)
- wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- eapply hsist_set_local_correct_dyn; eauto.
- generalize (sok_local_set_sreg_simp (Rload t0 m a)); simpl; eauto.
- - (* refines_dyn Istore *)
- wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- eapply hsist_set_local_correct_dyn; eauto.
- unfold sok_local; simpl; intuition.
- - (* refines_stat Icond *)
- wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- + unfold hsistate_refines_stat, hsiexits_refines_stat in *; simpl; intuition.
- constructor; simpl; eauto.
- constructor.
- + destruct REF as (EXREF & LREF & NEST).
- split.
- * constructor; simpl; auto.
- constructor; simpl; auto.
- intros; erewrite seval_condition_refines; eauto.
- * split; simpl; auto.
- destruct NEST as [|st0 se lse TOP NEST];
- econstructor; simpl; auto; constructor; auto.
-Qed.
-Global Opaque hsiexec_inst.
-Local Hint Resolve hsiexec_inst_correct: wlp.
-
-
-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 =>
- let pc := hst.(hsi_pc) in
- XDEBUG pc (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("- sym exec node: " +; name_pc)%string);;
- DO i <~ some_or_fail ((fn_code f)!pc) "hsiexec_path.internal_error.1";;
- DO ohst1 <~ hsiexec_inst i hst;;
- DO hst1 <~ some_or_fail ohst1 "hsiexec_path.internal_error.2";;
- hsiexec_path p f hst1
- end.
-
-Lemma hsiexec_path_correct path f: forall hst,
- WHEN hsiexec_path path f hst ~> hst' THEN forall st
- (RSTAT:hsistate_refines_stat hst st),
- exists st', siexec_path path f st = Some st'
- /\ hsistate_refines_stat hst' st'
- /\ (forall ge sp rs0 m0 (REF:hsistate_refines_dyn ge sp rs0 m0 hst st), hsistate_refines_dyn ge sp rs0 m0 hst' st').
-Proof.
- induction path; wlp_simplify; try_simplify_someHyps. clear IHpath.
- generalize RSTAT; intros (PCEQ & _) INSTEQ.
- rewrite <- PCEQ, INSTEQ; simpl.
- exploit H0; eauto. clear H0.
- intros (st0 & SINST & ISTAT & IDYN); erewrite SINST.
- exploit H1; eauto. clear H1.
- intros (st' & SPATH & PSTAT & PDYN).
- eexists; intuition eauto.
-Qed.
-Global Opaque hsiexec_path.
-Local Hint Resolve hsiexec_path_correct: wlp.
-
-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 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;;
- 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.
-
-Lemma hbuiltin_arg_correct hst arg:
- WHEN hbuiltin_arg hst arg ~> hargs THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- seval_builtin_sval ge sp (builtin_arg_map hsval_proj hargs) rs0 m0 = seval_builtin_sval ge sp (builtin_arg_map f arg) rs0 m0.
-Proof.
- induction arg; wlp_simplify.
- + erewrite H; eauto.
- + erewrite H; eauto.
- erewrite H0; eauto.
- + erewrite H; eauto.
- erewrite H0; eauto.
-Qed.
-Global Opaque hbuiltin_arg.
-Local Hint Resolve hbuiltin_arg_correct: wlp.
-
-Fixpoint hbuiltin_args (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.
-
-Lemma hbuiltin_args_correct hst args:
- WHEN hbuiltin_args hst args ~> hargs THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- bargs_refines ge sp rs0 m0 hargs (List.map (builtin_arg_map f) args).
-Proof.
- unfold bargs_refines, seval_builtin_args; induction args; wlp_simplify.
- erewrite H; eauto.
- erewrite H0; eauto.
-Qed.
-Global Opaque hbuiltin_args.
-Local Hint Resolve hbuiltin_args_correct: wlp.
-
-Definition hsum_left (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.
-
-Lemma hsum_left_correct hst ros:
- WHEN hsum_left hst ros ~> hsi THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- sum_refines ge sp rs0 m0 hsi (sum_left_map f ros).
-Proof.
- unfold sum_refines; destruct ros; wlp_simplify.
-Qed.
-Global Opaque hsum_left.
-Local Hint Resolve hsum_left_correct: wlp.
-
-Definition hsexec_final (i: instruction) (hst: PTree.t hsval): ?? hsfval :=
- match i with
- | Icall sig ros args 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 =>
- DO svos <~ hsum_left hst ros;;
- DO sargs <~ hlist_args hst args;;
- RET (HStailcall sig svos sargs)
- | Ibuiltin ef args res pc =>
- DO sargs <~ hbuiltin_args hst args;;
- RET (HSbuiltin ef sargs res pc)
- | Ijumptable reg tbl =>
- 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.
-
-Lemma hsexec_final_correct (hsl: hsistate_local) i:
- WHEN hsexec_final i hsl ~> hsf THEN forall ge sp rs0 m0 sl
- (OK: hsok_local ge sp rs0 m0 hsl)
- (REF: hsilocal_refines ge sp rs0 m0 hsl sl),
- hfinal_refines ge sp rs0 m0 hsf (sexec_final i sl).
-Proof.
- destruct i; wlp_simplify; try econstructor; simpl; eauto.
-Qed.
-Global Opaque hsexec_final.
-Local Hint Resolve hsexec_final_correct: wlp.
-
-Definition init_hsistate_local (_:unit): ?? hsistate_local
- := DO hm <~ hSinit ();;
- RET {| hsi_smem := hm; hsi_ok_lsval := nil; hsi_sreg := PTree.empty hsval |}.
-
-Lemma init_hsistate_local_correct:
- WHEN init_hsistate_local () ~> hsl THEN forall ge sp rs0 m0,
- hsilocal_refines ge sp rs0 m0 hsl init_sistate_local.
-Proof.
- unfold hsilocal_refines; wlp_simplify.
- - unfold hsok_local; simpl; intuition. erewrite H in *; congruence.
- - unfold hsok_local, sok_local; simpl in *; intuition; try congruence.
- - unfold hsi_sreg_eval, hsi_sreg_proj. rewrite PTree.gempty. reflexivity.
- - try_simplify_someHyps.
-Qed.
-Global Opaque init_hsistate_local.
-Local Hint Resolve init_hsistate_local_correct: wlp.
-
-Definition init_hsistate pc: ?? hsistate
- := DO hst <~ init_hsistate_local ();;
- RET {| hsi_pc := pc; hsi_exits := nil; hsi_local := hst |}.
-
-Lemma init_hsistate_correct pc:
- WHEN init_hsistate pc ~> hst THEN
- hsistate_refines_stat hst (init_sistate pc)
- /\ forall ge sp rs0 m0, hsistate_refines_dyn ge sp rs0 m0 hst (init_sistate pc).
-Proof.
- unfold hsistate_refines_stat, hsistate_refines_dyn, hsiexits_refines_dyn; wlp_simplify; constructor.
-Qed.
-Global Opaque init_hsistate.
-Local Hint Resolve init_hsistate_correct: wlp.
-
-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.
-
-Lemma hsexec_correct_aux f pc:
- WHEN hsexec f pc ~> hst THEN
- exists st, sexec f pc = Some st /\ hsstate_refines hst st.
-Proof.
- unfold hsstate_refines, sexec; wlp_simplify.
- - (* Some *)
- rewrite H; clear H.
- exploit H0; clear H0; eauto.
- intros (st0 & EXECPATH & SREF & DREF).
- rewrite EXECPATH; clear EXECPATH.
- generalize SREF. intros (EQPC & _).
- rewrite <- EQPC, H3; clear H3.
- exploit H4; clear H4; eauto.
- intros (st' & EXECL & SREF' & DREF').
- try_simplify_someHyps.
- eexists; intuition (simpl; eauto).
- constructor.
- - (* None *)
- rewrite H; clear H H4.
- exploit H0; clear H0; eauto.
- intros (st0 & EXECPATH & SREF & DREF).
- rewrite EXECPATH; clear EXECPATH.
- generalize SREF. intros (EQPC & _).
- rewrite <- EQPC, H3; clear H3.
- erewrite hsiexec_inst_None_correct; eauto.
- eexists; intuition (simpl; eauto).
-Qed.
-
-Global Opaque hsexec.
-
-End CanonBuilding.
-
-(** Correction of concrete symbolic execution wrt abstract symbolic execution *)
-Theorem hsexec_correct
- (hC_hsval : hashinfo hsval -> ?? hsval)
- (hC_list_hsval : hashinfo list_hsval -> ?? list_hsval)
- (hC_hsmem : hashinfo hsmem -> ?? hsmem)
- (f : function)
- (pc : node):
- WHEN hsexec hC_hsval hC_list_hsval hC_hsmem f pc ~> hst THEN forall
- (hC_hsval_correct: forall hs,
- WHEN hC_hsval hs ~> hs' THEN forall ge sp rs0 m0,
- seval_sval ge sp (hsval_proj (hdata hs)) rs0 m0 =
- seval_sval ge sp (hsval_proj hs') rs0 m0)
- (hC_list_hsval_correct: forall lh,
- WHEN hC_list_hsval lh ~> lh' THEN forall ge sp rs0 m0,
- seval_list_sval ge sp (hsval_list_proj (hdata lh)) rs0 m0 =
- seval_list_sval ge sp (hsval_list_proj lh') rs0 m0)
- (hC_hsmem_correct: forall hm,
- WHEN hC_hsmem hm ~> hm' THEN forall ge sp rs0 m0,
- seval_smem ge sp (hsmem_proj (hdata hm)) rs0 m0 =
- seval_smem ge sp (hsmem_proj hm') rs0 m0),
- exists st : sstate, sexec f pc = Some st /\ hsstate_refines hst st.
-Proof.
- wlp_simplify.
- eapply hsexec_correct_aux; eauto.
-Qed.
-Local Hint Resolve hsexec_correct: wlp.
-
-(** * Implementing the simulation test with concrete hash-consed symbolic execution *)
-
-Definition phys_check {A} (x y:A) (msg: pstring): ?? unit :=
- DO b <~ phys_eq x y;;
- assert_b b msg;;
- RET tt.
-
-Definition struct_check {A} (x y: A) (msg: pstring): ?? unit :=
- DO b <~ struct_eq x y;;
- assert_b b msg;;
- RET tt.
-
-Lemma struct_check_correct {A} (a b: A) msg:
- WHEN struct_check a b msg ~> _ THEN
- a = b.
-Proof. wlp_simplify. Qed.
-Global Opaque struct_check.
-Hint Resolve struct_check_correct: wlp.
-
-Definition 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.
-
-Lemma option_eq_check_correct A (o1 o2: option A): WHEN option_eq_check o1 o2 ~> _ THEN o1=o2.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque option_eq_check.
-Hint Resolve option_eq_check_correct:wlp.
-
-Import PTree.
-
-Fixpoint PTree_eq_check {A} (d1 d2: PTree.t A): ?? unit :=
- match d1, d2 with
- | Leaf, Leaf => RET tt
- | Node l1 o1 r1, Node l2 o2 r2 =>
- option_eq_check o1 o2;;
- PTree_eq_check l1 l2;;
- PTree_eq_check r1 r2
- | _, _ => FAILWITH "PTree_eq_check: some key is absent"
- end.
-
-Lemma PTree_eq_check_correct A d1: forall (d2: t A),
- WHEN PTree_eq_check d1 d2 ~> _ THEN forall x, PTree.get x d1 = PTree.get x d2.
-Proof.
- induction d1 as [|l1 Hl1 o1 r1 Hr1]; destruct d2 as [|l2 o2 r2]; simpl;
- wlp_simplify. destruct x; simpl; auto.
-Qed.
-Global Opaque PTree_eq_check.
-Local Hint Resolve PTree_eq_check_correct: wlp.
-
-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.
-
-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.
-Local Hint Resolve PTree_frame_eq_check_correct: wlp.
-
-Definition hsilocal_frame_simu_check frame hst1 hst2 : ?? unit :=
- DEBUG("? frame check");;
- phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_frame_simu_check: hsi_smem sets aren't equiv";;
- PTree_frame_eq_check frame (hsi_sreg hst1) (hsi_sreg hst2);;
- Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);;
- DEBUG("=> frame check: OK").
-
-Lemma setoid_in {A: Type} (a: A): forall l,
- SetoidList.InA (fun x y => x = y) a l ->
- In a l.
-Proof.
- induction l; intros; inv H.
- - constructor. reflexivity.
- - right. auto.
-Qed.
-
-Lemma regset_elements_in r rs:
- Regset.In r rs ->
- In r (Regset.elements rs).
-Proof.
- intros. exploit Regset.elements_1; eauto. intro SIN.
- apply setoid_in. assumption.
-Qed.
-Local Hint Resolve regset_elements_in: core.
-
-Lemma hsilocal_frame_simu_check_correct hst1 hst2 alive:
- WHEN hsilocal_frame_simu_check (Regset.elements alive) hst1 hst2 ~> _ THEN
- hsilocal_simu_spec alive hst1 hst2.
-Proof.
- unfold hsilocal_simu_spec; wlp_simplify. symmetry; eauto.
-Qed.
-Hint Resolve hsilocal_frame_simu_check_correct: wlp.
-Global Opaque hsilocal_frame_simu_check.
-
-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";;
- struct_check n res "revmap_check_single: n and res are physically different".
-
-Lemma revmap_check_single_correct dm pc1 pc2:
- WHEN revmap_check_single dm pc1 pc2 ~> _ THEN
- dm ! pc2 = Some pc1.
-Proof.
- wlp_simplify. congruence.
-Qed.
-Hint Resolve revmap_check_single_correct: wlp.
-Global Opaque revmap_check_single.
-
-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).
-
-Lemma hsiexit_simu_check_correct dm f hse1 hse2:
- WHEN hsiexit_simu_check dm f hse1 hse2 ~> _ THEN
- hsiexit_simu_spec dm f hse1 hse2.
-Proof.
- unfold hsiexit_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsiexit_simu_check_correct: wlp.
-Global Opaque hsiexit_simu_check.
-
-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.
-
-Lemma hsiexits_simu_check_correct dm f: forall le1 le2,
- WHEN hsiexits_simu_check dm f le1 le2 ~> _ THEN
- hsiexits_simu_spec dm f le1 le2.
-Proof.
- unfold hsiexits_simu_spec; induction le1; simpl; destruct le2; wlp_simplify; constructor; eauto.
-Qed.
-Hint Resolve hsiexits_simu_check_correct: wlp.
-Global Opaque hsiexits_simu_check.
-
-Definition hsistate_simu_check (dm: PTree.t node) (f: RTLpath.function) outframe (hst1 hst2: hsistate) :=
- hsiexits_simu_check dm f (hsi_exits hst1) (hsi_exits hst2);;
- hsilocal_frame_simu_check (Regset.elements outframe) (hsi_local hst1) (hsi_local hst2).
-
-Lemma hsistate_simu_check_correct dm f outframe hst1 hst2:
- WHEN hsistate_simu_check dm f outframe hst1 hst2 ~> _ THEN
- hsistate_simu_spec dm f outframe hst1 hst2.
-Proof.
- unfold hsistate_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsistate_simu_check_correct: wlp.
-Global Opaque hsistate_simu_check.
-
-
-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.
-
-Lemma revmap_check_list_correct dm: forall lpc lpc',
- WHEN revmap_check_list dm lpc lpc' ~> _ THEN
- ptree_get_list dm lpc' = Some lpc.
-Proof.
- induction lpc.
- - destruct lpc'; wlp_simplify.
- - destruct lpc'; wlp_simplify. try_simplify_someHyps.
-Qed.
-Global Opaque revmap_check_list.
-Hint Resolve revmap_check_list_correct: wlp.
-
-
-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.
-
-Lemma svos_simu_check_correct svos1 svos2:
- WHEN svos_simu_check svos1 svos2 ~> _ THEN
- svos1 = svos2.
-Proof.
- destruct svos1; destruct svos2; wlp_simplify.
-Qed.
-Global Opaque svos_simu_check.
-Hint Resolve svos_simu_check_correct: wlp.
-
-
-Fixpoint builtin_arg_simu_check (bs bs': builtin_arg hsval) :=
- match bs with
- | BA sv =>
- match bs' with
- | BA sv' => phys_check sv sv' "builtin_arg_simu_check: sval mismatch"
- | _ => FAILWITH "builtin_arg_simu_check: BA mismatch"
- end
- | BA_splitlong lo hi =>
- match bs' with
- | BA_splitlong lo' hi' =>
- builtin_arg_simu_check lo lo';;
- builtin_arg_simu_check hi hi'
- | _ => FAILWITH "builtin_arg_simu_check: BA_splitlong mismatch"
- end
- | BA_addptr b1 b2 =>
- match bs' with
- | BA_addptr b1' b2' =>
- builtin_arg_simu_check b1 b1';;
- builtin_arg_simu_check b2 b2'
- | _ => FAILWITH "builtin_arg_simu_check: BA_addptr mismatch"
- end
- | bs => struct_check bs bs' "builtin_arg_simu_check: basic mismatch"
- end.
-
-Lemma builtin_arg_simu_check_correct: forall bs1 bs2,
- WHEN builtin_arg_simu_check bs1 bs2 ~> _ THEN
- builtin_arg_map hsval_proj bs1 = builtin_arg_map hsval_proj bs2.
-Proof.
- induction bs1.
- all: try (wlp_simplify; subst; reflexivity).
- all: destruct bs2; wlp_simplify; congruence.
-Qed.
-Global Opaque builtin_arg_simu_check.
-Hint Resolve builtin_arg_simu_check_correct: wlp.
-
-Fixpoint list_builtin_arg_simu_check lbs1 lbs2 :=
- match lbs1, lbs2 with
- | nil, nil => RET tt
- | bs1::lbs1, bs2::lbs2 =>
- builtin_arg_simu_check bs1 bs2;;
- list_builtin_arg_simu_check lbs1 lbs2
- | _, _ => FAILWITH "list_builtin_arg_simu_check: length mismatch"
- end.
-
-Lemma list_builtin_arg_simu_check_correct: forall lbs1 lbs2,
- WHEN list_builtin_arg_simu_check lbs1 lbs2 ~> _ THEN
- List.map (builtin_arg_map hsval_proj) lbs1 = List.map (builtin_arg_map hsval_proj) lbs2.
-Proof.
- induction lbs1; destruct lbs2; wlp_simplify. congruence.
-Qed.
-Global Opaque list_builtin_arg_simu_check.
-Hint Resolve list_builtin_arg_simu_check_correct: wlp.
-
-Definition sfval_simu_check (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.
-
-Lemma sfval_simu_check_correct dm f opc1 opc2 fv1 fv2:
- WHEN sfval_simu_check dm f opc1 opc2 fv1 fv2 ~> _ THEN
- hfinal_simu_spec dm f opc1 opc2 fv1 fv2.
-Proof.
- unfold hfinal_simu_spec; destruct fv1; destruct fv2; wlp_simplify; try congruence.
-Qed.
-Hint Resolve sfval_simu_check_correct: wlp.
-Global Opaque sfval_simu_check.
-
-Definition hsstate_simu_check (dm: PTree.t node) (f: RTLpath.function) outframe (hst1 hst2: hsstate) :=
- hsistate_simu_check dm f outframe (hinternal hst1) (hinternal hst2);;
- sfval_simu_check dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2).
-
-Lemma hsstate_simu_check_correct dm f outframe hst1 hst2:
- WHEN hsstate_simu_check dm f outframe hst1 hst2 ~> _ THEN
- hsstate_simu_spec dm f outframe hst1 hst2.
-Proof.
- unfold hsstate_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsstate_simu_check_correct: wlp.
-Global Opaque hsstate_simu_check.
-
-Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node): ?? unit :=
- let (pc2, pc1) := m in
- (* creating the hash-consing tables *)
- DO hC_sval <~ hCons hSVAL;;
- DO hC_list_hsval <~ hCons hLSVAL;;
- DO hC_hsmem <~ hCons hSMEM;;
- let hsexec := hsexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in
- (* performing the hash-consed executions *)
- XDEBUG pc1 (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("entry-point of input superblock: " +; name_pc)%string);;
- DO hst1 <~ hsexec f pc1;;
- XDEBUG pc2 (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("entry-point of output superblock: " +; name_pc)%string);;
- DO hst2 <~ hsexec tf pc2;;
- DO path <~ some_or_fail ((fn_path f)!pc1) "simu_check_single.internal_error.1";;
- let outframe := path.(pre_output_regs) in
- (* comparing the executions *)
- hsstate_simu_check dm f outframe 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.
-Proof.
- unfold sexec_simu; wlp_simplify.
- exploit H2; clear H2. 1-3: wlp_simplify.
- intros (st2 & SEXEC2 & REF2). try_simplify_someHyps.
- exploit H3; clear H3. 1-3: wlp_simplify.
- intros (st3 & SEXEC3 & REF3). try_simplify_someHyps.
- eexists. eexists. split; eauto. split; eauto.
- intros ctx.
- eapply hsstate_simu_spec_correct; eauto.
-Qed.
-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 : ?? unit :=
- match lm with
- | 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 imp_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? unit :=
- simu_check_rec dm f tf (PTree.elements dm);;
- DEBUG("simu_check OK!").
-
-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). *)
-
-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.
-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.
diff --git a/scheduling/RTLpathSE_simu_specs.v b/scheduling/RTLpathSE_simu_specs.v
deleted file mode 100644
index 5051d805..00000000
--- a/scheduling/RTLpathSE_simu_specs.v
+++ /dev/null
@@ -1,937 +0,0 @@
-(** Low-level specifications of the simulation tests by symbolic execution with hash-consing *)
-
-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.
-Require Import RTLpathSE_theory RTLpathLivegenproof.
-Require Import Axioms.
-Require Import Lia.
-
-Local Open Scope error_monad_scope.
-Local Open Scope option_monad_scope.
-
-Require Export Impure.ImpHCons.
-Import HConsing.
-
-Import ListNotations.
-Local Open Scope list_scope.
-
-(** * Auxilary notions on simulation tests *)
-
-Definition silocal_simu (dm: PTree.t node) (f: RTLpath.function) outframe (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 outframe is1 is2.
-
-(* a kind of negation of sabort_local *)
-Definition sok_local (ge: RTL.genv) (sp:val) (rs0: regset) (m0: mem) (st: sistate_local): 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.
-
-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.
-
-Definition siexit_simu (dm: PTree.t node) (f: RTLpath.function) outframe (ctx: simu_proof_context f) (se1 se2: sistate_exit) :=
- (sok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (si_elocal 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)
- (si_smem (si_elocal se2)) (the_rs0 ctx) (the_m0 ctx)))
- /\ 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) ->
- 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 outframe is1 is2.
-
-Definition siexits_simu (dm: PTree.t node) (f: RTLpath.function) outframe (lse1 lse2: list sistate_exit) (ctx: simu_proof_context f) :=
- list_forall2 (siexit_simu dm f outframe ctx) lse1 lse2.
-
-
-(** * 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)
- | HSop (op: operation) (lhsv: list_hsval) (hid: hashcode) (** NB: does not depend on the memory ! *)
- | HSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (hid: hashcode)
-with list_hsval :=
- | HSnil (hid: hashcode)
- | HScons (hsv: hsval) (lhsv: list_hsval) (hid: hashcode)
-with hsmem :=
- | HSinit (hid: hashcode)
- | HSstore (hsm: hsmem) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (srce: hsval) (hid:hashcode).
-
-Scheme hsval_mut := Induction for hsval Sort Prop
-with list_hsval_mut := Induction for list_hsval Sort Prop
-with hsmem_mut := Induction for hsmem Sort Prop.
-
-
-
-(** 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: list_hsval) (res: reg) (pc: node)
- | HStailcall (sig: signature) (svos: hsval + ident) (lsv: list_hsval)
- | 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)
-.
-
-(** * gives the semantics of hash-consed symbolic values *)
-Fixpoint hsval_proj hsv :=
- match hsv with
- | HSinput r _ => Sinput r
- | HSop op hl _ => Sop op (hsval_list_proj hl) Sinit (** NB: use the initial memory of the path ! *)
- | HSload hm t chk addr hl _ => Sload (hsmem_proj hm) t chk addr (hsval_list_proj hl)
- end
-with hsval_list_proj hl :=
- match hl with
- | HSnil _ => Snil
- | HScons hv hl _ => Scons (hsval_proj hv) (hsval_list_proj hl)
- end
-with hsmem_proj hm :=
- match hm with
- | HSinit _ => Sinit
- | HSstore hm chk addr hl hv _ => Sstore (hsmem_proj hm) chk addr (hsval_list_proj hl) (hsval_proj hv)
- end.
-
-(** We use a Notation instead a Definition, in order to get more automation "for free" *)
-Notation "'seval_hsval' ge sp hsv" := (seval_sval ge sp (hsval_proj hsv))
- (only parsing, at level 0, ge at next level, sp at next level, hsv at next level): hse.
-
-Local Open Scope hse.
-
-Notation "'seval_list_hsval' ge sp lhv" := (seval_list_sval ge sp (hsval_list_proj lhv))
- (only parsing, at level 0, ge at next level, sp at next level, lhv at next level): hse.
-Notation "'seval_hsmem' ge sp hsm" := (seval_smem ge sp (hsmem_proj hsm))
- (only parsing, at level 0, ge at next level, sp at next level, hsm at next level): hse.
-
-Notation "'sval_refines' ge sp rs0 m0 hv sv" := (seval_hsval ge sp hv rs0 m0 = seval_sval ge sp sv rs0 m0)
- (only parsing, at level 0, ge at next level, sp at next level, rs0 at next level, m0 at next level, hv at next level, sv at next level): hse.
-Notation "'list_sval_refines' ge sp rs0 m0 lhv lsv" := (seval_list_hsval ge sp lhv rs0 m0 = seval_list_sval ge sp lsv rs0 m0)
- (only parsing, at level 0, ge at next level, sp at next level, rs0 at next level, m0 at next level, lhv at next level, lsv at next level): hse.
-Notation "'smem_refines' ge sp rs0 m0 hm sm" := (seval_hsmem ge sp hm rs0 m0 = seval_smem ge sp sm rs0 m0)
- (only parsing, at level 0, ge at next level, sp at next level, rs0 at next level, m0 at next level, hm at next level, sm at next level): hse.
-
-
-(** ** Implementation of symbolic states (with hash-consing) *)
-
-(** *** Syntax and semantics of symbolic internal local states
-
-The semantics is given by the refinement relation [hsilocal_refines] wrt to (abstract) symbolic internal local states
-
-*)
-
-(* NB: "h" stands for hash-consing *)
-Record hsistate_local :=
- {
- (** [hsi_smem] represents the current smem symbolic evaluations.
- (we also recover the history of smem in hsi_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_proj (hst: PTree.t hsval) r: sval :=
- match PTree.get r hst with
- | None => Sinput r
- | Some hsv => hsval_proj hsv
- end.
-
-Definition hsi_sreg_eval ge sp hst r := seval_sval ge sp (hsi_sreg_proj hst r).
-
-Definition hsok_local ge sp rs0 m0 (hst: hsistate_local) : Prop :=
- (forall hsv, List.In hsv (hsi_ok_lsval hst) -> seval_hsval ge sp hsv rs0 m0 <> None)
- /\ (seval_hsmem ge sp (hst.(hsi_smem)) rs0 m0 <> None).
-
-(* refinement link between a (st: sistate_local) and (hst: hsistate_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 -> smem_refines ge sp rs0 m0 (hsi_smem hst) (st.(si_smem)))
- /\ (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)
- /\ (* the below invariant allows to evaluate operations in the initial memory of the path instead of the current memory *)
- (forall m b ofs, seval_smem ge sp st.(si_smem) rs0 m0 = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
- .
-
-(** *** Syntax and semantics of symbolic exit states *)
-Record hsistate_exit := mk_hsistate_exit
- { hsi_cond: condition; hsi_scondargs: list_hsval; 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
-*)
-Definition hsiexit_refines_stat (hext: hsistate_exit) (ext: sistate_exit): Prop :=
- hsi_ifso hext = si_ifso ext.
-
-Definition hseval_condition ge sp cond hcondargs hmem rs0 m0 :=
- seval_condition ge sp cond (hsval_list_proj hcondargs) (hsmem_proj hmem) rs0 m0.
-
-Lemma hseval_condition_preserved ge ge' sp cond args mem rs0 m0:
- (forall s : ident, Genv.find_symbol ge' s = Genv.find_symbol ge s) ->
- hseval_condition ge sp cond args mem rs0 m0 = hseval_condition ge' sp cond args mem rs0 m0.
-Proof.
- intros. unfold hseval_condition. erewrite seval_condition_preserved; [|eapply H].
- reflexivity.
-Qed.
-
-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_local ge sp rs0 m0 (hsi_elocal hext) ->
- hseval_condition ge sp (hsi_cond hext) (hsi_scondargs hext) (hsi_smem (hsi_elocal hext)) rs0 m0
- = seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs0 m0).
-
-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 }.
-
-(* expresses the "monotony" of sok_local along sequences *)
-Inductive nested_sok ge sp rs0 m0: sistate_local -> list sistate_exit -> Prop :=
- nsok_nil st: nested_sok ge sp rs0 m0 st nil
- | nsok_cons st se lse:
- (sok_local ge sp rs0 m0 st -> sok_local ge sp rs0 m0 (si_elocal se)) ->
- nested_sok ge sp rs0 m0 (si_elocal se) lse ->
- nested_sok ge sp rs0 m0 st (se::lse).
-
-Lemma nested_sok_prop ge sp st sle rs0 m0:
- nested_sok ge sp rs0 m0 st sle ->
- sok_local ge sp rs0 m0 st ->
- forall se, In se sle -> sok_local ge sp rs0 m0 (si_elocal se).
-Proof.
- induction 1; simpl; intuition (subst; eauto).
-Qed.
-
-Lemma nested_sok_elocal ge sp rs0 m0 st2 exits:
- nested_sok ge sp rs0 m0 st2 exits ->
- forall st1, (sok_local ge sp rs0 m0 st1 -> sok_local ge sp rs0 m0 st2) ->
- nested_sok ge sp rs0 m0 st1 exits.
-Proof.
- induction 1; [intros; constructor|].
- intros. constructor; auto.
-Qed.
-
-Lemma nested_sok_tail ge sp rs0 m0 st lx exits:
- is_tail lx exits ->
- nested_sok ge sp rs0 m0 st exits ->
- nested_sok ge sp rs0 m0 st lx.
-Proof.
- induction 1; [auto|].
- intros. inv H0. eapply IHis_tail. eapply nested_sok_elocal; eauto.
-Qed.
-
-Definition hsistate_refines_stat (hst: hsistate) (st:sistate): Prop :=
- hsi_pc hst = si_pc st
- /\ hsiexits_refines_stat (hsi_exits hst) (si_exits st).
-
-Definition hsistate_refines_dyn ge sp rs0 m0 (hst: hsistate) (st:sistate): Prop :=
- 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)
- /\ nested_sok ge sp rs0 m0 (si_local st) (si_exits st) (* invariant necessary to prove "monotony" of sok_local along execution *)
- .
-
-(** *** Syntax and Semantics of symbolic state *)
-
-Definition hfinal_proj (hfv: hsfval) : sfval :=
- match hfv with
- | HSnone => Snone
- | HScall s hvi hlv r pc => Scall s (sum_left_map hsval_proj hvi) (hsval_list_proj hlv) r pc
- | HStailcall s hvi hlv => Stailcall s (sum_left_map hsval_proj hvi) (hsval_list_proj hlv)
- | HSbuiltin ef lbh br pc => Sbuiltin ef (List.map (builtin_arg_map hsval_proj) lbh) br pc
- | HSjumptable hv ln => Sjumptable (hsval_proj hv) ln
- | HSreturn oh => Sreturn (option_map hsval_proj oh)
- end.
-
-Section HFINAL_REFINES.
-
-Variable ge: RTL.genv.
-Variable sp: val.
-Variable rs0: regset.
-Variable m0: mem.
-
-Definition option_refines (ohsv: option hsval) (osv: option sval) :=
- match ohsv, osv with
- | Some hsv, Some sv => sval_refines ge sp rs0 m0 hsv sv
- | None, None => True
- | _, _ => False
- end.
-
-Definition sum_refines (hsi: hsval + ident) (si: sval + ident) :=
- match hsi, si with
- | inl hv, inl sv => sval_refines ge sp rs0 m0 hv sv
- | inr id, inr id' => id = id'
- | _, _ => False
- end.
-
-Definition bargs_refines (hargs: list (builtin_arg hsval)) (args: list (builtin_arg sval)): Prop :=
- seval_list_builtin_sval ge sp (List.map (builtin_arg_map hsval_proj) hargs) rs0 m0 = seval_list_builtin_sval ge sp args rs0 m0.
-
-Inductive hfinal_refines: hsfval -> sfval -> Prop :=
- | hsnone_ref: hfinal_refines HSnone Snone
- | hscall_ref: forall hros ros hargs args s r pc,
- sum_refines hros ros ->
- list_sval_refines ge sp rs0 m0 hargs args ->
- hfinal_refines (HScall s hros hargs r pc) (Scall s ros args r pc)
- | hstailcall_ref: forall hros ros hargs args s,
- sum_refines hros ros ->
- list_sval_refines ge sp rs0 m0 hargs args ->
- hfinal_refines (HStailcall s hros hargs) (Stailcall s ros args)
- | hsbuiltin_ref: forall ef lbha lba br pc,
- bargs_refines lbha lba ->
- hfinal_refines (HSbuiltin ef lbha br pc) (Sbuiltin ef lba br pc)
- | hsjumptable_ref: forall hsv sv lpc,
- sval_refines ge sp rs0 m0 hsv sv -> hfinal_refines (HSjumptable hsv lpc) (Sjumptable sv lpc)
- | hsreturn_ref: forall ohsv osv,
- option_refines ohsv osv -> hfinal_refines (HSreturn ohsv) (Sreturn osv).
-
-End HFINAL_REFINES.
-
-(* TODO gourdinl Leave this here ? *)
-Section FAKE_HSVAL.
-(* BEGIN "fake" hsval without real hash-consing *)
-(* TODO:
- 2) reuse these definitions in hSinput, hSop, etc
- in order to factorize proofs ?
-*)
-
-Definition fSinput (r: reg): hsval :=
- HSinput r unknown_hid.
-
-Lemma fSinput_correct r ge sp rs0 m0: (* useless trivial lemma ? *)
- sval_refines ge sp rs0 m0 (fSinput r) (Sinput r).
-Proof.
- auto.
-Qed.
-
-Definition fSop (op:operation) (lhsv: list_hsval): hsval :=
- HSop op lhsv unknown_hid.
-
-Lemma fSop_correct op lhsv ge sp rs0 m0 lsv sm m: forall
- (MEM: seval_smem ge sp sm rs0 m0 = Some m)
- (MVALID: forall b ofs, Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv),
- sval_refines ge sp rs0 m0 (fSop op lhsv) (Sop op lsv sm).
-Proof.
- intros; simpl. rewrite <- LR, MEM.
- destruct (seval_list_sval _ _ _ _); try congruence.
- eapply op_valid_pointer_eq; eauto.
-Qed.
-
-Definition fsi_sreg_get (hst: PTree.t hsval) r: hsval :=
- match PTree.get r hst with
- | None => fSinput r
- | Some sv => sv
- end.
-
-Lemma fsi_sreg_get_correct hst r ge sp rs0 m0 (f: reg -> sval): forall
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- sval_refines ge sp rs0 m0 (fsi_sreg_get hst r) (f r).
-Proof.
- unfold hsi_sreg_eval, hsi_sreg_proj, fsi_sreg_get; intros; simpl.
- rewrite <- RR. destruct (hst ! r); simpl; auto.
-Qed.
-
-Definition fSnil: list_hsval :=
- HSnil unknown_hid.
-
-(* TODO: Lemma fSnil_correct *)
-
-Definition fScons (hsv: hsval) (lhsv: list_hsval): list_hsval :=
- HScons hsv lhsv unknown_hid.
-
-(* TODO: Lemma fScons_correct *)
-
-(* END "fake" hsval ... *)
-
-End FAKE_HSVAL.
-
-
-Record hsstate := { hinternal:> hsistate; hfinal: hsfval }.
-
-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))
- /\ (forall ge sp rs0 m0, hsok_local ge sp rs0 m0 (hsi_local (hinternal hst)) -> hfinal_refines ge sp rs0 m0 (hfinal hst) (final st))
- .
-
-(** * Intermediate specifications of the simulation tests *)
-
-(** ** Specification of the simulation test on [hsistate_local].
- It is motivated by [hsilocal_simu_spec_correct theorem] below
-*)
-Definition hsilocal_simu_spec (alive: Regset.t) (hst1 hst2: hsistate_local) :=
- List.incl (hsi_ok_lsval hst2) (hsi_ok_lsval hst1)
- /\ (forall r, Regset.In r alive -> PTree.get r hst2 = PTree.get r hst1)
- /\ hsi_smem hst1 = hsi_smem hst2.
-
-Definition seval_sval_partial ge sp rs0 m0 hsv :=
- match seval_hsval ge sp hsv rs0 m0 with
- | Some v => v
- | None => Vundef
- end.
-
-Definition select_first (ox oy: option val) :=
- match ox with
- | Some v => Some v
- | None => oy
- end.
-
-(** If the register was computed by hrs, evaluate the symbolic value from hrs.
- Else, take the value directly from rs0 *)
-Definition seval_partial_regset ge sp rs0 m0 hrs :=
- let hrs_eval := PTree.map1 (seval_sval_partial ge sp rs0 m0) hrs in
- (fst rs0, PTree.combine select_first hrs_eval (snd rs0)).
-
-Lemma seval_partial_regset_get ge sp rs0 m0 hrs r:
- (seval_partial_regset ge sp rs0 m0 hrs) # r =
- match (hrs ! r) with Some sv => seval_sval_partial ge sp rs0 m0 sv | None => (rs0 # r) end.
-Proof.
- unfold seval_partial_regset. unfold Regmap.get. simpl.
- rewrite PTree.gcombine; [| simpl; reflexivity]. rewrite PTree.gmap1.
- destruct (hrs ! r); simpl; [reflexivity|].
- destruct ((snd rs0) ! r); reflexivity.
-Qed.
-
-Lemma ssem_local_refines_hok ge sp rs0 m0 hst st rs m:
- 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.
-
-Lemma hsilocal_simu_spec_nofail ge1 ge2 of sp rs0 m0 hst1 hst2:
- hsilocal_simu_spec of 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 (RSOK & _ & MEMOK) GFS (OKV & OKM). constructor.
- - intros sv INS. apply RSOK in INS. apply OKV in INS. erewrite seval_preserved; eauto.
- - erewrite MEMOK in OKM. erewrite smem_eval_preserved; eauto.
-Qed.
-
-Theorem hsilocal_simu_spec_correct hst1 hst2 alive ge1 ge2 sp rs0 m0 rs m st1 st2:
- hsilocal_simu_spec alive 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 ->
- let rs' := seval_partial_regset ge2 sp rs0 m0 (hsi_sreg hst2)
- in ssem_local ge2 sp st2 rs0 m0 rs' m /\ eqlive_reg (fun r => Regset.In r alive) rs rs'.
-Proof.
- intros CORE HREF1 HREF2 GFS SEML.
- refine (modusponens _ _ (ssem_local_refines_hok _ _ _ _ _ _ _ _ _ _) _); eauto.
- intro HOK1.
- refine (modusponens _ _ (hsilocal_simu_spec_nofail _ _ _ _ _ _ _ _ _ _ _) _); eauto.
- intro HOK2.
- destruct SEML as (PRE & MEMEQ & RSEQ).
- assert (SIPRE: si_pre st2 ge2 sp rs0 m0). { destruct HREF2 as (OKEQ & _ & _). rewrite <- OKEQ in HOK2. apply HOK2. }
- assert (SMEMEVAL: seval_smem ge2 sp (si_smem st2) rs0 m0 = Some m). {
- destruct HREF2 as (_ & MEMEQ2 & _). destruct HREF1 as (_ & MEMEQ1 & _).
- destruct CORE as (_ & _ & MEMEQ3).
- rewrite <- MEMEQ2; auto. rewrite <- MEMEQ3.
- erewrite smem_eval_preserved; [| eapply GFS].
- rewrite MEMEQ1; auto. }
- constructor.
- + constructor; [assumption | constructor; [assumption|]].
- destruct HREF2 as (B & _ & A & _).
- (** B is used for the auto below. *)
- assert (forall r : positive, hsi_sreg_eval ge2 sp hst2 r rs0 m0 = seval_sval ge2 sp (si_sreg st2 r) rs0 m0) by auto.
- intro r. rewrite <- H. clear H.
- generalize (A HOK2 r). unfold hsi_sreg_eval.
- rewrite seval_partial_regset_get.
- unfold hsi_sreg_proj.
- destruct (hst2 ! r) eqn:HST2; [| simpl; reflexivity].
- unfold seval_sval_partial. generalize HOK2; rewrite <- B; intros (_ & _ & C) D.
- assert (seval_sval ge2 sp (hsval_proj h) rs0 m0 <> None) by congruence.
- destruct (seval_sval ge2 sp _ rs0 m0); [reflexivity | contradiction].
- + intros r ALIVE. destruct HREF2 as (_ & _ & A & _). destruct HREF1 as (_ & _ & B & _).
- destruct CORE as (_ & C & _). rewrite seval_partial_regset_get.
- assert (OPT: forall (x y: val), Some x = Some y -> x = y) by congruence.
- destruct (hst2 ! r) eqn:HST2; apply OPT; clear OPT.
- ++ unfold seval_sval_partial.
- assert (seval_sval ge2 sp (hsval_proj h) rs0 m0 = hsi_sreg_eval ge2 sp hst2 r rs0 m0). {
- unfold hsi_sreg_eval, hsi_sreg_proj. rewrite HST2. reflexivity. }
- rewrite H. clear H. unfold hsi_sreg_eval, hsi_sreg_proj. rewrite C; [|assumption].
- erewrite seval_preserved; [| eapply GFS].
- unfold hsi_sreg_eval, hsi_sreg_proj in B; rewrite B; [|assumption]. rewrite RSEQ. reflexivity.
- ++ rewrite <- RSEQ. rewrite <- B; [|assumption]. unfold hsi_sreg_eval, hsi_sreg_proj.
- rewrite <- C; [|assumption]. rewrite HST2. reflexivity.
-Qed.
-
-(** ** Specification of the simulation test on [hsistate_exit].
- It is motivated by [hsiexit_simu_spec_correct theorem] below
-*)
-Definition hsiexit_simu_spec dm f (hse1 hse2: hsistate_exit) :=
- (exists path, (fn_path f) ! (hsi_ifso hse1) = Some path
- /\ hsilocal_simu_spec path.(input_regs) (hsi_elocal hse1) (hsi_elocal hse2))
- /\ dm ! (hsi_ifso hse2) = Some (hsi_ifso hse1)
- /\ hsi_cond hse1 = hsi_cond hse2
- /\ hsi_scondargs hse1 = hsi_scondargs hse2.
-
-Definition hsiexit_simu dm f outframe (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 ->
- hsiexit_refines_dyn (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hse2 se2 ->
- siexit_simu dm f outframe ctx se1 se2.
-
-Lemma hsiexit_simu_spec_nofail dm f hse1 hse2 ge1 ge2 sp rs m:
- hsiexit_simu_spec dm f hse1 hse2 ->
- (forall s, Genv.find_symbol ge1 s = Genv.find_symbol ge2 s) ->
- hsok_local ge1 sp rs m (hsi_elocal hse1) ->
- hsok_local ge2 sp rs m (hsi_elocal hse2).
-Proof.
- intros CORE GFS HOK1.
- destruct CORE as ((p & _ & CORE') & _ & _ & _).
- eapply hsilocal_simu_spec_nofail; eauto.
-Qed.
-
-Theorem hsiexit_simu_spec_correct dm f outframe hse1 hse2 ctx:
- hsiexit_simu_spec dm f hse1 hse2 ->
- hsiexit_simu dm f outframe ctx hse1 hse2.
-Proof.
- intros SIMUC st1 st2 HREF1 HREF2 HDYN1 HDYN2.
- assert (SEVALC:
- sok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (si_elocal st1) ->
- (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 ((OKEQ1 & _) & SCOND1).
- rewrite OKEQ1; intro OK1. rewrite <- SCOND1 by assumption. clear SCOND1.
- generalize (genv_match ctx).
- intro GFS; exploit hsiexit_simu_spec_nofail; eauto.
- destruct HDYN2 as (_ & SCOND2). intro OK2. rewrite <- SCOND2 by assumption. clear OK1 OK2 SCOND2.
- destruct SIMUC as ((path & _ & LSIMU) & _ & CONDEQ & ARGSEQ). destruct LSIMU as (_ & _ & MEMEQ).
- rewrite CONDEQ. rewrite ARGSEQ. rewrite MEMEQ. erewrite <- hseval_condition_preserved; eauto.
- }
- constructor; [assumption|]. intros is1 ICONT SSEME.
- assert (OK1: sok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (si_elocal st1)). {
- destruct SSEME as (_ & SSEML & _). eapply ssem_local_sok; eauto. }
- assert (HOK1: hsok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (hsi_elocal hse1)). {
- destruct HDYN1 as (LREF & _). destruct LREF as (OKEQ & _ & _). rewrite <- OKEQ. assumption. }
- exploit hsiexit_simu_spec_nofail. 2: eapply ctx. all: eauto. intro HOK2.
- destruct SSEME as (SCOND & SLOC & PCEQ). destruct SIMUC as ((path & PATH & LSIMU) & REVEQ & _ & _); eauto.
- destruct HDYN1 as (LREF1 & _). destruct HDYN2 as (LREF2 & _).
- exploit hsilocal_simu_spec_correct; eauto; [apply ctx|]. simpl.
- intros (SSEML & EQREG).
- eexists (mk_istate (icontinue is1) (si_ifso st2) _ (imem is1)). simpl. constructor.
- - constructor; intuition congruence || eauto.
- - unfold istate_simu. rewrite ICONT.
- simpl. assert (PCEQ': hsi_ifso hse1 = ipc is1) by congruence.
- exists path. constructor; [|constructor]; [congruence| |congruence].
- constructor; [|constructor]; simpl; auto.
-Qed.
-
-Remark hsiexit_simu_siexit dm f outframe ctx hse1 hse2 se1 se2:
- hsiexit_simu dm f outframe 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 outframe ctx se1 se2.
-Proof.
- auto.
-Qed.
-
-(** ** Specification of the simulation test on [list hsistate_exit].
- It is motivated by [hsiexit_simu_spec_correct theorem] below
-*)
-
-Definition hsiexits_simu dm f outframe (ctx: simu_proof_context f) (lhse1 lhse2: list hsistate_exit): Prop :=
- list_forall2 (hsiexit_simu dm f outframe ctx) lhse1 lhse2.
-
-Definition hsiexits_simu_spec dm f lhse1 lhse2: Prop :=
- list_forall2 (hsiexit_simu_spec dm f) lhse1 lhse2.
-
-Theorem hsiexits_simu_spec_correct dm f outframe lhse1 lhse2 ctx:
- hsiexits_simu_spec dm f lhse1 lhse2 ->
- hsiexits_simu dm f outframe ctx lhse1 lhse2.
-Proof.
- induction 1; [constructor|].
- constructor; [|apply IHlist_forall2; assumption].
- apply hsiexit_simu_spec_correct; assumption.
-Qed.
-
-
-Lemma siexits_simu_all_fallthrough dm f outframe ctx: forall lse1 lse2,
- siexits_simu dm f outframe lse1 lse2 ctx ->
- all_fallthrough (the_ge1 ctx) (the_sp ctx) lse1 (the_rs0 ctx) (the_m0 ctx) ->
- (forall se1, In se1 lse1 -> sok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (si_elocal se1)) ->
- all_fallthrough (the_ge2 ctx) (the_sp ctx) lse2 (the_rs0 ctx) (the_m0 ctx).
-Proof.
- induction 1; [unfold all_fallthrough; contradiction|]; simpl.
- intros X OK 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.
- - intros; intuition.
-Qed.
-
-
-Lemma siexits_simu_all_fallthrough_upto dm f outframe ctx lse1 lse2:
- siexits_simu dm f outframe lse1 lse2 ctx ->
- forall ext1 lx1,
- (forall se1, In se1 lx1 -> sok_local (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) (si_elocal se1)) ->
- 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 ext lx1. intros OK H. destruct H as (ITAIL & ALLFU). eapply is_tail_false in ITAIL. contradiction.
- - simpl; intros ext lx1 OK 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.
- * intuition. apply OK. eassumption.
- * constructor; eauto.
- * intros (ext2 & lx2 & ALLFUE2 & LENEQ).
- eexists; eexists. constructor; eauto.
- eapply all_fallthrough_upto_exit_cons; eauto.
-Qed.
-
-
-Lemma hsiexits_simu_siexits dm f outframe ctx lhse1 lhse2:
- hsiexits_simu dm f outframe 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 outframe lse1 lse2 ctx.
-Proof.
- 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.
-
-
-(** ** Specification of the simulation test on [hsistate].
- It is motivated by [hsistate_simu_spec_correct theorem] below
-*)
-
-Definition hsistate_simu_spec dm f outframe (hse1 hse2: hsistate) :=
- list_forall2 (hsiexit_simu_spec dm f) (hsi_exits hse1) (hsi_exits hse2)
- /\ hsilocal_simu_spec outframe (hsi_local hse1) (hsi_local hse2).
-
-Definition hsistate_simu dm f outframe (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 outframe st1 st2 ctx.
-
-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. lia.
-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; lia). rewrite H. clear H.
- inv ITAIL.
- + assert (forall n, (n - n)%nat = 0%nat) by (intro; lia). 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; lia).
- exploit (is_tail_length (x::l1)); eauto. intro. simpl in H2.
- assert ((length l2 > length l1)%nat) by lia. clear H2.
- rewrite H0; auto.
-Qed.
-
-Theorem hsistate_simu_spec_correct dm f outframe hst1 hst2 ctx:
- hsistate_simu_spec dm f outframe hst1 hst2 ->
- hsistate_simu dm f outframe hst1 hst2 ctx.
-Proof.
- intros (ESIMU & LSIMU) st1 st2 (PCREF1 & EREF1) (PCREF2 & EREF2) DREF1 DREF2 is1 SEMI.
- destruct DREF1 as (DEREF1 & LREF1 & NESTED). destruct DREF2 as (DEREF2 & LREF2 & _).
- exploit hsiexits_simu_spec_correct; eauto. intro HESIMU.
- unfold ssem_internal in SEMI. destruct (icontinue _) eqn:ICONT.
- - destruct SEMI as (SSEML & PCEQ & ALLFU).
- exploit hsilocal_simu_spec_correct; eauto; [apply ctx|]. simpl. intro SSEML2.
- exists (mk_istate (icontinue is1) (si_pc st2) (seval_partial_regset (the_ge2 ctx) (the_sp ctx)
- (the_rs0 ctx) (the_m0 ctx) (hsi_local hst2)) (imem is1)). constructor.
- + unfold ssem_internal. simpl. rewrite ICONT.
- destruct SSEML2 as [SSEMLP EQLIVE].
- constructor; [assumption | constructor; [reflexivity |]].
- eapply siexits_simu_all_fallthrough; eauto.
- * eapply hsiexits_simu_siexits; eauto.
- * eapply nested_sok_prop; eauto.
- eapply ssem_local_sok; eauto.
- + unfold istate_simu. rewrite ICONT.
- destruct SSEML2 as [SSEMLP EQLIVE].
- constructor; simpl; auto.
- - destruct SEMI as (ext & lx & SSEME & ALLFU).
- assert (SESIMU: siexits_simu dm f outframe (si_exits st1) (si_exits st2) ctx) by (eapply hsiexits_simu_siexits; eauto).
- exploit siexits_simu_all_fallthrough_upto; eauto.
- * destruct ALLFU as (ITAIL & ALLF).
- exploit nested_sok_tail; eauto. intros NESTED2.
- inv NESTED2. destruct SSEME as (_ & SSEML & _). eapply ssem_local_sok in SSEML.
- eapply nested_sok_prop; eauto.
- * intros (ext2 & lx2 & ALLFU2 & LENEQ).
- assert (EXTSIMU: siexit_simu dm f outframe 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. }
- destruct EXTSIMU as (CONDEVAL & EXTSIMU).
- apply EXTSIMU in SSEME; [|assumption]. 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).
- destruct ISIMULIVE as (CONTEQ & REGEQ & MEMEQ).
- exists path. repeat (constructor; auto).
-Qed.
-
-
-(** ** Specification of the simulation test on [sfval].
- It is motivated by [hfinal_simu_spec_correct theorem] below
-*)
-
-
-Definition final_simu_spec (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (f1 f2: sfval): Prop :=
- match f1 with
- | Scall sig1 svos1 lsv1 res1 pc1 =>
- match f2 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 f2 with
- | Sbuiltin ef2 lbs2 br2 pc2 =>
- dm ! pc2 = Some pc1 /\ ef1 = ef2 /\ lbs1 = lbs2 /\ br1 = br2
- | _ => False
- end
- | Sjumptable sv1 lpc1 =>
- match f2 with
- | Sjumptable sv2 lpc2 =>
- ptree_get_list dm lpc2 = Some lpc1 /\ sv1 = sv2
- | _ => False
- end
- | Snone =>
- match f2 with
- | Snone => dm ! pc2 = Some pc1
- | _ => False
- end
- (* Stailcall, Sreturn *)
- | _ => f1 = f2
- end.
-
-Definition hfinal_simu_spec (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (hf1 hf2: hsfval): Prop :=
- final_simu_spec dm f pc1 pc2 (hfinal_proj hf1) (hfinal_proj hf2).
-
-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.
-
-Lemma list_proj_refines_eq ge ge' sp rs0 m0 lsv lhsv:
- (forall s, Genv.find_symbol ge s = Genv.find_symbol ge' s) ->
- list_sval_refines ge sp rs0 m0 lhsv lsv ->
- forall lhsv' lsv',
- list_sval_refines ge' sp rs0 m0 lhsv' lsv' ->
- hsval_list_proj lhsv = hsval_list_proj lhsv' ->
- seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv' rs0 m0.
-Proof.
- intros GFS H lhsv' lsv' H' H0.
- erewrite <- H, H0.
- erewrite list_sval_eval_preserved; eauto.
-Qed.
-
-Lemma seval_builtin_sval_preserved ge ge' sp sv rs0 m0:
- (forall s : ident, Genv.find_symbol ge' s = Genv.find_symbol ge s) ->
- seval_builtin_sval ge sp sv rs0 m0 =
- seval_builtin_sval ge' sp sv rs0 m0.
-Proof.
- induction sv; intro FIND; cbn.
- all: try (erewrite seval_preserved by eauto); trivial.
- all: erewrite IHsv1 by eauto; erewrite IHsv2 by eauto; reflexivity.
-Qed.
-
-Lemma seval_list_builtin_sval_preserved ge ge' sp lsv rs0 m0:
- (forall s : ident, Genv.find_symbol ge' s = Genv.find_symbol ge s) ->
- seval_list_builtin_sval ge sp lsv rs0 m0 =
- seval_list_builtin_sval ge' sp lsv rs0 m0.
-Proof.
- induction lsv; intro FIND; cbn. { trivial. }
- erewrite seval_builtin_sval_preserved by eauto.
- erewrite IHlsv by eauto.
- reflexivity.
-Qed.
-
-Lemma barg_proj_refines_eq ge ge' sp rs0 m0:
- (forall s, Genv.find_symbol ge s = Genv.find_symbol ge' s) ->
- forall lhsv lsv, bargs_refines ge sp rs0 m0 lhsv lsv ->
- forall lhsv' lsv', bargs_refines ge' sp rs0 m0 lhsv' lsv' ->
- List.map (builtin_arg_map hsval_proj) lhsv = List.map (builtin_arg_map hsval_proj) lhsv' ->
- seval_list_builtin_sval ge sp lsv rs0 m0 = seval_list_builtin_sval ge' sp lsv' rs0 m0.
-Proof.
- unfold bargs_refines; intros GFS lhsv lsv H lhsv' lsv' H' H0.
- erewrite <- H, H0.
- erewrite seval_list_builtin_sval_preserved; eauto.
-Qed.
-
-Lemma sval_refines_proj ge ge' sp rs m hsv sv hsv' sv':
- (forall s, Genv.find_symbol ge s = Genv.find_symbol ge' s) ->
- sval_refines ge sp rs m hsv sv ->
- sval_refines ge' sp rs m hsv' sv' ->
- hsval_proj hsv = hsval_proj hsv' ->
- seval_sval ge sp sv rs m = seval_sval ge' sp sv' rs m.
-Proof.
- intros GFS REF REF' PROJ.
- rewrite <- REF, PROJ.
- erewrite <- seval_preserved; eauto.
-Qed.
-
-Theorem hfinal_simu_spec_correct dm f ctx opc1 opc2 hf1 hf2 f1 f2:
- hfinal_simu_spec dm f opc1 opc2 hf1 hf2 ->
- hfinal_refines (the_ge1 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hf1 f1 ->
- hfinal_refines (the_ge2 ctx) (the_sp ctx) (the_rs0 ctx) (the_m0 ctx) hf2 f2 ->
- sfval_simu dm f opc1 opc2 ctx f1 f2.
-Proof.
- assert (GFS: forall s : ident, Genv.find_symbol (the_ge1 ctx) s = Genv.find_symbol (the_ge2 ctx) s) by apply ctx.
- intros CORE FREF1 FREF2.
- destruct hf1; inv FREF1.
- (* Snone *)
- - destruct hf2; try contradiction. inv FREF2.
- inv CORE. constructor. assumption.
- (* Scall *)
- - rename H5 into SREF1. rename H6 into LREF1.
- destruct hf2; try contradiction. inv FREF2.
- rename H5 into SREF2. rename H6 into LREF2.
- destruct CORE as (PCEQ & ? & ? & ? & ?). subst.
- rename H0 into SVOSEQ. rename H1 into LSVEQ.
- constructor; [assumption | |].
- + destruct svos.
- * destruct svos0; try discriminate. destruct ros; try contradiction.
- destruct ros0; try contradiction. constructor.
- simpl in SVOSEQ. inv SVOSEQ.
- simpl in SREF1. simpl in SREF2.
- rewrite <- SREF1. rewrite <- SREF2.
- erewrite <- seval_preserved; [| eapply GFS]. congruence.
- * destruct svos0; try discriminate. destruct ros; try contradiction.
- destruct ros0; try contradiction. constructor.
- simpl in SVOSEQ. inv SVOSEQ. congruence.
- + erewrite list_proj_refines_eq; eauto.
- (* Stailcall *)
- - rename H3 into SREF1. rename H4 into LREF1.
- destruct hf2; try (inv CORE; fail). inv FREF2.
- rename H4 into LREF2. rename H3 into SREF2.
- inv CORE. rename H1 into SVOSEQ. rename H2 into LSVEQ.
- constructor.
- + destruct svos. (** Copy-paste from Scall *)
- * destruct svos0; try discriminate. destruct ros; try contradiction.
- destruct ros0; try contradiction. constructor.
- simpl in SVOSEQ. inv SVOSEQ.
- simpl in SREF1. simpl in SREF2.
- rewrite <- SREF1. rewrite <- SREF2.
- erewrite <- seval_preserved; [| eapply GFS]. congruence.
- * destruct svos0; try discriminate. destruct ros; try contradiction.
- destruct ros0; try contradiction. constructor.
- simpl in SVOSEQ. inv SVOSEQ. congruence.
- + erewrite list_proj_refines_eq; eauto.
- (* Sbuiltin *)
- - rename H4 into BREF1. destruct hf2; try (inv CORE; fail). inv FREF2.
- rename H4 into BREF2. inv CORE. destruct H0 as (? & ? & ?). subst.
- rename H into PCEQ. rename H1 into ARGSEQ. constructor; [assumption|].
- erewrite barg_proj_refines_eq; eauto. constructor.
- (* Sjumptable *)
- - rename H2 into SREF1. destruct hf2; try contradiction. inv FREF2.
- rename H2 into SREF2. destruct CORE as (A & B). constructor; [assumption|].
- erewrite sval_refines_proj; eauto.
- (* Sreturn *)
- - rename H0 into SREF1.
- destruct hf2; try discriminate. inv CORE.
- inv FREF2. destruct osv; destruct res; inv SREF1.
- + destruct res0; try discriminate. destruct osv0; inv H1.
- constructor. simpl in H0. inv H0. erewrite sval_refines_proj; eauto.
- + destruct res0; try discriminate. destruct osv0; inv H1. constructor.
-Qed.
-
-
-(** ** Specification of the simulation test on [hsstate].
- It is motivated by [hsstate_simu_spec_correct theorem] below
-*)
-
-Definition hsstate_simu_spec (dm: PTree.t node) (f: RTLpath.function) outframe (hst1 hst2: hsstate) :=
- hsistate_simu_spec dm f outframe (hinternal hst1) (hinternal hst2)
- /\ hfinal_simu_spec dm f (hsi_pc (hinternal hst1)) (hsi_pc (hinternal hst2)) (hfinal hst1) (hfinal hst2).
-
-Definition hsstate_simu dm f outframe (hst1 hst2: hsstate) ctx: Prop :=
- forall st1 st2,
- hsstate_refines hst1 st1 ->
- hsstate_refines hst2 st2 -> sstate_simu dm f outframe st1 st2 ctx.
-
-Theorem hsstate_simu_spec_correct dm f outframe ctx hst1 hst2:
- hsstate_simu_spec dm f outframe hst1 hst2 ->
- hsstate_simu dm f outframe hst1 hst2 ctx.
-Proof.
- intros (SCORE & FSIMU) st1 st2 (SREF1 & DREF1 & FREF1) (SREF2 & DREF2 & FREF2).
- generalize SCORE. intro SIMU; eapply hsistate_simu_spec_correct in SIMU; eauto.
- constructor; auto.
- intros is1 SEM1 CONT1.
- unfold hsistate_simu in SIMU. exploit SIMU; clear SIMU; eauto.
- unfold istate_simu, ssem_internal in *; intros (is2 & SEM2 & SIMU).
- rewrite! CONT1 in *. destruct SIMU as (CONT2 & _).
- rewrite! CONT1, <- CONT2 in *.
- destruct SEM1 as (SEM1 & _ & _).
- destruct SEM2 as (SEM2 & _ & _).
- eapply hfinal_simu_spec_correct in FSIMU; eauto.
- - destruct SREF1 as (PC1 & _). destruct SREF2 as (PC2 & _). rewrite <- PC1. rewrite <- PC2.
- eapply FSIMU.
- - eapply FREF1. exploit DREF1. intros (_ & (OK & _) & _). rewrite <- OK. eapply ssem_local_sok; eauto.
- - eapply FREF2. exploit DREF2. intros (_ & (OK & _) & _). rewrite <- OK. eapply ssem_local_sok; eauto.
-Qed.
diff --git a/scheduling/RTLpathSE_theory.v b/scheduling/RTLpathSE_theory.v
deleted file mode 100644
index 2a791feb..00000000
--- a/scheduling/RTLpathSE_theory.v
+++ /dev/null
@@ -1,1876 +0,0 @@
-(* A theory of symbolic execution on RTLpath
-
-NB: an efficient implementation with hash-consing will be defined in RTLpathSE_impl.v
-
-*)
-
-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.
-
-Local Open Scope error_monad_scope.
-
-(* Enhanced from kvx/Asmblockgenproof.v *)
-Ltac explore_hyp :=
- 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
- | [ 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. *)
-
-(** * Syntax and semantics of symbolic values *)
-
-(* symbolic value *)
-Inductive sval :=
- | Sinput (r: reg)
- | Sop (op:operation) (lsv: list_sval) (sm: smem)
- | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval)
-with list_sval :=
- | Snil
- | Scons (sv: sval) (lsv: list_sval)
-(* symbolic memory *)
-with smem :=
- | Sinit
- | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval).
-
-Scheme sval_mut := Induction for sval Sort Prop
-with list_sval_mut := Induction for list_sval Sort Prop
-with smem_mut := Induction for smem Sort Prop.
-
-Fixpoint list_sval_inj (l: list sval): list_sval :=
- match l with
- | nil => Snil
- | v::l => Scons v (list_sval_inj l)
- end.
-
-Local Open Scope option_monad_scope.
-
-Fixpoint 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 <- 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 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 Vundef
- | Some a =>
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- match (Mem.loadv chunk m a) with
- | None => Some Vundef
- | 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
- | Snil => Some nil
- | Scons sv lsv' =>
- 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 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 <- 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
- 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 *)
-(* [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 }.
-
-(* 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).
-
-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 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
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- eval_condition cond args m.
-
-Definition all_fallthrough ge sp (lx: list sistate_exit) rs0 m0: Prop :=
- forall ext, List.In ext lx ->
- seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some false.
-
-Lemma all_fallthrough_revcons ge sp ext rs m lx:
- all_fallthrough ge sp (ext::lx) rs m ->
- seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false
- /\ all_fallthrough ge sp lx rs m.
-Proof.
- intros ALLFU. constructor.
- - assert (In ext (ext::lx)) by (constructor; auto). apply ALLFU in H. assumption.
- - intros ext' INEXT. assert (In ext' (ext::lx)) by (apply in_cons; auto).
- apply ALLFU in H. assumption.
-Qed.
-
-(** Semantic of an exit in pseudo code:
- if si_cond (si_condargs)
- si_elocal; goto if_so
- else ()
-*)
-
-Definition ssem_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop :=
- seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true
- /\ ssem_local ge sp (si_elocal ext) rs m rs' m'
- /\ (si_ifso ext) = pc'.
-
-(* Either an abort on the condition evaluation OR an abort on the sistate_local IF the condition was true *)
-Definition sabort_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) : Prop :=
- let sev_cond := seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m in
- sev_cond = None
- \/ (sev_cond = Some true /\ sabort_local ge sp ext.(si_elocal) rs m).
-
-(** * Syntax and Semantics of symbolic internal state *)
-Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }.
-
-Definition all_fallthrough_upto_exit ge sp ext lx' lx rs m : Prop :=
- is_tail (ext::lx') lx /\ all_fallthrough ge sp lx' rs m.
-
-(** Semantic of a sistate in pseudo code:
- si_exit1; si_exit2; ...; si_exitn;
- si_local; goto si_pc *)
-
-(* Note: in RTLpath, is.(icontinue) = false iff we took an early exit *)
-
-Definition ssem_internal (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop :=
- if (is.(icontinue))
- then
- ssem_local ge sp st.(si_local) rs m is.(irs) is.(imem)
- /\ st.(si_pc) = is.(ipc)
- /\ all_fallthrough ge sp st.(si_exits) rs m
- else exists ext lx,
- ssem_exit ge sp ext rs m is.(irs) is.(imem) is.(ipc)
- /\ all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m.
-
-Definition sabort (ge: RTL.genv) (sp: val) (st: sistate) (rs: regset) (m: mem): Prop :=
- (* No early exit was met but we aborted on the si_local *)
- (all_fallthrough ge sp st.(si_exits) rs m /\ sabort_local ge sp st.(si_local) rs m)
- (* OR we aborted on an evaluation of one of the early exits *)
- \/ (exists ext lx, all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m /\ sabort_exit ge sp ext rs m).
-
-Definition ssem_internal_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop :=
- match ois with
- | Some is => ssem_internal ge sp st rs0 m0 is
- | None => sabort ge sp st rs0 m0
- end.
-
-Definition ssem_internal_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop :=
- match ost with
- | Some st => ssem_internal_opt ge sp st rs0 m0 ois
- | None => ois=None
- end.
-
-(** * An internal state represents a parallel program !
-
- We prove below that the semantics [ssem_internal_opt] is deterministic.
-
- *)
-
-Definition istate_eq ist1 ist2 :=
- ist1.(icontinue) = ist2.(icontinue) /\
- ist1.(ipc) = ist2.(ipc) /\
- (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\
- ist1.(imem) = ist2.(imem).
-
-Lemma all_fallthrough_noexit ge sp ext lx rs0 m0 rs m pc:
- ssem_exit ge sp ext rs0 m0 rs m pc ->
- In ext lx ->
- all_fallthrough ge sp lx rs0 m0 ->
- False.
-Proof.
- Local Hint Resolve is_tail_in: core.
- intros SSEM INE ALLF.
- destruct SSEM as (SSEM & SSEM').
- unfold all_fallthrough in ALLF. rewrite ALLF in SSEM; eauto.
- discriminate.
-Qed.
-
-Lemma ssem_internal_exclude_incompatible_continue ge sp st rs m is1 is2:
- is1.(icontinue) = true ->
- is2.(icontinue) = false ->
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- False.
-Proof.
- Local Hint Resolve all_fallthrough_noexit: core.
- unfold ssem_internal.
- intros CONT1 CONT2.
- rewrite CONT1, CONT2; simpl.
- intuition eauto.
- destruct H0 as (ext & lx & SSEME & ALLFU).
- destruct ALLFU as (ALLFU & ALLFU').
- eapply all_fallthrough_noexit; eauto.
-Qed.
-
-Lemma ssem_internal_determ_continue ge sp st rs m is1 is2:
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- is1.(icontinue) = is2.(icontinue).
-Proof.
- Local Hint Resolve ssem_internal_exclude_incompatible_continue: core.
- destruct (Bool.bool_dec is1.(icontinue) is2.(icontinue)) as [|H]; auto.
- intros H1 H2. assert (absurd: False); intuition.
- destruct (icontinue is1) eqn: His1, (icontinue is2) eqn: His2; eauto.
-Qed.
-
-Lemma ssem_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2:
- ssem_local ge sp st rs0 m0 rs1 m1 ->
- ssem_local ge sp st rs0 m0 rs2 m2 ->
- (forall r, rs1#r = rs2#r) /\ m1 = m2.
-Proof.
- unfold ssem_local. intuition try congruence.
- generalize (H5 r); rewrite H4; congruence.
-Qed.
-
-(* TODO: lemma to move in Coqlib *)
-Lemma is_tail_bounded_total {A} (l1 l2 l3: list A): is_tail l1 l3 -> is_tail l2 l3
- -> is_tail l1 l2 \/ is_tail l2 l1.
-Proof.
- Local Hint Resolve is_tail_cons: core.
- induction 1 as [|i l1 l3 T1 IND]; simpl; auto.
- intros T2; inversion T2; subst; auto.
-Qed.
-
-Lemma exit_cond_determ ge sp rs0 m0 l1 l2:
- is_tail l1 l2 -> forall ext1 lx1 ext2 lx2,
- l1=(ext1 :: lx1) ->
- l2=(ext2 :: lx2) ->
- all_fallthrough ge sp lx1 rs0 m0 ->
- seval_condition ge sp (si_cond ext1) (si_scondargs ext1) (si_smem (si_elocal ext1)) rs0 m0 = Some true ->
- all_fallthrough ge sp lx2 rs0 m0 ->
- ext1=ext2.
-Proof.
- destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst;
- inversion EQ2; subst; auto.
- intros D1 EVAL NYE.
- Local Hint Resolve is_tail_in: core.
- unfold all_fallthrough in NYE.
- rewrite NYE in EVAL; eauto.
- try congruence.
-Qed.
-
-Lemma ssem_exit_determ ge sp ext rs0 m0 rs1 m1 pc1 rs2 m2 pc2:
- ssem_exit ge sp ext rs0 m0 rs1 m1 pc1 ->
- ssem_exit ge sp ext rs0 m0 rs2 m2 pc2 ->
- pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2.
-Proof.
- Local Hint Resolve exit_cond_determ eq_sym: core.
- intros SSEM1 SSEM2. destruct SSEM1 as (SEVAL1 & SLOC1 & PCEQ1). destruct SSEM2 as (SEVAL2 & SLOC2 & PCEQ2). subst.
- destruct (ssem_local_determ ge sp (si_elocal ext) rs0 m0 rs1 m1 rs2 m2); auto.
-Qed.
-
-Remark is_tail_inv_left {A: Type} (a a': A) l l':
- is_tail (a::l) (a'::l') ->
- (a = a' /\ l = l') \/ (In a l' /\ is_tail l (a'::l')).
-Proof.
- intros. inv H.
- - left. eauto.
- - right. econstructor.
- + eapply is_tail_in; eauto.
- + eapply is_tail_cons_left; eauto.
-Qed.
-
-Lemma ssem_internal_determ ge sp st rs m is1 is2:
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- istate_eq is1 is2.
-Proof.
- unfold istate_eq.
- intros SEM1 SEM2.
- exploit (ssem_internal_determ_continue ge sp st rs m is1 is2); eauto.
- intros CONTEQ. unfold ssem_internal in * |-. rewrite CONTEQ in * |- *.
- destruct (icontinue is2).
- - destruct (ssem_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2));
- intuition (try congruence).
- - destruct SEM1 as (ext1 & lx1 & SSEME1 & ALLFU1). destruct SEM2 as (ext2 & lx2 & SSEME2 & ALLFU2).
- destruct ALLFU1 as (ALLFU1 & ALLFU1'). destruct ALLFU2 as (ALLFU2 & ALLFU2').
- destruct SSEME1 as (SSEME1 & SSEME1' & SSEME1''). destruct SSEME2 as (SSEME2 & SSEME2' & SSEME2'').
- assert (X:ext1=ext2).
- { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2 :: lx2) (si_exits st)) as [TAIL|TAIL]; eauto. }
- subst. destruct (ssem_local_determ ge sp (si_elocal ext2) rs m (irs is1) (imem is1) (irs is2) (imem is2)); auto.
- intuition. congruence.
-Qed.
-
-Lemma ssem_local_exclude_sabort_local ge sp loc rs m rs' m':
- ssem_local ge sp loc rs m rs' m' ->
- sabort_local ge sp loc rs m ->
- False.
-Proof.
- intros SIML ABORT. inv SIML. destruct H0 as (H0 & H0').
- inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence.
-Qed.
-
-Lemma ssem_local_exclude_sabort ge sp st rs m rs' m':
- ssem_local ge sp (si_local st) rs m rs' m' ->
- all_fallthrough ge sp (si_exits st) rs m ->
- sabort ge sp st rs m ->
- False.
-Proof.
- intros SIML ALLF ABORT.
- inv ABORT.
- - intuition; eapply ssem_local_exclude_sabort_local; eauto.
- - destruct H as (ext & lx & ALLFU & SABORT).
- destruct ALLFU as (TAIL & _). eapply is_tail_in in TAIL.
- eapply ALLF in TAIL.
- destruct SABORT as [CONDFAIL | (CONDTRUE & ABORTL)]; congruence.
-Qed.
-
-Lemma ssem_exit_fallthrough_upto_exit ge sp ext ext' lx lx' exits rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- all_fallthrough_upto_exit ge sp ext lx exits rs m ->
- all_fallthrough_upto_exit ge sp ext' lx' exits rs m ->
- is_tail (ext'::lx') (ext::lx).
-Proof.
- intros SSEME ALLFU ALLFU'.
- destruct ALLFU as (ISTAIL & ALLFU). destruct ALLFU' as (ISTAIL' & ALLFU').
- destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto.
- inv H.
- - econstructor; eauto.
- - eapply is_tail_in in H2. eapply ALLFU' in H2.
- destruct SSEME as (SEVAL & _). congruence.
-Qed.
-
-Lemma ssem_exit_exclude_sabort_exit ge sp ext rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- sabort_exit ge sp ext rs m ->
- False.
-Proof.
- intros A B. destruct A as (A & A' & A''). inv B.
- - congruence.
- - destruct H as (_ & H). eapply ssem_local_exclude_sabort_local; eauto.
-Qed.
-
-Lemma ssem_exit_exclude_sabort ge sp ext st lx rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m ->
- sabort ge sp st rs m ->
- False.
-Proof.
- intros SSEM ALLFU ABORT.
- inv ABORT.
- - destruct H as (ALLF & _). destruct ALLFU as (TAIL & _).
- eapply is_tail_in in TAIL.
- destruct SSEM as (SEVAL & _ & _).
- eapply ALLF in TAIL. congruence.
- - destruct H as (ext' & lx' & ALLFU' & ABORT).
- exploit ssem_exit_fallthrough_upto_exit; eauto. intros ITAIL.
- destruct ALLFU as (ALLFU1 & ALLFU2). destruct ALLFU' as (ALLFU1' & ALLFU2').
- exploit (is_tail_inv_left ext' ext lx' lx); eauto. intro. inv H.
- + inv H0. eapply ssem_exit_exclude_sabort_exit; eauto.
- + destruct H0 as (INE & TAIL). eapply ALLFU2 in INE. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence.
-Qed.
-
-Lemma ssem_internal_exclude_sabort ge sp st rs m is:
- sabort ge sp st rs m ->
- ssem_internal ge sp st rs m is -> False.
-Proof.
- intros ABORT SEM.
- unfold ssem_internal in SEM. destruct icontinue.
- - destruct SEM as (SEM1 & SEM2 & SEM3).
- eapply ssem_local_exclude_sabort; eauto.
- - destruct SEM as (ext & lx & SEM1 & SEM2). eapply ssem_exit_exclude_sabort; eauto.
-Qed.
-
-Definition istate_eq_opt ist1 oist :=
- exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2.
-
-Lemma ssem_internal_opt_determ ge sp st rs m ois is:
- ssem_internal_opt ge sp st rs m ois ->
- ssem_internal ge sp st rs m is ->
- istate_eq_opt is ois.
-Proof.
- destruct ois as [is1|]; simpl; eauto.
- - intros; eexists; intuition; eapply ssem_internal_determ; eauto.
- - intros; exploit ssem_internal_exclude_sabort; eauto. destruct 1.
-Qed.
-
-(** * Symbolic execution of one internal step *)
-
-Definition slocal_set_sreg (st:sistate_local) (r:reg) (sv:sval) :=
- {| si_pre:=(fun ge sp rs m => 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 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_set_local (st: sistate) (pc: node) (nxt: sistate_local): sistate :=
- {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}.
-
-Definition slocal_store st chunk addr args src : sistate_local :=
- let args := list_sval_inj (List.map (si_sreg st) args) in
- let src := si_sreg st src in
- let sm := Sstore (si_smem st) chunk addr args src
- in slocal_set_smem st sm.
-
-Definition siexec_inst (i: instruction) (st: sistate): option sistate :=
- match i with
- | Inop pc' =>
- Some (sist_set_local st pc' st.(si_local))
- | Iop op args dst pc' =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let next := slocal_set_sreg prev dst (Sop op vargs prev.(si_smem)) in
- Some (sist_set_local st pc' next)
- | Iload trap chunk addr args dst pc' =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let next := slocal_set_sreg prev dst (Sload prev.(si_smem) trap chunk addr vargs) in
- Some (sist_set_local st pc' next)
- | Istore chunk addr args src pc' =>
- let next := slocal_store st.(si_local) chunk addr args src in
- Some (sist_set_local st pc' next)
- | Icond cond args ifso ifnot _ =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let ex := {| si_cond:=cond; si_scondargs:=vargs; si_elocal := prev; si_ifso := ifso |} in
- Some {| si_pc := ifnot; si_exits := ex::st.(si_exits); si_local := prev |}
- | _ => None
- end.
-
-Lemma 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.
- inversion_SOME lv.
- generalize (H r).
- try_simplify_someHyps.
-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.
-Proof.
- unfold sabort_local. simpl; intuition.
- destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST.
- - subst; rewrite H; intuition.
- - right. right. exists r1. rewrite HTEST. auto.
-Qed.
-
-Lemma slocal_set_smem_preserves_sabort_local ge sp st rs0 m0 m:
- sabort_local ge sp st rs0 m0 ->
- sabort_local ge sp (slocal_set_smem st m) rs0 m0.
-Proof.
- unfold sabort_local. simpl; intuition.
-Qed.
-
-Lemma all_fallthrough_upto_exit_cons ge sp ext lx ext' exits rs m:
- all_fallthrough_upto_exit ge sp ext lx exits rs m ->
- all_fallthrough_upto_exit ge sp ext lx (ext'::exits) rs m.
-Proof.
- intros. inv H. econstructor; eauto.
-Qed.
-
-Lemma all_fallthrough_cons ge sp exits rs m ext:
- all_fallthrough ge sp exits rs m ->
- seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false ->
- all_fallthrough ge sp (ext::exits) rs m.
-Proof.
- intros. unfold all_fallthrough in *. intros.
- inv H1; eauto.
-Qed.
-
-Lemma siexec_inst_preserves_sabort i ge sp rs m st st':
- siexec_inst i st = Some st' ->
- sabort ge sp st rs m -> sabort ge sp st' rs m.
-Proof.
- intros SISTEP ABORT.
- destruct i; simpl in SISTEP; try discriminate; inv SISTEP; unfold sabort; simpl.
- (* NOP *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* OP *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* LOAD *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* STORE *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_smem_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* COND *)
- * remember ({| si_cond := _; si_scondargs := _; si_elocal := _; si_ifso := _ |}) as ext.
- destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - destruct (seval_condition ge sp (si_cond ext) (si_scondargs ext)
- (si_smem (si_elocal ext)) rs m) eqn:SEVAL; [destruct b|].
- (* case true *)
- + right. exists ext, (si_exits st).
- constructor.
- ++ constructor. econstructor; eauto. eauto.
- ++ unfold sabort_exit. right. constructor; eauto.
- subst. simpl. eauto.
- (* case false *)
- + left. constructor; eauto. eapply all_fallthrough_cons; eauto.
- (* case None *)
- + right. exists ext, (si_exits st). constructor.
- ++ constructor. econstructor; eauto. eauto.
- ++ unfold sabort_exit. left. eauto.
- - right. exists ext0, lx0. constructor; eauto. eapply all_fallthrough_upto_exit_cons; eauto.
-Qed.
-
-Lemma siexec_inst_WF i st:
- siexec_inst i st = None -> default_succ i = None.
-Proof.
- destruct i; simpl; unfold sist_set_local; simpl; congruence.
-Qed.
-
-Lemma siexec_inst_default_succ i st st':
- siexec_inst i st = Some st' -> default_succ i = Some (st'.(si_pc)).
-Proof.
- destruct i; simpl; unfold sist_set_local; simpl; try congruence;
- intro H; inversion_clear H; simpl; auto.
-Qed.
-
-
-Lemma seval_list_sval_inj_not_none ge sp st rs0 m0: forall l,
- (forall r, List.In r l -> seval_sval ge sp (si_sreg st r) rs0 m0 = None -> False) ->
- seval_list_sval ge sp (list_sval_inj (map (si_sreg st) l)) rs0 m0 = None -> False.
-Proof.
- induction l.
- - intuition discriminate.
- - intros ALLR. simpl.
- inversion_SOME v.
- + intro SVAL. inversion_SOME lv; [discriminate|].
- assert (forall r : reg, In r l -> seval_sval ge sp (si_sreg st r) rs0 m0 = None -> False).
- { intros r INR. eapply ALLR. right. assumption. }
- intro SVALLIST. intro. eapply IHl; eauto.
- + intros. exploit (ALLR a); simpl; eauto.
-Qed.
-
-Lemma siexec_inst_correct ge sp i st rs0 m0 rs m:
- ssem_local ge sp st.(si_local) rs0 m0 rs m ->
- all_fallthrough ge sp st.(si_exits) rs0 m0 ->
- ssem_internal_opt2 ge sp (siexec_inst i st) rs0 m0 (istep ge i sp rs m).
-Proof.
- intros (PRE & MEM & REG) NYE.
- destruct i; simpl; auto.
- + (* Nop *)
- constructor; [|constructor]; simpl; auto.
- constructor; auto.
- + (* Op *)
- inversion_SOME v; intros OP; simpl.
- - constructor; [|constructor]; simpl; auto.
- constructor; simpl; auto.
- * constructor; auto. congruence.
- * constructor; auto.
- intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst. rewrite Regmap.gss; simpl; auto.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- - left. constructor; simpl; auto.
- unfold sabort_local. right. right.
- simpl. exists r. destruct (Pos.eq_dec r r); try congruence.
- simpl. erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- + (* LOAD *)
- inversion_SOME a0; intro ADD.
- { inversion_SOME v; intros LOAD; simpl.
- - explore_destruct; unfold ssem_internal, ssem_local; simpl; intuition.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- inversion_SOME args; intros ARGS.
- 2: { exploit seval_list_sval_inj_not_none; eauto; intuition congruence. }
- 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.
- simpl. erewrite seval_list_sval_inj; simpl; auto.
- rewrite ADD; simpl; auto. try_simplify_someHyps.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- } { rewrite ADD. destruct t.
- - simpl. left; eauto. simpl. econstructor; eauto.
- right. right. simpl. exists r. destruct (Pos.eq_dec r r); [|contradiction].
- simpl. inversion_SOME args. intro SLS.
- eapply 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.
- }
- + (* STORE *)
- inversion_SOME a0; intros ADD.
- { inversion_SOME m'; intros STORE; simpl.
- - unfold ssem_internal, ssem_local; simpl; intuition.
- * congruence.
- * erewrite seval_list_sval_inj; simpl; auto.
- erewrite REG.
- try_simplify_someHyps.
- - unfold sabort, sabort_local; simpl.
- left. constructor; auto. right. left.
- erewrite seval_list_sval_inj; simpl; auto.
- erewrite REG.
- try_simplify_someHyps. }
- { unfold sabort, sabort_local; simpl.
- left. constructor; auto. right. left.
- erewrite seval_list_sval_inj; simpl; auto.
- erewrite ADD; simpl; auto. }
- + (* COND *)
- Local Hint Resolve is_tail_refl: core.
- Local Hint Unfold ssem_local: core.
- inversion_SOME b; intros COND.
- { destruct b; simpl; unfold ssem_internal, ssem_local; simpl.
- - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st).
- constructor; constructor; subst; simpl; auto.
- unfold seval_condition. subst; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- - intuition. unfold all_fallthrough in * |- *. simpl.
- intuition. subst. 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. }
-Qed.
-
-
-Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m:
- ssem_local ge sp (st.(si_local)) rs0 m0 rs m ->
- siexec_inst i st = None ->
- istep ge i sp rs m = None.
-Proof.
- intros (PRE & MEM & REG).
- destruct i; simpl; unfold sist_set_local, ssem_internal, ssem_local; simpl; try_simplify_someHyps.
-Qed.
-
-(** * Symbolic execution of the internal steps of a path *)
-Fixpoint siexec_path (path:nat) (f: function) (st: sistate): option sistate :=
- match path with
- | O => Some st
- | S p =>
- SOME i <- (fn_code f)!(st.(si_pc)) IN
- SOME st1 <- siexec_inst i st IN
- siexec_path p f st1
- end.
-
-Lemma siexec_inst_add_exits i st st':
- siexec_inst i st = Some st' ->
- ( si_exits st' = si_exits st \/ exists ext, si_exits st' = ext :: si_exits st ).
-Proof.
- destruct i; simpl; intro SISTEP; inversion_clear SISTEP; unfold siexec_inst; simpl; (discriminate || eauto).
-Qed.
-
-Lemma siexec_inst_preserves_allfu ge sp ext lx rs0 m0 st st' i:
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs0 m0 ->
- siexec_inst i st = Some st' ->
- all_fallthrough_upto_exit ge sp ext lx (si_exits st') rs0 m0.
-Proof.
- intros ALLFU SISTEP. destruct ALLFU as (ISTAIL & ALLF).
- constructor; eauto.
- destruct i; simpl in SISTEP; inversion_clear SISTEP; simpl; (discriminate || eauto).
-Qed.
-
-Lemma siexec_path_correct_false ge sp f rs0 m0 st' is:
- forall path,
- is.(icontinue)=false ->
- forall st, ssem_internal ge sp st rs0 m0 is ->
- siexec_path path f st = Some st' ->
- ssem_internal ge sp st' rs0 m0 is.
-Proof.
- induction path; simpl.
- - intros. congruence.
- - intros ICF st SSEM STEQ'.
- destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate].
- destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate].
- eapply IHpath. 3: eapply STEQ'. eauto.
- unfold ssem_internal in SSEM. rewrite ICF in SSEM.
- destruct SSEM as (ext & lx & SEXIT & ALLFU).
- unfold ssem_internal. rewrite ICF. exists ext, lx.
- constructor; auto. eapply siexec_inst_preserves_allfu; eauto.
-Qed.
-
-Lemma siexec_path_preserves_sabort ge sp path f rs0 m0 st': forall st,
- siexec_path path f st = Some st' ->
- sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0.
-Proof.
- Local Hint Resolve siexec_inst_preserves_sabort: core.
- induction path; simpl.
- + unfold sist_set_local; try_simplify_someHyps.
- + intros st; inversion_SOME i.
- inversion_SOME st1; eauto.
-Qed.
-
-Lemma siexec_path_WF path f: forall st,
- siexec_path path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None.
-Proof.
- induction path; simpl.
- + unfold sist_set_local. intuition congruence.
- + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try tauto.
- destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl.
- - intros; erewrite siexec_inst_default_succ; eauto.
- - intros; erewrite siexec_inst_WF; eauto.
-Qed.
-
-Lemma siexec_path_default_succ path f st': forall st,
- siexec_path path f st = Some st' -> nth_default_succ (fn_code f) path st.(si_pc) = Some st'.(si_pc).
-Proof.
- induction path; simpl.
- + unfold sist_set_local. intros st H. inversion_clear H; simpl; try congruence.
- + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try congruence.
- destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl; try congruence.
- intros; erewrite siexec_inst_default_succ; eauto.
-Qed.
-
-Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is,
- is.(icontinue)=true ->
- ssem_internal ge sp st rs0 m0 is ->
- nth_default_succ (fn_code f) path st.(si_pc) <> None ->
- ssem_internal_opt2 ge sp (siexec_path path f st) rs0 m0
- (isteps ge path f sp is.(irs) is.(imem) is.(ipc))
- .
-Proof.
- Local Hint Resolve siexec_path_correct_false siexec_path_preserves_sabort siexec_path_WF: core.
- induction path; simpl.
- + intros st is CONT INV WF;
- unfold ssem_internal, sist_set_local in * |- *;
- try_simplify_someHyps. simpl.
- destruct is; simpl in * |- *; subst; intuition auto.
- + intros st is CONT; unfold ssem_internal at 1; rewrite CONT.
- intros (LOCAL & PC & NYE) WF.
- rewrite <- PC.
- inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto.
- exploit siexec_inst_correct; eauto.
- inversion_SOME st1; intros Hst1; erewrite Hst1; simpl.
- - inversion_SOME is1; intros His1;rewrite His1; simpl.
- * destruct (icontinue is1) eqn:CONT1.
- (* icontinue is0 = true *)
- intros; eapply IHpath; eauto.
- destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps.
- (* icontinue is0 = false -> EARLY EXIT *)
- destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto.
- destruct WF. erewrite siexec_inst_default_succ; eauto.
- (* try_simplify_someHyps; eauto. *)
- * destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto.
- - intros His1;rewrite His1; simpl; auto.
-Qed.
-
-(** REM: in the following two unused lemmas *)
-
-Lemma siexec_path_right_assoc_decompose f path: forall st st',
- siexec_path (S path) f st = Some st' ->
- exists st0, siexec_path path f st = Some st0 /\ siexec_path 1%nat f st0 = Some st'.
-Proof.
- induction path; simpl; eauto.
- intros st st'.
- inversion_SOME i1.
- inversion_SOME st1.
- try_simplify_someHyps; eauto.
-Qed.
-
-Lemma siexec_path_right_assoc_compose f path: forall st st0 st',
- siexec_path path f st = Some st0 ->
- siexec_path 1%nat f st0 = Some st' ->
- siexec_path (S path) f st = Some st'.
-Proof.
- induction path.
- + intros st st0 st' H. simpl in H.
- try_simplify_someHyps; auto.
- + intros st st0 st'.
- assert (X:exists x, x=(S path)); eauto.
- destruct X as [x X].
- intros H1 H2. rewrite <- X.
- generalize H1; clear H1. simpl.
- inversion_SOME i1. intros Hi1; rewrite Hi1.
- inversion_SOME st1. intros Hst1; rewrite Hst1.
- subst; eauto.
-Qed.
-
-(** * Symbolic (final) value of a path *)
-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 (ef:external_function) (sargs: list (builtin_arg sval)) (res: builtin_res reg) (pc:node)
- | Sjumptable (sv: sval) (tbl: list node)
- | Sreturn: option sval -> sfval
-.
-
-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_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.
-
-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 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 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 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 ->
- 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 ->
- 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 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 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 npc stack f rs0 m0 (Sreturn osv) rs m
- E0 (Returnstate stack v m')
-.
-
-Record sstate := { internal:> sistate; final: sfval }.
-
-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 ->
- 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 ->
- 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
- .
-
-(* NB: generic function that could be put into [AST] file *)
-Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B :=
- match arg with
- | BA x => BA (f x)
- | BA_int n => BA_int n
- | BA_long n => BA_long n
- | BA_float f => BA_float f
- | BA_single s => BA_single s
- | BA_loadstack chunk ptr => BA_loadstack chunk ptr
- | BA_addrstack ptr => BA_addrstack ptr
- | BA_loadglobal chunk id ptr => BA_loadglobal chunk id ptr
- | BA_addrglobal id ptr => BA_addrglobal id ptr
- | BA_splitlong ba1 ba2 => BA_splitlong (builtin_arg_map f ba1) (builtin_arg_map f ba2)
- | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2)
- end.
-
-Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg: forall arg varg,
- (forall r, 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_map sreg arg) varg.
-Proof.
- induction arg.
- all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence).
- - intros varg SEVAL BARG. inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
- - intros varg SEVAL BARG. inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
-Qed.
-
-Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs:
- (forall r, 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_map sreg) args) vargs.
-Proof.
- induction 2.
- - constructor.
- - simpl. constructor; [| assumption].
- eapply seval_builtin_arg_correct; eauto.
-Qed.
-
-Lemma seval_builtin_arg_complete ge sp rs m rs0 m0 sreg: forall arg varg,
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg ->
- eval_builtin_arg ge (fun r => rs # r) sp m arg varg.
-Proof.
- induction arg.
- all: intros varg SEVAL BARG; try (inv BARG; constructor; congruence).
- - inv BARG. rewrite SEVAL in H0. inv H0. constructor.
- - inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
- - inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
-Qed.
-
-Lemma seval_builtin_args_complete ge sp rs m rs0 m0 sreg: forall args vargs,
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs ->
- eval_builtin_args ge (fun r => rs # r) sp m args vargs.
-Proof.
- induction args.
- - simpl. intros. inv H0. constructor.
- - intros vargs SEVAL BARG. simpl in BARG. inv BARG.
- constructor; [| eapply IHargs; eauto].
- eapply seval_builtin_arg_complete; eauto.
-Qed.
-
-(** * Symbolic execution of final step *)
-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
- 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
- | Ibuiltin ef args res pc =>
- 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
- Sreturn sor
- | Ijumptable reg tbl =>
- let sv := prev.(si_sreg) reg in
- Sjumptable sv tbl
- | _ => Snone
- end.
-
-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 ->
- 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&REG) LAST. destruct LAST; subst; try_simplify_someHyps; simpl.
- + (* Snone *) intro Hi; 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 seval_list_sval_inj; simpl; auto.
- + (* Itailcall *) intros. eapply exec_Stailcall; auto.
- - destruct ros; simpl in * |- *; auto.
- rewrite REG; auto.
- - erewrite seval_list_sval_inj; simpl; auto.
- + (* Ibuiltin *) intros. eapply exec_Sbuiltin; eauto.
- eapply seval_builtin_args_correct; eauto.
- + (* Ijumptable *) intros. eapply exec_Sjumptable; eauto. congruence.
- + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto.
- destruct or; simpl; auto.
-Qed.
-
-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 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.
- intros PC1 PC2 (PRE&MEM&REG) 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.
- + (* 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 *) intros HPC SMEM.
- eapply exec_Ibuiltin; eauto.
- eapply seval_builtin_args_complete; eauto.
- + (* 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.
-Qed.
-
-(** * Main function of the symbolic execution *)
-
-Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}.
-
-Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}.
-
-Lemma init_ssem_internal ge sp pc rs m: ssem_internal ge sp (init_sistate pc) rs m (mk_istate true pc rs m).
-Proof.
- unfold ssem_internal, ssem_local, all_fallthrough; simpl. intuition.
-Qed.
-
-Definition sexec (f: function) (pc:node): option sstate :=
- SOME path <- (fn_path f)!pc IN
- SOME st <- siexec_path path.(psize) f (init_sistate pc) IN
- SOME i <- (fn_code f)!(st.(si_pc)) IN
- Some (match siexec_inst i st with
- | Some st' => {| internal := st'; final := Snone |}
- | None => {| internal := st; final := sexec_final i st.(si_local) |}
- end).
-
-Lemma final_node_path_simpl f path pc:
- (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(psize) pc <> None.
-Proof.
- intros; exploit final_node_path; eauto.
- intros (i & NTH & DUM).
- congruence.
-Qed.
-
-Lemma symb_path_last_step i st st' ge pge stack (f:function) sp pc rs m t s:
- (fn_code f) ! pc = Some i ->
- pc = st.(si_pc) ->
- siexec_inst i st = Some st' ->
- path_last_step ge pge stack f sp pc rs m t s ->
- exists mk_istate,
- istep ge i sp rs m = Some mk_istate
- /\ t = E0
- /\ s = (State stack f sp mk_istate.(ipc) mk_istate.(RTLpath.irs) mk_istate.(imem)).
-Proof.
- intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl.
-Qed.
-
-(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec])
-(sexec is a correct over-approximation)
-*)
-Theorem sexec_correct f pc pge ge sp path stack rs m t s:
- (fn_path f)!pc = Some path ->
- path_step ge pge path.(psize) stack f sp rs m pc t s ->
- exists st, sexec f pc = Some st /\ ssem pge ge sp st stack f rs m t s.
-Proof.
- Local Hint Resolve init_ssem_internal: core.
- intros PATH STEP; unfold sexec; rewrite PATH; simpl.
- lapply (final_node_path_simpl f path pc); eauto. intro WF.
- exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto.
- { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. }
- (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]).
- destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST];
- (* intro Hst *)
- (rewrite STEPS; unfold ssem_internal_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence);
- (* intro SEM *)
- (simpl; unfold ssem_internal; simpl; rewrite CONT; intro SEM);
- (* intro Hi' *)
- ( assert (Hi': (fn_code f) ! (si_pc st) = Some i);
- [ unfold nth_default_succ_inst in Hi;
- exploit siexec_path_default_succ; eauto; simpl;
- intros DEF; rewrite DEF in Hi; auto
- | clear Hi; rewrite Hi' ]);
- (* eexists *)
- (eexists; constructor; eauto).
- - (* early *)
- eapply ssem_early; eauto.
- unfold ssem_internal; simpl; rewrite CONT.
- destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl; eauto.
- destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx.
- constructor; auto. eapply siexec_inst_preserves_allfu; eauto.
- - destruct SEM as (SEM & PC & HNYE).
- destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl.
- + (* normal on Snone *)
- rewrite <- PC in LAST.
- exploit symb_path_last_step; eauto; simpl.
- intros (mk_istate & ISTEP & Ht & Hs); subst.
- exploit siexec_inst_correct; eauto. simpl.
- erewrite Hst', ISTEP; simpl.
- clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i.
- intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT.
- { (* icontinue mk_istate = true *)
- eapply ssem_normal; simpl; eauto.
- unfold ssem_internal in SEM.
- rewrite CONT in SEM.
- destruct SEM as (SEM & PC & HNYE).
- rewrite <- PC.
- eapply exec_Snone. }
- { eapply ssem_early; eauto. }
- + (* normal non-Snone instruction *)
- eapply ssem_normal; eauto.
- * unfold ssem_internal; simpl; rewrite CONT; intuition.
- * simpl. eapply sexec_final_correct; eauto.
- rewrite PC; auto.
-Qed.
-
-(* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *)
-Inductive equiv_stackframe: stackframe -> stackframe -> Prop :=
- | equiv_stackframe_intro res f sp pc rs1 rs2
- (EQUIV: forall r : positive, rs1 !! r = rs2 !! r):
- equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2).
-
-Inductive equiv_state: state -> state -> Prop :=
- | State_equiv stack f sp pc rs1 m rs2
- (EQUIV: forall r, rs1#r = rs2#r):
- equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m)
- | Call_equiv stk stk' f args m
- (STACKS: list_forall2 equiv_stackframe stk stk'):
- equiv_state (Callstate stk f args m) (Callstate stk' f args m)
- | Return_equiv stk stk' v m
- (STACKS: list_forall2 equiv_stackframe stk stk'):
- equiv_state (Returnstate stk v m) (Returnstate stk' v m).
-
-Lemma equiv_stackframe_refl stf: equiv_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: core.
- induction stk; simpl; constructor; auto.
-Qed.
-
-Lemma equiv_state_refl s: equiv_state s s.
-Proof.
- Local Hint Resolve equiv_stack_refl: core.
- induction s; simpl; constructor; auto.
-Qed.
-
-(*
-Lemma equiv_stackframe_trans stf1 stf2 stf3:
- equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3.
-Proof.
- destruct 1; intros EQ; inv EQ; try econstructor; eauto.
- intros; eapply eq_trans; eauto.
-Qed.
-
-Lemma equiv_stack_trans stk1 stk2:
- list_forall2 equiv_stackframe stk1 stk2 ->
- forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
- list_forall2 equiv_stackframe stk1 stk3.
-Proof.
- Local Hint Resolve equiv_stackframe_trans.
- induction 1; intros stk3 EQ; inv EQ; econstructor; eauto.
-Qed.
-
-Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3.
-Proof.
- Local Hint Resolve equiv_stack_trans.
- destruct 1; intros EQ; inv EQ; econstructor; eauto.
- intros; eapply eq_trans; eauto.
-Qed.
-*)
-
-Lemma regmap_setres_eq (rs rs': regset) res vres:
- (forall r, rs # r = rs' # r) ->
- forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r.
-Proof.
- intros RSEQ r. destruct res; simpl; try congruence.
- destruct (peq x r).
- - subst. repeat (rewrite Regmap.gss). reflexivity.
- - repeat (rewrite Regmap.gso); auto.
-Qed.
-
-Lemma 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' /\ 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.
- - (* Snone *) intros; eexists; econstructor.
- + eapply State_equiv; eauto.
- + eapply exec_Snone.
- - (* Scall *)
- intros; eexists; econstructor.
- 2: { eapply exec_Scall; eauto. }
- apply Call_equiv; auto.
- repeat (constructor; auto).
- - (* Stailcall *)
- intros; eexists; econstructor; [| eapply exec_Stailcall; eauto].
- apply Call_equiv; auto.
- - (* Sbuiltin *)
- intros; eexists; econstructor; [| eapply exec_Sbuiltin; eauto].
- constructor. eapply regmap_setres_eq; eauto.
- - (* Sjumptable *)
- intros; eexists; econstructor; [| eapply exec_Sjumptable; eauto].
- constructor. assumption.
- - (* Sreturn *)
- intros; eexists; econstructor; [| eapply exec_Sreturn; eauto].
- eapply equiv_state_refl; eauto.
-Qed.
-
-Lemma siexec_inst_early_exit_absurd i st st' ge sp rs m rs' m' pc':
- siexec_inst i st = Some st' ->
- (exists ext lx, ssem_exit ge sp ext rs m rs' m' pc' /\
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m) ->
- all_fallthrough ge sp (si_exits st') rs m ->
- False.
-Proof.
- intros SIEXEC (ext & lx & SSEME & ALLFU) ALLF. destruct ALLFU as (TAIL & _).
- exploit siexec_inst_add_exits; eauto. destruct 1 as [SIEQ | (ext0 & SIEQ)].
- - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. eassumption.
- - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in.
- constructor. eassumption.
-Qed.
-
-Lemma is_tail_false {A: Type}: forall (l: list A) a, is_tail (a::l) nil -> False.
-Proof.
- intros. eapply is_tail_incl in H. unfold incl in H. pose (H a).
- assert (In a (a::l)) by (constructor; auto). assert (In a nil) by auto. apply in_nil in H1.
- contradiction.
-Qed.
-
-Lemma cons_eq_false {A: Type}: forall (l: list A) a,
- a :: l = l -> False.
-Proof.
- induction l; intros.
- - discriminate.
- - inv H. apply IHl in H2. contradiction.
-Qed.
-
-Lemma app_cons_nil_eq {A: Type}: forall l' l (a:A),
- (l' ++ a :: nil) ++ l = l' ++ a::l.
-Proof.
- induction l'; intros.
- - simpl. reflexivity.
- - simpl. rewrite IHl'. reflexivity.
-Qed.
-
-Lemma app_eq_false {A: Type}: forall l (l': list A) a,
- l' ++ a :: l = l -> False.
-Proof.
- induction l; intros.
- - apply app_eq_nil in H. destruct H as (_ & H). apply cons_eq_false in H. contradiction.
- - destruct l' as [|a' l'].
- + simpl in H. apply cons_eq_false in H. contradiction.
- + rewrite <- app_comm_cons in H. inv H.
- apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption.
-Qed.
-
-Lemma is_tail_false_gen {A: Type}: forall (l: list A) l' a, is_tail (l'++(a::l)) l -> False.
-Proof.
- induction l.
- - intros. destruct l' as [|a' l'].
- + simpl in H. apply is_tail_false in H. contradiction.
- + rewrite <- app_comm_cons in H. apply is_tail_false in H. contradiction.
- - intros. inv H.
- + apply app_eq_false in H2. contradiction.
- + apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption.
-Qed.
-
-Lemma is_tail_eq {A: Type}: forall (l l': list A),
- is_tail l' l ->
- is_tail l l' ->
- l = l'.
-Proof.
- destruct l as [|a l]; intros l' ITAIL ITAIL'.
- - destruct l' as [|i' l']; auto. apply is_tail_false in ITAIL. contradiction.
- - inv ITAIL; auto.
- destruct l' as [|i' l']. { apply is_tail_false in ITAIL'. contradiction. }
- exploit is_tail_trans. eapply ITAIL'. eauto. intro ABSURD.
- apply (is_tail_false_gen l nil a) in ABSURD. contradiction.
-Qed.
-
-(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution
- (sexec is exact).
-*)
-Theorem sexec_exact f pc pge ge sp path stack st rs m t s1:
- (fn_path f)!pc = Some path ->
- sexec f pc = Some st ->
- ssem pge ge sp st stack f rs m t s1 ->
- exists s2, path_step ge pge path.(psize) stack f sp rs m pc t s2 /\
- equiv_state s1 s2.
-Proof.
- Local Hint Resolve init_ssem_internal: core.
- unfold sexec; intros PATH SSTEP SEM; rewrite PATH in SSTEP.
- lapply (final_node_path_simpl f path pc); eauto. intro WF.
- exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto.
- { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. }
- (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]).
- unfold nth_default_succ_inst in Hi.
- destruct (siexec_path path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl.
- 2:{ (* absurd case *)
- exploit siexec_path_WF; eauto.
- simpl; intros NDS; rewrite NDS in Hi; congruence. }
- exploit siexec_path_default_succ; eauto; simpl.
- intros NDS; rewrite NDS in Hi.
- rewrite Hi in SSTEP.
- intros ISTEPS. try_simplify_someHyps.
- destruct (siexec_inst i st0) as [st'|] eqn:Hst'; simpl.
- + (* exit on Snone instruction *)
- assert (SEM': t = E0 /\ exists is, ssem_internal ge sp st' rs m is
- /\ s1 = (State stack f sp (if (icontinue is) then (si_pc st') else (ipc is)) (irs is) (imem is))).
- { destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *.
- - repeat (econstructor; eauto).
- rewrite CONT; eauto.
- - inversion SEM2. repeat (econstructor; eauto).
- rewrite CONT; eauto. }
- clear SEM; subst. destruct SEM' as [X (is & SEM & X')]; subst.
- intros.
- destruct (isteps ge (psize path) f sp rs m pc) as [is0|] eqn:RISTEPS; simpl in *.
- * unfold ssem_internal in ISTEPS. destruct (icontinue is0) eqn: ICONT0.
- ** (* icontinue is0=true: path_step by normal_exit *)
- destruct ISTEPS as (SEMis0&H1&H2).
- rewrite H1 in * |-.
- exploit siexec_inst_correct; eauto.
- rewrite Hst'; simpl.
- intros; exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- eexists. econstructor 1.
- *** eapply exec_normal_exit; eauto.
- eapply exec_istate; eauto.
- *** rewrite EQ1.
- enough ((ipc st) = (if icontinue st then si_pc st' else ipc is)) as ->.
- { rewrite EQ2, EQ4. eapply State_equiv; auto. }
- destruct (icontinue st) eqn:ICONT; auto.
- exploit siexec_inst_default_succ; eauto.
- erewrite istep_normal_exit; eauto.
- try_simplify_someHyps.
- ** (* The concrete execution has not reached "i" => early exit *)
- unfold ssem_internal in SEM.
- destruct (icontinue is) eqn:ICONT.
- { destruct SEM as (SEML & SIPC & ALLF).
- exploit siexec_inst_early_exit_absurd; eauto. contradiction. }
-
- eexists. econstructor 1.
- *** eapply exec_early_exit; eauto.
- *** destruct ISTEPS as (ext & lx & SSEME & ALLFU). destruct SEM as (ext' & lx' & SSEME' & ALLFU').
- eapply siexec_inst_preserves_allfu in ALLFU; eauto.
- exploit ssem_exit_fallthrough_upto_exit; eauto.
- exploit ssem_exit_fallthrough_upto_exit. eapply SSEME. eapply ALLFU. eapply ALLFU'.
- intros ITAIL ITAIL'. apply is_tail_eq in ITAIL; auto. clear ITAIL'.
- inv ITAIL. exploit ssem_exit_determ. eapply SSEME. eapply SSEME'. intros (IPCEQ & IRSEQ & IMEMEQ).
- rewrite <- IPCEQ. rewrite <- IMEMEQ. constructor. congruence.
- * (* The concrete execution has not reached "i" => abort case *)
- eapply siexec_inst_preserves_sabort in ISTEPS; eauto.
- exploit ssem_internal_exclude_sabort; eauto. contradiction.
- + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *.
- - (* early exit *)
- intros.
- exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- eexists. econstructor 1.
- * eapply exec_early_exit; eauto.
- * rewrite EQ2, EQ4; eapply State_equiv. auto.
- - (* normal exit non-Snone instruction *)
- intros.
- exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- unfold ssem_internal in SEM1.
- rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0).
- exploit ssem_final_equiv; eauto.
- clear SEM2; destruct 1 as (s' & Ms' & SEM2).
- rewrite ! EQ4 in * |-; clear EQ4.
- rewrite ! EQ2 in * |-; clear EQ2.
- exists s'; intuition.
- eapply exec_normal_exit; eauto.
- eapply sexec_final_complete; eauto.
- * congruence.
- * unfold ssem_local in * |- *.
- destruct SEM1 as (A & B & C). constructor; [|constructor]; eauto.
- intro r. congruence.
- * congruence.
-Qed.
-
-(** * Simulation of RTLpath code w.r.t symbolic execution *)
-
-Section SymbValPreserved.
-
-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.
- Local Hint Resolve symbols_preserved_RTL: core.
- 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 (seval_list_sval _ _ _ _); auto.
- erewrite <- eval_addressing_preserved; eauto.
- destruct (eval_addressing _ sp _ _); auto.
- rewrite IHsv; auto.
- + rewrite IHsv; clear IHsv. destruct (seval_sval _ _ _ _); auto.
- rewrite IHsv0; 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 (seval_smem _ _ _ _); auto.
- rewrite IHsv1; auto.
-Qed.
-
-Lemma seval_builtin_arg_preserved sp m rs0 m0:
- forall bs varg,
- seval_builtin_arg ge sp m rs0 m0 bs varg ->
- seval_builtin_arg ge' sp m rs0 m0 bs varg.
-Proof.
- induction 1.
- all: try (constructor; auto).
- - rewrite <- 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.
- induction lsv; simpl; auto.
- rewrite seval_preserved. destruct (seval_sval _ _ _ _); auto.
- rewrite IHlsv; auto.
-Qed.
-
-Lemma smem_eval_preserved 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 (seval_list_sval _ _ _ _); auto.
- erewrite <- eval_addressing_preserved; eauto.
- destruct (eval_addressing _ sp _ _); auto.
- rewrite IHsm; clear IHsm. destruct (seval_smem _ _ _ _); auto.
- rewrite seval_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 (seval_list_sval _ _ _ _); auto.
- rewrite smem_eval_preserved; auto.
-Qed.
-
-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)
- /\ is1.(imem) = is2.(imem).
-
-Definition istate_simu f (srce: PTree.t node) outframe is1 is2: Prop :=
- if is1.(icontinue) then
- istate_simulive (fun r => Regset.In r outframe) srce is1 is2
- else
- exists path, f.(fn_path)!(is1.(ipc)) = Some path
- /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2
- /\ srce!(is2.(ipc)) = Some is1.(ipc).
-
-Record simu_proof_context {f1: RTLpath.function} := {
- liveness_hyps: liveness_ok_function f1;
- the_ge1: RTL.genv;
- the_ge2: RTL.genv;
- genv_match: forall s, Genv.find_symbol the_ge1 s = Genv.find_symbol the_ge2 s;
- the_sp: val;
- the_rs0: regset;
- the_m0: mem
-}.
-Arguments simu_proof_context: clear implicits.
-
-(* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *)
-Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) outframe (st1 st2: sistate) (ctx: simu_proof_context f): Prop :=
- forall is1, ssem_internal (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) is1 ->
- exists is2, ssem_internal (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) is2
- /\ istate_simu f dm outframe is1 is2.
-
-Inductive svident_simu (f: RTLpath.function) (ctx: simu_proof_context f): (sval + ident) -> (sval + ident) -> Prop :=
- | Sleft_simu sv1 sv2:
- (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.
-
-Fixpoint seval_builtin_sval ge sp bsv rs0 m0 :=
- match bsv with
- | BA sv => SOME v <- seval_sval ge sp sv rs0 m0 IN Some (BA v)
- | BA_splitlong sv1 sv2 =>
- SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN
- SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN
- Some (BA_splitlong v1 v2)
- | BA_addptr sv1 sv2 =>
- SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN
- SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN
- Some (BA_addptr v1 v2)
- | BA_int i => Some (BA_int i)
- | BA_long l => Some (BA_long l)
- | BA_float f => Some (BA_float f)
- | BA_single s => Some (BA_single s)
- | BA_loadstack chk ptr => Some (BA_loadstack chk ptr)
- | BA_addrstack ptr => Some (BA_addrstack ptr)
- | BA_loadglobal chk id ptr => Some (BA_loadglobal chk id ptr)
- | BA_addrglobal id ptr => Some (BA_addrglobal id ptr)
- end.
-
-
-Fixpoint seval_list_builtin_sval ge sp lbsv rs0 m0 :=
- match lbsv with
- | nil => Some nil
- | bsv::lbsv => SOME v <- seval_builtin_sval ge sp bsv rs0 m0 IN
- SOME lv <- seval_list_builtin_sval ge sp lbsv rs0 m0 IN
- Some (v::lv)
- end.
-
-Lemma seval_list_builtin_sval_nil ge sp rs0 m0 lbs2:
- seval_list_builtin_sval ge sp lbs2 rs0 m0 = Some nil ->
- lbs2 = nil.
-Proof.
- destruct lbs2; simpl; auto.
- intros. destruct (seval_builtin_sval _ _ _ _ _);
- try destruct (seval_list_builtin_sval _ _ _ _ _); discriminate.
-Qed.
-
-Lemma seval_builtin_sval_arg (ge:RTL.genv) sp rs0 m0 bs:
- forall ba m v,
- seval_builtin_sval ge sp bs rs0 m0 = Some ba ->
- eval_builtin_arg ge (fun id => id) sp m ba v ->
- seval_builtin_arg ge sp m rs0 m0 bs v.
-Proof.
- induction bs; simpl;
- try (intros ba m v H; inversion H; subst; clear H;
- intros H; inversion H; subst;
- econstructor; auto; fail).
- - intros ba m v; destruct (seval_sval _ _ _ _ _) eqn: SV;
- intros H; inversion H; subst; clear H.
- intros H; inversion H; subst.
- econstructor; auto.
- - intros ba m v.
- destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence.
- destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence.
- intros H; inversion H; subst; clear H.
- intros H; inversion H; subst.
- econstructor; eauto.
- - intros ba m v.
- destruct (seval_builtin_sval _ _ bs1 _ _) eqn: SV1; try congruence.
- destruct (seval_builtin_sval _ _ bs2 _ _) eqn: SV2; try congruence.
- intros H; inversion H; subst; clear H.
- intros H; inversion H; subst.
- econstructor; eauto.
-Qed.
-
-Lemma seval_builtin_arg_sval ge sp m rs0 m0 v: forall bs,
- seval_builtin_arg ge sp m rs0 m0 bs v ->
- exists ba,
- seval_builtin_sval ge sp bs rs0 m0 = Some ba
- /\ eval_builtin_arg ge (fun id => id) sp m ba v.
-Proof.
- induction 1.
- all: try (eexists; constructor; [simpl; reflexivity | constructor]).
- 2-3: try assumption.
- - eexists. constructor.
- + simpl. rewrite H. reflexivity.
- + constructor.
- - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1).
- destruct IHseval_builtin_arg2 as (ba2 & A2 & B2).
- eexists. constructor.
- + simpl. rewrite A1. rewrite A2. reflexivity.
- + constructor; assumption.
- - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1).
- destruct IHseval_builtin_arg2 as (ba2 & A2 & B2).
- eexists. constructor.
- + simpl. rewrite A1. rewrite A2. reflexivity.
- + constructor; assumption.
-Qed.
-
-Lemma seval_builtin_sval_args (ge:RTL.genv) sp rs0 m0 lbs:
- forall lba m v,
- seval_list_builtin_sval ge sp lbs rs0 m0 = Some lba ->
- list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba v ->
- seval_builtin_args ge sp m rs0 m0 lbs v.
-Proof.
- unfold seval_builtin_args; induction lbs; simpl; intros lba m v.
- - intros H; inversion H; subst; clear H.
- intros H; inversion H. econstructor.
- - destruct (seval_builtin_sval _ _ _ _ _) eqn:SV; try congruence.
- destruct (seval_list_builtin_sval _ _ _ _ _) eqn: SVL; try congruence.
- intros H; inversion H; subst; clear H.
- intros H; inversion H; subst; clear H.
- econstructor; eauto.
- eapply seval_builtin_sval_arg; eauto.
-Qed.
-
-Lemma seval_builtin_args_sval ge sp m rs0 m0 lv: forall lbs,
- seval_builtin_args ge sp m rs0 m0 lbs lv ->
- exists lba,
- seval_list_builtin_sval ge sp lbs rs0 m0 = Some lba
- /\ list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba lv.
-Proof.
- induction 1.
- - eexists. constructor.
- + simpl. reflexivity.
- + constructor.
- - destruct IHlist_forall2 as (lba & A & B).
- apply seval_builtin_arg_sval in H. destruct H as (ba & A' & B').
- eexists. constructor.
- + simpl. rewrite A'. rewrite A. reflexivity.
- + constructor; assumption.
-Qed.
-
-Lemma seval_builtin_sval_correct ge sp m rs0 m0: forall bs1 v bs2,
- seval_builtin_arg ge sp m rs0 m0 bs1 v ->
- (seval_builtin_sval ge sp bs1 rs0 m0) = (seval_builtin_sval ge sp bs2 rs0 m0) ->
- seval_builtin_arg ge sp m rs0 m0 bs2 v.
-Proof.
- intros. exploit seval_builtin_arg_sval; eauto.
- intros (ba & X1 & X2).
- eapply seval_builtin_sval_arg; eauto.
- congruence.
-Qed.
-
-Lemma seval_list_builtin_sval_correct ge sp m rs0 m0 vargs: forall lbs1,
- seval_builtin_args ge sp m rs0 m0 lbs1 vargs ->
- forall lbs2, (seval_list_builtin_sval ge sp lbs1 rs0 m0) = (seval_list_builtin_sval ge sp lbs2 rs0 m0) ->
- seval_builtin_args ge sp m rs0 m0 lbs2 vargs.
-Proof.
- intros. exploit seval_builtin_args_sval; eauto.
- intros (ba & X1 & X2).
- eapply seval_builtin_sval_args; eauto.
- congruence.
-Qed.
-
-(* 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 ->
- (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 ->
- (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 ->
- (seval_list_builtin_sval (the_ge1 ctx) (the_sp ctx) lbs1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_list_builtin_sval (the_ge2 ctx) (the_sp ctx) lbs2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Sbuiltin ef lbs1 br pc1) (Sbuiltin ef lbs2 br pc2)
- | Sjumptable_simu sv1 sv2 lpc1 lpc2:
- ptree_get_list dm lpc2 = Some lpc1 ->
- (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_none: sfval_simu dm f opc1 opc2 ctx (Sreturn None) (Sreturn None)
- | Sreturn_simu_some sv1 sv2:
- (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 (Sreturn (Some sv1)) (Sreturn (Some sv2)).
-
-Definition sstate_simu dm f outframe (s1 s2: sstate) (ctx: simu_proof_context f): Prop :=
- sistate_simu dm f outframe s1.(internal) s2.(internal) ctx
- /\ forall is1,
- ssem_internal (the_ge1 ctx) (the_sp ctx) s1 (the_rs0 ctx) (the_m0 ctx) is1 ->
- is1.(icontinue) = true ->
- sfval_simu dm f s1.(si_pc) s2.(si_pc) ctx s1.(final) s2.(final).
-
-Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop :=
- forall st1, sexec f1 pc1 = Some st1 ->
- exists path st2, (fn_path f1)!pc1 = Some path /\ sexec f2 pc2 = Some st2
- /\ forall ctx, sstate_simu dm f1 path.(pre_output_regs) st1 st2 ctx.
diff --git a/scheduling/RTLpathScheduler.v b/scheduling/RTLpathScheduler.v
deleted file mode 100644
index 31680256..00000000
--- a/scheduling/RTLpathScheduler.v
+++ /dev/null
@@ -1,329 +0,0 @@
-(** RTLpath Scheduling from an external oracle.
-
-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 RTLpathWFcheck.
-
-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.
-
-(** 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 !
-
-*)
-
-(* 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".
-
-Program Definition function_builder (tfr: RTL.function) (tpm: path_map) :
- { r : res RTLpath.function | forall f', r = OK f' -> fn_RTL f' = tfr} :=
- match RTLpathWFcheck.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 RTLpathWFcheck.function_checker_path_entry. auto.
-Defined. Next Obligation.
- apply RTLpathWFcheck.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)")
- | 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 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.
-
-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 tf;
- OK tt.
-
-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.
- eapply entrypoint_check_correct; eauto.
-Qed.
-
-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 (fn_path tf) pc2.
-Proof.
- unfold function_equiv_checker. intros. explore.
- exploit path_entry_check_correct. eassumption. all: eauto. intuition.
-Qed.
-
-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.
- unfold function_equiv_checker. intros. explore.
- exploit path_entry_check_correct. eassumption. all: eauto. intuition.
-Qed.
-
-Lemma function_equiv_checker_correct f tf dm:
- function_equiv_checker dm f tf = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- sexec_simu dm f tf pc1 pc2.
-Proof.
- unfold function_equiv_checker. intros. explore.
- eapply simu_check_correct; eauto.
-Qed.
-
-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
- 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 (fn_path tf) 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
- function_equiv_checker_pathentry1 function_equiv_checker_pathentry2
- function_equiv_checker_correct: core.
- 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.
-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 -> sexec_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'} :=
- match (verified_scheduler f) with
- | Error e => Error e
- | OK (tf, dm) => OK tf
- end.
-Next Obligation.
- exploit verified_scheduler_correct; eauto.
- intros (A & B & C & D & E & F & G (* & H *)).
- 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.
- destruct H as [SIG PARAM SIZE ENTRY CORRECT].
- intuition.
-Qed.
-
-Definition transf_fundef (f: fundef) : res fundef :=
- transf_partial_fundef (fun f => proj1_sig (transf_function f)) f.
-
-Definition transf_program (p: program) : res program :=
- transform_partial_program transf_fundef p.
-
-(** * 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' path
- (TRANSF: match_function dupmap f f')
- (DUPLIC: dupmap!pc' = Some pc)
- (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' path
- (STACKS: list_forall2 match_stackframes st st')
- (TRANSF: match_function dupmap f f')
- (DUPLIC: dupmap!pc' = Some pc)
- (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 st st' f f' args m
- (STACKS: list_forall2 match_stackframes st st')
- (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 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 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:
- list_forall2 match_stackframes stk1 stk2 ->
- forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
- list_forall2 match_stackframes stk1 stk3.
-Proof.
- Local Hint Resolve match_stackframes_equiv: core.
- induction 1; intros stk3 EQ; inv EQ; 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: core.
- 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.
-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: core.
- 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':
- transf_fundef f = OK f' -> match_fundef 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.
- eapply match_Internal; eauto.
- + eapply match_External.
-Qed.
diff --git a/scheduling/RTLpathScheduleraux.ml b/scheduling/RTLpathScheduleraux.ml
deleted file mode 100644
index f3f09954..00000000
--- a/scheduling/RTLpathScheduleraux.ml
+++ /dev/null
@@ -1,497 +0,0 @@
-open DebugPrint
-open Machine
-open RTLpathLivegenaux
-open RTLpath
-open RTLpathCommon
-open RTL
-open Maps
-open Registers
-open ExpansionOracle
-
-let config = Machine.config
-
-let print_superblock (sb: superblock) code =
- let insts = sb.instructions in
- let li = sb.liveins in
- let outs = sb.s_output_regs in
- begin
- debug "{ instructions = "; print_instructions (Array.to_list insts) code; debug "\n";
- debug " liveins = "; print_ptree_regset li; debug "\n";
- debug " output_regs = "; print_regset outs; debug "\n}"
- end
-
-let print_superblocks lsb code =
- let rec f = function
- | [] -> ()
- | sb :: lsb -> (print_superblock sb code; debug ",\n"; f lsb)
- in begin
- debug "[\n";
- f lsb;
- debug "]"
- end
-
-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
- let rec follow pc n =
- let inst = get_some @@ PTree.get pc code in
- 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
- 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 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;
- s_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
- (* debug_flag := true; *)
- debug "Superblocks identified:"; print_superblocks lsb code; debug "\n";
- (* debug_flag := false; *)
- lsb
-end
-
-(** the useful one. Returns a hashtable with bindings of shape
- ** [(r,(t, n)], where [r] is a pseudo-register (Registers.reg),
- ** [t] is its class (according to [typing]), and [n] the number of
- ** times it's referenced as an argument in instructions of [seqa] ;
- ** and an arrray containg the list of regs referenced by each
- ** instruction, with a boolean to know whether it's as arg or dest *)
-let reference_counting (seqa : (instruction * Regset.t) array)
- (out_regs : Registers.Regset.t) (typing : RTLtyping.regenv) :
- (Registers.reg, int * int) Hashtbl.t *
- (Registers.reg * bool) list array =
- let retl = Hashtbl.create 42 in
- let retr = Array.make (Array.length seqa) [] in
- (* retr.(i) : (r, b) -> (r', b') -> ...
- * where b = true if seen as arg, false if seen as dest
- *)
- List.iter (fun reg ->
- Hashtbl.add retl
- reg (Machregsaux.class_of_type (typing reg), 1)
- ) (Registers.Regset.elements out_regs);
- let add_reg reg =
- match Hashtbl.find_opt retl reg with
- | Some (t, n) -> Hashtbl.add retl reg (t, n+1)
- | None -> Hashtbl.add retl reg (Machregsaux.class_of_type
- (typing reg), 1)
- in
- let map_true = List.map (fun r -> r, true) in
- Array.iteri (fun i (ins, _) ->
- match ins with
- | Iop(_,args,dest,_) | Iload(_,_,_,args,dest,_) ->
- List.iter (add_reg) args;
- retr.(i) <- (dest, false)::(map_true args)
- | Icond(_,args,_,_,_) ->
- List.iter (add_reg) args;
- retr.(i) <- map_true args
- | Istore(_,_,args,src,_) ->
- List.iter (add_reg) args;
- add_reg src;
- retr.(i) <- (src, true)::(map_true args)
- | Icall(_,fn,args,dest,_) ->
- List.iter (add_reg) args;
- retr.(i) <- (match fn with
- | Datatypes.Coq_inl reg ->
- add_reg reg;
- (dest,false)::(reg, true)::(map_true args)
- | _ -> (dest,false)::(map_true args))
-
- | Itailcall(_,fn,args) ->
- List.iter (add_reg) args;
- retr.(i) <- (match fn with
- | Datatypes.Coq_inl reg ->
- add_reg reg;
- (reg, true)::(map_true args)
- | _ -> map_true args)
- | Ibuiltin(_,args,dest,_) ->
- let rec bar = function
- | AST.BA r -> add_reg r;
- retr.(i) <- (r, true)::retr.(i)
- | AST.BA_splitlong (hi, lo) | AST.BA_addptr (hi, lo) ->
- bar hi; bar lo
- | _ -> ()
- in
- List.iter (bar) args;
- let rec bad = function
- | AST.BR r -> retr.(i) <- (r, false)::retr.(i)
- | AST.BR_splitlong (hi, lo) ->
- bad hi; bad lo
- | _ -> ()
- in
- bad dest;
- | Ijumptable (reg,_) | Ireturn (Some reg) ->
- add_reg reg;
- retr.(i) <- [reg, true]
- | _ -> ()
- ) seqa;
- (* print_string "mentions\n";
- * Array.iteri (fun i l ->
- * print_int i;
- * print_string ": [";
- * List.iter (fun (r, b) ->
- * print_int (Camlcoq.P.to_int r);
- * print_string ":";
- * print_string (if b then "a:" else "d");
- * if b then print_int (snd (Hashtbl.find retl r));
- * print_string ", "
- * ) l;
- * print_string "]\n";
- * flush stdout;
- * ) retr; *)
- retl, retr
-
-
-let get_live_regs_entry (sb : superblock) code =
- (if !Clflags.option_debug_compcert > 6
- then debug_flag := true);
- debug "getting live regs for superblock:\n";
- print_superblock sb code;
- debug "\n";
- let seqa = Array.map (fun i ->
- (match PTree.get i code with
- | Some ii -> ii
- | None -> failwith "RTLpathScheduleraux.get_live_regs_entry"
- ),
- (match PTree.get i sb.liveins with
- | Some s -> s
- | None -> Regset.empty))
- sb.instructions in
- let ret =
- Array.fold_right (fun (ins, liveins) regset_i ->
- let regset = Registers.Regset.union liveins regset_i in
- match ins with
- | Inop _ -> regset
- | Iop (_, args, dest, _)
- | Iload (_, _, _, args, dest, _) ->
- List.fold_left (fun set reg -> Registers.Regset.add reg set)
- (Registers.Regset.remove dest regset) args
- | Istore (_, _, args, src, _) ->
- List.fold_left (fun set reg -> Registers.Regset.add reg set)
- (Registers.Regset.add src regset) args
- | Icall (_, fn, args, dest, _) ->
- List.fold_left (fun set reg -> Registers.Regset.add reg set)
- ((match fn with
- | Datatypes.Coq_inl reg -> (Registers.Regset.add reg)
- | Datatypes.Coq_inr _ -> (fun x -> x))
- (Registers.Regset.remove dest regset))
- args
- | Itailcall (_, fn, args) ->
- List.fold_left (fun set reg -> Registers.Regset.add reg set)
- (match fn with
- | Datatypes.Coq_inl reg -> (Registers.Regset.add reg regset)
- | Datatypes.Coq_inr _ -> regset)
- args
- | Ibuiltin (_, args, dest, _) ->
- List.fold_left (fun set arg ->
- let rec add reg set =
- match reg with
- | AST.BA r -> Registers.Regset.add r set
- | AST.BA_splitlong (hi, lo)
- | AST.BA_addptr (hi, lo) -> add hi (add lo set)
- | _ -> set
- in add arg set)
- (let rec rem dest set =
- match dest with
- | AST.BR r -> Registers.Regset.remove r set
- | AST.BR_splitlong (hi, lo) -> rem hi (rem lo set)
- | _ -> set
- in rem dest regset)
- args
- | Icond (_, args, _, _, _) ->
- List.fold_left (fun set reg ->
- Registers.Regset.add reg set)
- regset args
- | Ijumptable (reg, _)
- | Ireturn (Some reg) ->
- Registers.Regset.add reg regset
- | _ -> regset
- ) seqa sb.s_output_regs
- in debug "live in regs: ";
- print_regset ret;
- debug "\n";
- debug_flag := false;
- ret
-
-(* TODO David *)
-let schedule_superblock sb code =
- 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; *)
- let nr_instr = Array.length sb.instructions in
- let trailer_length =
- match PTree.get (sb.instructions.(nr_instr-1)) code with
- | None -> 0
- | Some ii ->
- match predicted_successor ii with
- | Some _ -> 0
- | None -> 1 in
- debug "hello\n";
- let live_regs_entry = get_live_regs_entry sb code in
- let seqa =
- Array.map (fun i ->
- (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)) in
- match PrepassSchedulingOracle.schedule_sequence
- seqa
- live_regs_entry
- sb.typing
- (reference_counting seqa sb.s_output_regs sb.typing) with
- | None -> sb.instructions
- | Some order ->
- 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;
- debug_flag := old_flag;
- flush stdout; *)
- assert ((Array.length sb.instructions) = (Array.length ins'));
- (*sb.instructions; *)
- ins';;
-
- (* 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 *)
-
-(**
- * 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 is_a_cb = function Icond _ -> true | _ -> false
-let is_a_load = function Iload _ -> true | _ -> false
-
-let find_array arr n =
- let index = ref None in
- begin
- Array.iteri (fun i n' ->
- if n = n' then
- match !index with
- | Some _ -> failwith "More than one element present"
- | None -> index := Some i
- ) arr;
- !index
- end
-
-let rec hashedset_from_list = function
- | [] -> HashedSet.PSet.empty
- | n::ln -> HashedSet.PSet.add n (hashedset_from_list ln)
-
-let hashedset_map f hs = hashedset_from_list @@ List.map f @@ HashedSet.PSet.elements hs
-
-let apply_schedule code sb new_order =
- let tc = ref code in
- let old_order = sb.instructions in
- let count_cbs order code =
- let current_cbs = ref HashedSet.PSet.empty in
- let cbs_above = ref PTree.empty in
- Array.iter (fun n ->
- let inst = get_some @@ PTree.get n code in
- if is_a_cb inst then current_cbs := HashedSet.PSet.add n !current_cbs
- else if is_a_load inst then cbs_above := PTree.set n !current_cbs !cbs_above
- ) order;
- !cbs_above
- in let fmap n =
- let index = get_some @@ find_array new_order n in
- old_order.(index)
- in begin
- check_order code old_order new_order;
- (* First pass - modify the positions, nothing else *)
- 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;
- (* Second pass - turn the loads back into trapping when it was not needed *)
- (* 1) We remember which CBs are "above" a given load *)
- let cbs_above = count_cbs old_order code in
- (* 2) We do the same for new_order *)
- let cbs_above' = count_cbs (Array.map fmap new_order) !tc in
- (* 3) We examine each load, turn it back into trapping if cbs_above is included in cbs_above' *)
- Array.iter (fun n ->
- let n' = fmap n in
- let inst' = get_some @@ PTree.get n' !tc in
- match inst' with
- | Iload (t,a,b,c,d,s) ->
- let pset = hashedset_map fmap @@ get_some @@ PTree.get n cbs_above in
- let pset' = get_some @@ PTree.get n' cbs_above' in
- if HashedSet.PSet.is_subset pset pset' then tc := PTree.set n' (Iload (AST.TRAP,a,b,c,d,s)) !tc
- else assert !config.has_non_trapping_loads
- | _ -> ()
- ) old_order;
- !tc
- end
-
-let turn_all_loads_nontrap sb code =
- if not !config.has_non_trapping_loads then code
- else begin
- let code' = ref code in
- Array.iter (fun n ->
- let inst = get_some @@ PTree.get n code in
- match inst with
- | Iload (t,a,b,c,d,s) -> code' := PTree.set n (Iload (AST.NOTRAP,a,b,c,d,s)) !code'
- | _ -> ()
- ) sb.instructions;
- !code'
- end
-
-let rec do_schedule code pm = function
- | [] -> (code, pm)
- | sb :: lsb ->
- (*debug_flag := true;*)
- let (code_exp, pm) = expanse sb code pm in
- (*debug_flag := false;*)
- (* Trick: instead of turning loads into non trap as needed..
- * First, we turn them all into non-trap.
- * Then, we turn back those who didn't need to be turned, into TRAP again
- * This is because the scheduler (rightfully) refuses to schedule ahead of a branch
- * operations that might trap *)
- let code' = turn_all_loads_nontrap sb code_exp in
- let schedule = schedule_superblock sb code' in
- let new_code = apply_schedule code' sb schedule in
- begin
- (*debug_flag := true;*)
- if code != code_exp then (
- debug "Old Code: "; print_code code;
- debug "Exp Code: "; print_code code_exp);
- debug "\nSchedule to apply: "; print_arrayp schedule;
- debug "\nNew Code: "; print_code new_code;
- debug "\n";
- do_schedule new_code pm 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 typing = get_ok @@ RTLtyping.type_function f.fn_RTL in
- let lsb = get_superblocks code entry pm typing in
- begin
- (* debug_flag := true; *)
- debug "Pathmap:\n"; debug "\n";
- print_path_map pm;
- debug "Superblocks:\n";
- (*print_code code; flush stdout; flush stderr;*)
- (*debug_flag := false;*)
- (*print_superblocks lsb code; debug "\n";*)
- find_last_node_reg (PTree.elements code);
- let (tc, pm) = do_schedule code pm lsb in
- (((tc, entry), pm), id_ptree)
- end
diff --git a/scheduling/RTLpathSchedulerproof.v b/scheduling/RTLpathSchedulerproof.v
deleted file mode 100644
index a9c2fa76..00000000
--- a/scheduling/RTLpathSchedulerproof.v
+++ /dev/null
@@ -1,509 +0,0 @@
-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 RTLpathScheduler.
-
-Definition match_prog (p tp: program) :=
- match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
-
-Lemma transf_program_match:
- forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
-Proof.
- intros. eapply match_transform_partial_program_contextual; eauto.
-Qed.
-
-Section PRESERVATION.
-
-Variable prog: program.
-Variable tprog: program.
-
-Hypothesis TRANSL: match_prog prog tprog.
-
-Let pge := Genv.globalenv prog.
-Let tpge := Genv.globalenv tprog.
-
-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.
- rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
-Qed.
-
-Lemma senv_preserved:
- Senv.equiv pge tpge.
-Proof.
- eapply (Genv.senv_match TRANSL).
-Qed.
-
-Lemma functions_preserved:
- forall (v: val) (f: fundef),
- Genv.find_funct pge v = Some f ->
- exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tpge v = Some tf /\ linkorder cunit prog.
-Proof.
- intros. exploit (Genv.find_funct_match TRANSL); eauto.
- intros (cu & tf & A & B & C).
- repeat eexists; intuition eauto.
- + unfold incl; auto.
- + eapply linkorder_refl.
-Qed.
-
-Lemma function_ptr_preserved:
- forall v f,
- Genv.find_funct_ptr pge v = Some f ->
- exists tf,
- Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf.
-Proof.
- intros.
- exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
-Qed.
-
-Lemma function_sig_preserved:
- forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f.
-Proof.
- intros. destruct f.
- - 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.
-
-Theorem transf_initial_states:
- forall s1, initial_state prog s1 ->
- exists s2, initial_state tprog s2 /\ match_states s1 s2.
-Proof.
- intros. inv H.
- exploit function_ptr_preserved; eauto. intros (tf & FIND & TRANSF).
- exists (Callstate nil tf nil m0).
- split.
- - 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.
- + destruct f.
- * 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.
- + apply transf_fundef_correct; auto.
-(* + eapply all_fundef_liveness_ok; eauto. *)
-Qed.
-
-Theorem transf_final_states s1 s2 r:
- final_state s1 r -> match_states s1 s2 -> final_state s2 r.
-Proof.
- 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.
-
-
-Let ge := Genv.globalenv (RTLpath.transf_program prog).
-Let tge := Genv.globalenv (RTLpath.transf_program 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 symbols_preserved_RTL s: Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof.
- unfold tge, ge. erewrite RTLpath.symbols_preserved; eauto.
- rewrite symbols_preserved.
- erewrite RTLpath.symbols_preserved; eauto.
-Qed.
-
-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.
- erewrite symbols_preserved_RTL. eauto.
-Qed.
-
-Lemma s_find_function_fundef f sp svos rs0 m0 fd
- (LIVE: liveness_ok_function f):
- sfind_function pge ge sp svos rs0 m0 = Some fd ->
- liveness_ok_fundef fd.
-Proof.
- unfold sfind_function. destruct svos; simpl.
- + destruct (seval_sval _ _ _ _); try congruence.
- eapply find_funct_liveness_ok; eauto.
- + destruct (Genv.find_symbol _ _); try congruence.
- intros. eapply all_fundef_liveness_ok; eauto.
-Qed.
-Local Hint Resolve s_find_function_fundef: core.
-
-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'.
-Proof.
- Local Hint Resolve symbols_preserved_RTL: core.
- 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 & (X1 & X2 & X3)). eexists.
- repeat split; eauto.
- + subst. rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence.
- intros; exploit function_ptr_preserved; eauto.
-Qed.
-
-Lemma sistate_simu f dupmap outframe sp st st' rs m is
- (LIVE: liveness_ok_function f):
- ssem_internal ge sp st rs m is ->
- sistate_simu dupmap f outframe st st' (mkctx sp rs m LIVE)->
- exists is',
- ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap outframe is is'.
-Proof.
- intros SEM X; eapply X; eauto.
-Qed.
-
-Lemma seval_builtin_sval_preserved sp rs m:
- forall bs, seval_builtin_sval ge sp bs rs m = seval_builtin_sval tge sp bs rs m.
-Proof.
- induction bs.
- all: try (simpl; try reflexivity; erewrite seval_preserved by eapply symbols_preserved_RTL; reflexivity).
- all: simpl; rewrite IHbs1; rewrite IHbs2; reflexivity.
-Qed.
-
-Lemma seval_list_builtin_sval_preserved sp rs m:
- forall lbs,
- seval_list_builtin_sval ge sp lbs rs m = seval_list_builtin_sval tge sp lbs rs m.
-Proof.
- induction lbs; [simpl; reflexivity|].
- simpl. rewrite seval_builtin_sval_preserved. rewrite IHlbs.
- reflexivity.
-Qed.
-
-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' ->
- list_forall2 match_stackframes stk stk' ->
- 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 STK SFV. destruct SFV; intros SEM; inv SEM; simpl in *.
- - (* Snone *)
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- eapply eqlive_reg_refl.
- - (* Scall *)
- exploit s_find_function_preserved; eauto.
- intros (fd' & FIND & TRANSF).
- erewrite <- function_sig_preserved; eauto.
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- + eapply eq_trans; try eassumption; auto.
- + simpl. repeat (econstructor; eauto).
- - (* Stailcall *)
- exploit s_find_function_preserved; eauto.
- intros (fd' & FIND & TRANSF).
- erewrite <- function_sig_preserved; eauto.
- eexists; split; econstructor; eauto.
- + erewrite <- preserv_fnstacksize; eauto.
- + eapply eq_trans; try eassumption; auto.
- - (* 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 seval_list_builtin_sval_correct; eauto.
- rewrite H0.
- erewrite seval_list_builtin_sval_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. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- + eapply eq_trans; try eassumption; auto.
- + eapply eqlive_reg_refl.
- - (* Sreturn *)
- eexists; split; econstructor; eauto.
- erewrite <- preserv_fnstacksize; eauto.
- - (* Sreturn bis *)
- eexists; split; econstructor; eauto.
- + erewrite <- preserv_fnstacksize; eauto.
- + rewrite <- H. erewrite <- seval_preserved; eauto.
-Qed.
-
-Lemma siexec_snone_por_correct rs' is t s alive path0 i sp s0 st0 stk stk' f rs0 m0: forall
- (SSEM2 : ssem_final pge ge sp (si_pc s0) stk f rs0 m0 Snone
- (irs is) (imem is) t s)
- (SIEXEC : siexec_inst i st0 = Some s0)
- (ICHK : inst_checker (fn_path f) alive (pre_output_regs path0) i = Some tt),
- (liveness_ok_function f) ->
- list_forall2 match_stackframes stk stk' ->
- eqlive_reg (fun r : Regset.elt => Regset.In r (pre_output_regs path0)) (irs is) rs' ->
- exists s' : state,
- ssem_final pge ge sp (si_pc s0) stk f rs0 m0 Snone rs' (imem is) t s' /\
- eqlive_states s s'.
-Proof.
- Local Hint Resolve eqlive_stacks_refl: core.
- intros ? ? ? LIVE STK EQLIVE.
- inversion SSEM2; subst; clear SSEM2.
- eexists; split.
- * econstructor.
- * generalize ICHK.
- unfold inst_checker. destruct i; simpl in *;
- unfold exit_checker; try discriminate.
- all:
- try destruct (list_mem _ _); simpl;
- try (destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence; fail).
- 4,5:
- destruct (Regset.mem _ _); destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence.
- 1,2,3,4: assert (NPC: n=(si_pc s0)).
- all: try (inv SIEXEC; simpl; auto; fail).
- 1,2,3,4:
- try (destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence);
- simpl; inversion_SOME p;
- destruct (Regset.subset (input_regs p) (pre_output_regs path0)) eqn:SUB_PATH; try congruence;
- intros NPATH _; econstructor; eauto;
- try (instantiate (1:=p); rewrite <- NPC; auto; fail).
- 1,2,3,4:
- eapply eqlive_reg_monotonic; eauto; simpl;
- intros; apply Regset.subset_2 in SUB_PATH;
- unfold Regset.Subset in SUB_PATH;
- apply SUB_PATH in H; auto.
- assert (NPC: n0=(si_pc s0)). { inv SIEXEC; simpl; auto. }
- inversion_SOME p.
- 2: { destruct (Regset.subset _ _) eqn:?; try congruence. }
- destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence.
- 2: { destruct (Regset.subset (pre_output_regs path0) alive) eqn:?; try congruence. }
- simpl.
- destruct (Regset.subset (pre_output_regs path0) alive) eqn:SUB_ALIVE'; try congruence.
- inversion_SOME p'.
- destruct (Regset.subset (input_regs p') (pre_output_regs path0)) eqn:SUB_PATH; try congruence.
- intros NPATH NPATH' _. econstructor; eauto.
- instantiate (1:=p'). rewrite <- NPC; auto.
- eapply eqlive_reg_monotonic; eauto; simpl.
- intros. apply Regset.subset_2 in SUB_PATH.
- unfold Regset.Subset in SUB_PATH.
- apply SUB_PATH in H; auto.
-Qed.
-
-Lemma pre_output_regs_correct f pc0 path0 stk stk' sp (st:sstate) rs0 m0 t s is rs':
- (liveness_ok_function f) ->
- (fn_path f) ! pc0 = Some path0 ->
- sexec f pc0 = Some st ->
- list_forall2 match_stackframes stk stk' ->
- ssem_final pge ge sp (si_pc st) stk f rs0 m0 (final st) (irs is) (imem is) t s ->
- eqlive_reg (fun r : Regset.elt => Regset.In r (pre_output_regs path0)) (irs is) rs' ->
- exists s', ssem_final pge ge sp (si_pc st) stk f rs0 m0 (final st) rs' (imem is) t s' /\ eqlive_states s s'.
-Proof.
- Local Hint Resolve eqlive_stacks_refl: core.
- intros LIVE PATH0 SEXEC STK SSEM2 EQLIVE.
- (* start decomposing path_checker *)
- generalize (LIVE pc0 path0 PATH0).
- unfold path_checker.
- inversion_SOME res; intros IPCHK.
- inversion_SOME i; intros INST ICHK.
- exploit ipath_checker_default_succ; eauto. intros DEFSUCC.
- (* start decomposing SEXEC *)
- generalize SEXEC; clear SEXEC.
- unfold sexec; rewrite PATH0.
- inversion_SOME st0; intros SEXEC_PATH.
- exploit siexec_path_default_succ; eauto.
- simpl. rewrite DEFSUCC.
- clear DEFSUCC. destruct res as [alive pc1]. simpl in *.
- try_simplify_someHyps.
- destruct (siexec_inst i st0) eqn: SIEXEC; try_simplify_someHyps; intros.
- (* Snone *)
- eapply siexec_snone_por_correct; eauto.
- destruct i; try_simplify_someHyps; try congruence;
- inversion SSEM2; subst; clear SSEM2; simpl in *.
- + (* Scall *)
- eexists; split.
- * econstructor; eauto.
- * econstructor; eauto.
- econstructor; eauto.
- (* wf *)
- generalize ICHK.
- unfold inst_checker; simpl in *.
- destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence.
- destruct (list_mem _ _); try congruence.
- destruct (reg_sum_mem _ _); try congruence.
- intros EXIT.
- exploit exit_checker_eqlive_ext1; eauto.
- intros. destruct H as [p [PATH EQLIVE']].
- econstructor; eauto.
- + (* Stailcall *)
- eexists; split.
- * econstructor; eauto.
- * econstructor; eauto.
- + (* Sbuiltin *)
- eexists; split.
- * econstructor; eauto.
- * (* wf *)
- generalize ICHK.
- unfold inst_checker; simpl in *.
- destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence.
- destruct (list_mem _ _); try congruence.
- intros EXIT.
- exploit exit_checker_eqlive_builtin_res; eauto.
- intros. destruct H as [p [PATH EQLIVE']].
- econstructor; eauto.
- + (* Sjumptable *)
- eexists; split.
- * econstructor; eauto.
- * (* wf *)
- generalize ICHK.
- unfold inst_checker; simpl in *.
- destruct (Regset.subset _ _) eqn:SUB_ALIVE; try congruence.
- destruct (Regset.mem _ _); try congruence.
- destruct (exit_list_checker _ _ _) eqn:EQL; try congruence.
- exploit exit_list_checker_eqlive; eauto.
- intros. destruct H as [p [PATH EQLIVE']].
- econstructor; eauto.
- + (* Sreturn *)
- eexists; split.
- * econstructor; eauto.
- * econstructor; eauto.
-Qed.
-
-(* The main theorem on simulation of symbolic states ! *)
-Theorem ssem_sstate_simu dm f f' pc0 path0 stk stk' sp st st' rs m t s:
- (fn_path f) ! pc0 = Some path0 ->
- sexec f pc0 = Some st ->
- match_function dm f f' ->
- liveness_ok_function f ->
- list_forall2 match_stackframes stk stk' ->
- ssem pge ge sp st stk f rs m t s ->
- (forall ctx: simu_proof_context f, sstate_simu dm f (pre_output_regs path0) st st' ctx) ->
- exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'.
-Proof.
- intros PATH0 SEXEC 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 in *.
- - (* sem_early *)
- exploit sistate_simu; eauto.
- unfold istate_simu; rewrite CONT.
- intros (is' & SEM' & (path & PATH & (CONT' & RS' & M') & PC')).
- exists (State stk' f' sp (ipc is') (irs is') (imem is')).
- split.
- + 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')).
- exploit pre_output_regs_correct; eauto.
- clear SEM2; intros (s0 & SEM2 & EQLIVE).
- exploit ssem_final_simu; eauto.
- clear SEM2; intros (s1 & SEM2 & MATCH0).
- exploit ssem_final_equiv; eauto.
- clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s2 & EQ & SEM2).
- exists s2; split.
- + eapply ssem_normal; eauto.
- + eapply eqlive_match_states; 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 ->
- 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' ->
- 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 initialize_path. { eapply dupmap_path_entry2; eauto. }
- intros (path' & PATH').
- exists path'.
- exploit (sexec_correct f pc pge ge sp path stk rs m t s); eauto.
- intros (st & SYMB & SEM).
- exploit dupmap_correct; eauto.
- intros (path0 & st' & PATH0 & SYMB' & SIMU).
- rewrite PATH0 in PATH; inversion PATH; subst.
- exploit ssem_sstate_simu; eauto.
- intros (s0 & SEM0 & MATCH).
- exploit (sexec_exact f'); eauto.
- intros (s' & STEP' & EQ).
- exists s'; intuition.
- eapply match_states_equiv; eauto.
-Qed.
-
-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 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.
- 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').
- exists s'; split.
- + eapply exec_path; eauto.
- + eapply eqlive_match_states; eauto.
-(* exec_function_internal *)
- - 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; 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.
- + constructor. assumption.
-(* exec_return *)
- - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split.
- + constructor.
- + inv H1. econstructor; eauto.
-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.
- - intros; eapply transf_final_states; eauto.
- - intros; eapply step_simulation; eauto.
-Qed.
-
-End PRESERVATION.
diff --git a/scheduling/RTLpathWFcheck.v b/scheduling/RTLpathWFcheck.v
deleted file mode 100644
index 63b914ec..00000000
--- a/scheduling/RTLpathWFcheck.v
+++ /dev/null
@@ -1,187 +0,0 @@
-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.
-Require RTLpathLivegen.
-
-Local Open Scope lazy_bool_scope.
-
-Local Open Scope option_monad_scope.
-
-Definition exit_checker {A} (pm: path_map) (pc: node) (v:A): option A :=
- SOME path <- pm!pc IN
- Some v.
-
-Lemma exit_checker_path_entry A (pm: path_map) (pc: node) (v:A) res:
- exit_checker pm 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) (pc: node) (v:A) res:
- exit_checker pm pc v = Some res -> v=res.
-Proof.
- unfold exit_checker, path_entry.
- inversion_SOME path; try_simplify_someHyps.
-Qed.
-
-Definition iinst_checker (pm: path_map) (i: instruction): option (node) :=
- match i with
- | Inop pc' | Iop _ _ _ pc' | Iload _ _ _ _ _ pc'
- | Istore _ _ _ _ pc' => Some (pc')
- | Icond cond args ifso ifnot _ =>
- exit_checker pm ifso ifnot
- | _ => None
- end.
-
-Local Hint Resolve exit_checker_path_entry: core.
-
-Lemma iinst_checker_path_entry (pm: path_map) (i: instruction) res pc:
- iinst_checker pm i = Some res ->
- early_exit i = Some pc -> path_entry pm pc.
-Proof.
- destruct i; simpl; try_simplify_someHyps; subst.
-Qed.
-
-Lemma iinst_checker_default_succ (pm: path_map) (i: instruction) res pc:
- iinst_checker pm i = Some res ->
- pc = 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) (pc:node): option (node) :=
- match ps with
- | O => Some (pc)
- | S p =>
- SOME i <- f.(fn_code)!pc IN
- SOME res <- iinst_checker pm i IN
- ipath_checker p f pm res
- end.
-
-Lemma ipath_checker_wellformed f pm ps: forall pc res,
- ipath_checker ps f pm pc = Some res ->
- wellformed_path f.(fn_code) pm 0 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.
-
-Fixpoint exit_list_checker (pm: path_map) (l: list node): bool :=
- match l with
- | nil => true
- | pc::l' => exit_checker pm pc tt &&& exit_list_checker pm l'
- end.
-
-Lemma exit_list_checker_correct pm l pc:
- exit_list_checker pm l = true -> List.In pc l -> exit_checker pm pc tt = Some tt.
-Proof.
- intros EXIT PC; induction l; intuition.
- simpl in * |-. rewrite RTLpathLivegen.lazy_and_Some_tt_true in EXIT.
- firstorder (subst; eauto).
-Qed.
-
-Local Hint Resolve exit_list_checker_correct: core.
-
-Definition inst_checker (pm: path_map) (i: instruction): option unit :=
- match i with
- | Icall sig ros args res pc' =>
- exit_checker pm pc' tt
- | Itailcall sig ros args =>
- Some tt
- | Ibuiltin ef args res pc' =>
- exit_checker pm pc' tt
- | Ijumptable arg tbl =>
- ASSERT exit_list_checker pm tbl IN
- Some tt
- | Ireturn optarg =>
- Some tt
- | _ =>
- SOME res <- iinst_checker pm i IN
- exit_checker pm res tt
- end.
-
-Lemma inst_checker_wellformed (c:code) pc (pm: path_map) (i: instruction):
- inst_checker pm 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 pc IN
- SOME i <- f.(fn_code)!res IN
- inst_checker pm 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.
-
-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 RTLpathLivegen.lazy_and_Some_tt_true in CHECKER. intuition (subst; auto).
-Qed.
-
-Definition function_checker (f: RTL.function) (pm: path_map): 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 RTLpathLivegen.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 RTLpathLivegen.lazy_and_Some_true;
- unfold path_entry. firstorder congruence.
-Qed.
diff --git a/scheduling/RTLpathproof.v b/scheduling/RTLpathproof.v
deleted file mode 100644
index 20eded97..00000000
--- a/scheduling/RTLpathproof.v
+++ /dev/null
@@ -1,50 +0,0 @@
-Require Import Coqlib Maps.
-Require Import AST Integers Values Events Memory Globalenvs Smallstep.
-Require Import Op Registers.
-Require Import RTL Linking.
-Require Import RTLpath.
-
-Definition match_prog (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_prog p (transf_program p).
-Proof.
- intros. eapply match_transform_program; eauto.
-Qed.
-
-Lemma match_program_transf:
- forall p tp, match_prog p tp -> transf_program 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 transf_program. 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_prog 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.
-
-
diff --git a/scheduling/RTLtoBTL.v b/scheduling/RTLtoBTL.v
new file mode 100644
index 00000000..309c616e
--- /dev/null
+++ b/scheduling/RTLtoBTL.v
@@ -0,0 +1,27 @@
+Require Import Coqlib Maps.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad BTL.
+Require Export BTLmatchRTL BTL_Livecheck.
+
+Require Import Errors Linking.
+
+(** External oracle *)
+Axiom rtl2btl: RTL.function -> BTL.code * node * (PTree.t node).
+
+Extract Constant rtl2btl => "RTLtoBTLaux.rtl2btl".
+
+Local Open Scope error_monad_scope.
+
+Definition transf_function (f: RTL.function) : res BTL.function :=
+ let (tcte, dupmap) := rtl2btl f in
+ let (tc, te) := tcte in
+ let f' := BTL.mkfunction (RTL.fn_sig f) (RTL.fn_params f) (RTL.fn_stacksize f) tc te in
+ do u <- verify_function dupmap f' f;
+ do u <- liveness_checker f';
+ OK f'.
+
+Definition transf_fundef (f: RTL.fundef) : res fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: RTL.program) : res program :=
+ transform_partial_program transf_fundef p.
diff --git a/scheduling/RTLtoBTLaux.ml b/scheduling/RTLtoBTLaux.ml
new file mode 100644
index 00000000..3c06f4cb
--- /dev/null
+++ b/scheduling/RTLtoBTLaux.ml
@@ -0,0 +1,119 @@
+open Maps
+open RTL
+open BTL
+open PrintBTL
+open RTLcommonaux
+open DebugPrint
+open BTLtypes
+open BTLcommonaux
+
+let encaps_final inst osucc =
+ match inst with
+ | BF _ -> inst
+ | _ -> Bseq (inst, BF (Bgoto (get_some @@ osucc), def_iinfo ()))
+
+let translate_inst (iinfo : BTL.inst_info) inst is_final =
+ let osucc = ref None in
+ let btli =
+ match inst with
+ | Inop s ->
+ osucc := Some s;
+ Bnop (Some iinfo)
+ | Iop (op, lr, rd, s) ->
+ osucc := Some s;
+ Bop (op, lr, rd, iinfo)
+ | Iload (trap, chk, addr, lr, rd, s) ->
+ osucc := Some s;
+ Bload (trap, chk, addr, lr, rd, iinfo)
+ | Istore (chk, addr, lr, src, s) ->
+ osucc := Some s;
+ Bstore (chk, addr, lr, src, iinfo)
+ | Icall (sign, fn, lr, rd, s) -> BF (Bcall (sign, fn, lr, rd, s), iinfo)
+ | Itailcall (sign, fn, lr) -> BF (Btailcall (sign, fn, lr), iinfo)
+ | Ibuiltin (ef, lr, rd, s) -> BF (Bbuiltin (ef, lr, rd, s), iinfo)
+ | Icond (cond, lr, ifso, ifnot, info) ->
+ osucc := Some ifnot;
+ iinfo.pcond <- info;
+ Bcond
+ ( cond,
+ lr,
+ BF (Bgoto ifso, def_iinfo ()),
+ Bnop None,
+ iinfo )
+ | Ijumptable (arg, tbl) -> BF (Bjumptable (arg, tbl), iinfo)
+ | Ireturn oreg -> BF (Breturn oreg, iinfo)
+ in
+ if is_final then encaps_final btli !osucc else btli
+
+let translate_function code entry join_points liveness (typing : RTLtyping.regenv) =
+ let reached = ref (PTree.map (fun n i -> false) code) in
+ let btl_code = ref PTree.empty in
+ let rec build_btl_tree e =
+ if get_some @@ PTree.get e !reached then ()
+ else (
+ reached := PTree.set e true !reached;
+ let next_nodes = ref [] in
+ let last = ref None in
+ let rec build_btl_block n =
+ let inst = get_some @@ PTree.get n code in
+ let psucc = predicted_successor inst in
+ let iinfo = mk_iinfo (p2i n) None in
+ let succ =
+ match psucc with
+ | Some ps ->
+ if get_some @@ PTree.get ps join_points then None else Some ps
+ | None -> None
+ in
+ match succ with
+ | Some s -> (
+ match inst with
+ | Icond (cond, lr, ifso, ifnot, info) ->
+ assert (s = ifnot);
+ next_nodes := !next_nodes @ non_predicted_successors inst psucc;
+ iinfo.pcond <- info;
+ Bseq
+ ( Bcond
+ (cond, lr, BF (Bgoto ifso, def_iinfo ()), Bnop None, iinfo),
+ build_btl_block s )
+ | _ -> Bseq (translate_inst iinfo inst false, build_btl_block s))
+ | None ->
+ debug "BLOCK END.\n";
+ next_nodes := !next_nodes @ successors_inst inst;
+ last := Some inst;
+ translate_inst iinfo inst true
+ in
+ let ib = build_btl_block e in
+ let succs = !next_nodes in
+
+ let inputs = get_some @@ PTree.get e liveness in
+ let soutput = get_outputs liveness e (get_some @@ !last) in
+
+ let bi = mk_binfo (p2i e) soutput typing in
+ let ibf = { entry = ib; input_regs = inputs; binfo = bi } in
+ btl_code := PTree.set e ibf !btl_code;
+ List.iter build_btl_tree succs)
+ in
+ build_btl_tree entry;
+ !btl_code
+
+let rtl2btl (f : RTL.coq_function) =
+ let code = f.fn_code in
+ let entry = f.fn_entrypoint in
+ let join_points = get_join_points code entry in
+ let liveness = analyze f in
+ let typing = get_ok @@ RTLtyping.type_function f in
+ let btl = translate_function code entry join_points liveness typing in
+ let dm = PTree.map (fun n i -> n) btl in
+ (*debug_flag := true;*)
+ debug "Entry %d\n" (p2i entry);
+ debug "RTL Code: ";
+ print_code code;
+ debug "BTL Code:\n";
+ print_btl_code stderr btl;
+ (*debug_flag := false;*)
+ debug "Dupmap:\n";
+ print_ptree print_pint dm;
+ debug "Join points: ";
+ print_true_nodes join_points;
+ debug "\n";
+ ((btl, entry), dm)
diff --git a/scheduling/RTLtoBTLproof.v b/scheduling/RTLtoBTLproof.v
new file mode 100644
index 00000000..0ca93bce
--- /dev/null
+++ b/scheduling/RTLtoBTLproof.v
@@ -0,0 +1,758 @@
+Require Import Coqlib Maps Lia.
+Require Import AST Integers Values Events Memory Globalenvs Smallstep.
+Require Import RTL Op Registers OptionMonad BTL.
+
+Require Import Errors Linking RTLtoBTL.
+
+Require Import Linking.
+
+(** * Normalization of BTL iblock for simulation of RTL
+
+Below [normRTL] normalizes the representation of BTL blocks,
+in order to represent as sequences of RTL instructions.
+
+This eases the
+
+*)
+
+Definition is_RTLatom (ib: iblock): bool :=
+ match ib with
+ | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None => false
+ | _ => true
+ end.
+
+Definition is_RTLbasic (ib: iblock): bool :=
+ match ib with
+ | Bseq _ _ | Bcond _ _ _ _ _ | Bnop None | BF _ _ => false
+ | _ => true
+ end.
+
+(** The strict [is_normRTL] property specifying the ouput of [normRTL] below *)
+Inductive is_normRTL: iblock -> Prop :=
+ | norm_Bseq ib1 ib2:
+ is_RTLbasic ib1 = true ->
+ is_normRTL ib2 ->
+ is_normRTL (Bseq ib1 ib2)
+ | norm_Bcond cond args ib1 ib2 i:
+ is_normRTL ib1 ->
+ is_normRTL ib2 ->
+ is_normRTL (Bcond cond args ib1 ib2 i)
+ | norm_others ib:
+ is_RTLatom ib = true ->
+ is_normRTL ib
+ .
+Local Hint Constructors is_normRTL: core.
+
+(** Weaker version allowing for trailing [Bnop None]. *)
+Inductive is_wnormRTL: iblock -> Prop :=
+ | wnorm_Bseq ib1 ib2:
+ is_RTLbasic ib1 = true ->
+ (ib2 <> Bnop None -> is_wnormRTL ib2) ->
+ is_wnormRTL (Bseq ib1 ib2)
+ | wnorm_Bcond cond args ib1 ib2 iinfo:
+ (ib1 <> Bnop None -> is_wnormRTL ib1) ->
+ (ib2 <> Bnop None -> is_wnormRTL ib2) ->
+ is_wnormRTL (Bcond cond args ib1 ib2 iinfo)
+ | wnorm_others ib:
+ is_RTLatom ib = true ->
+ is_wnormRTL ib
+ .
+Local Hint Constructors is_wnormRTL: core.
+
+(* NB: [k] is a "continuation" (e.g. semantically [normRTLrec ib k] is like [Bseq ib k]) *)
+Fixpoint normRTLrec (ib: iblock) (k: iblock): iblock :=
+ match ib with
+ | Bseq ib1 ib2 => normRTLrec ib1 (normRTLrec ib2 k)
+ | Bcond cond args ib1 ib2 iinfo =>
+ Bcond cond args (normRTLrec ib1 k) (normRTLrec ib2 k) iinfo
+ | BF fin iinfo => BF fin iinfo
+ | Bnop None => k
+ | ib => Bseq ib k
+ end.
+
+Definition normRTL ib := normRTLrec ib (Bnop None).
+
+Lemma normRTLrec_wcorrect ib: forall k,
+ (k <> (Bnop None) -> is_wnormRTL k) ->
+ (normRTLrec ib k) <> Bnop None ->
+ is_wnormRTL (normRTLrec ib k).
+Proof.
+ induction ib; simpl; intros; repeat autodestruct; auto.
+Qed.
+
+Lemma normRTL_wcorrect ib:
+ (normRTL ib) <> Bnop None ->
+ is_wnormRTL (normRTL ib).
+Proof.
+ intros; eapply normRTLrec_wcorrect; eauto.
+Qed.
+
+Lemma is_join_opt_None {A} (opc1 opc2: option A):
+ is_join_opt opc1 opc2 None -> opc1 = None /\ opc2 = None.
+Proof.
+ intros X. inv X; auto.
+Qed.
+
+Lemma match_iblock_None_not_Bnop dupmap cfg isfst pc ib:
+ match_iblock dupmap cfg isfst pc ib None -> ib <> Bnop None.
+Proof.
+ intros X; inv X; try congruence.
+Qed.
+Local Hint Resolve match_iblock_None_not_Bnop: core.
+
+Lemma is_wnormRTL_normRTL dupmap cfg ib:
+ is_wnormRTL ib ->
+ forall isfst pc
+ (MIB: match_iblock dupmap cfg isfst pc ib None),
+ is_normRTL ib.
+Proof.
+ induction 1; simpl; intros; auto; try (inv MIB); eauto.
+ (* Bcond *)
+ destruct (is_join_opt_None opc1 opc2); subst; eauto.
+ econstructor; eauto.
+Qed.
+
+Local Hint Constructors iblock_istep: core.
+Lemma normRTLrec_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin1:
+ forall (ISTEP: iblock_istep tge sp rs0 m0 ib rs1 m1 ofin1)
+ k ofin2 rs2 m2
+ (CONT: match ofin1 with
+ | None => iblock_istep tge sp rs1 m1 k rs2 m2 ofin2
+ | Some fin1 => rs2=rs1 /\ m2=m1 /\ ofin2=Some fin1
+ end),
+ iblock_istep tge sp rs0 m0 (normRTLrec ib k) rs2 m2 ofin2.
+Proof.
+ induction 1; simpl; intuition subst; eauto.
+ - (* Bnop *) autodestruct; eauto.
+ - (* Bop *) repeat econstructor; eauto.
+ - (* Bload *) inv LOAD.
+ + repeat econstructor; eauto.
+ + do 2 (econstructor; eauto).
+ eapply has_loaded_default; eauto.
+ - (* Bcond *) repeat econstructor; eauto.
+ destruct ofin; intuition subst;
+ destruct b; eapply IHISTEP; eauto.
+Qed.
+
+Lemma normRTL_iblock_istep_correct tge sp ib rs0 m0 rs1 m1 ofin:
+ iblock_istep tge sp rs0 m0 ib rs1 m1 ofin ->
+ iblock_istep tge sp rs0 m0 (normRTL ib) rs1 m1 ofin.
+Proof.
+ intros; eapply normRTLrec_iblock_istep_correct; eauto.
+ destruct ofin; simpl; auto.
+Qed.
+
+Lemma normRTLrec_iblock_istep_run_None tge sp ib:
+ forall rs0 m0 k
+ (CONT: match iblock_istep_run tge sp ib rs0 m0 with
+ | Some (out rs1 m1 ofin) =>
+ ofin = None /\
+ iblock_istep_run tge sp k rs1 m1 = None
+ | _ => True
+ end),
+ iblock_istep_run tge sp (normRTLrec ib k) rs0 m0 = None.
+Proof.
+ induction ib; simpl; intros; subst; intuition (try discriminate).
+ - (* Bnop *)
+ intros. autodestruct; auto.
+ - (* Bop *)
+ intros; repeat autodestruct; simpl; intuition congruence.
+ - (* Bload *)
+ intros; repeat autodestruct; simpl; intuition congruence.
+ - (* Bstore *)
+ intros; repeat autodestruct; simpl; intuition congruence.
+ - (* Bseq *)
+ intros.
+ eapply IHib1; eauto.
+ autodestruct; simpl in *; destruct o; simpl in *; intuition eauto.
+ + destruct _fin; intuition eauto.
+ + destruct _fin; intuition congruence || eauto.
+ - (* Bcond *)
+ intros; repeat autodestruct; simpl; intuition congruence || eauto.
+Qed.
+
+Lemma normRTL_preserves_iblock_istep_run_None tge sp ib:
+ forall rs m, iblock_istep_run tge sp ib rs m = None
+ -> iblock_istep_run tge sp (normRTL ib) rs m = None.
+Proof.
+ intros; eapply normRTLrec_iblock_istep_run_None; eauto.
+ rewrite H; simpl; auto.
+Qed.
+
+Lemma normRTL_preserves_iblock_istep_run tge sp ib:
+ forall rs m, iblock_istep_run tge sp ib rs m =
+ iblock_istep_run tge sp (normRTL ib) rs m.
+Proof.
+ intros.
+ destruct (iblock_istep_run tge sp ib rs m) eqn:ISTEP.
+ - destruct o. symmetry.
+ rewrite <- iblock_istep_run_equiv in *.
+ apply normRTL_iblock_istep_correct; auto.
+ - symmetry.
+ apply normRTL_preserves_iblock_istep_run_None; auto.
+Qed.
+
+Local Hint Constructors match_iblock: core.
+Lemma normRTLrec_matchiblock_correct dupmap cfg ib pc isfst:
+ forall opc1
+ (MIB: match_iblock dupmap cfg isfst pc ib opc1) k opc2
+ (CONT: match opc1 with
+ | Some pc' =>
+ match_iblock dupmap cfg false pc' k opc2
+ | None => opc2=opc1
+ end),
+ match_iblock dupmap cfg isfst pc (normRTLrec ib k) opc2.
+Proof.
+ induction 1; simpl; intros; subst; eauto.
+ (* Bcond *)
+ intros. inv H0;
+ econstructor; eauto; try econstructor.
+ destruct opc0; econstructor.
+Qed.
+
+Lemma normRTL_matchiblock_correct dupmap cfg ib pc isfst opc:
+ match_iblock dupmap cfg isfst pc ib opc ->
+ match_iblock dupmap cfg isfst pc (normRTL ib) opc.
+Proof.
+ intros.
+ eapply normRTLrec_matchiblock_correct; eauto.
+ destruct opc; simpl; auto.
+Qed.
+
+Lemma is_normRTL_correct dupmap cfg ib pc
+ (MI : match_iblock dupmap cfg true pc ib None):
+ is_normRTL (normRTL ib).
+Proof.
+ exploit normRTL_matchiblock_correct; eauto.
+ intros MI2.
+ eapply is_wnormRTL_normRTL; eauto.
+ apply normRTL_wcorrect; try congruence.
+ inv MI2; discriminate.
+Qed.
+
+(** * Matching relation on functions *)
+
+(* we simply switch [f] and [tf] in the usual way *)
+Record match_function dupmap (f:RTL.function) (tf: BTL.function): Prop := {
+ matchRTL:> BTLmatchRTL.match_function dupmap tf f;
+ liveness_ok: liveness_ok_function tf;
+}.
+
+Local Hint Resolve matchRTL: core.
+
+Inductive match_fundef: RTL.fundef -> BTL.fundef -> Prop :=
+ | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f')
+ | match_External ef: match_fundef (External ef) (External ef).
+
+Inductive match_stackframes: RTL.stackframe -> BTL.stackframe -> Prop :=
+ | match_stackframe_intro
+ dupmap res f sp pc rs f' pc'
+ (TRANSF: match_function dupmap f f')
+ (DUPLIC: dupmap!pc' = Some pc)
+ : match_stackframes (RTL.Stackframe res f sp pc rs) (BTL.Stackframe res f' sp pc' rs).
+
+Lemma transf_function_correct f f':
+ transf_function f = OK f' -> exists dupmap, match_function dupmap f f'.
+Proof.
+ unfold transf_function; unfold bind. repeat autodestruct.
+ intros LIVE VER _ _ X. inv X. eexists; econstructor.
+ - eapply verify_function_correct; simpl; eauto.
+ - unfold liveness_ok_function; destruct u0; auto.
+Qed.
+
+Lemma transf_fundef_correct f f':
+ transf_fundef f = OK f' -> match_fundef f f'.
+Proof.
+ intros TRANSF; destruct f; simpl; monadInv TRANSF.
+ + exploit transf_function_correct; eauto.
+ intros (dupmap & MATCH_F).
+ eapply match_Internal; eauto.
+ + eapply match_External.
+Qed.
+
+Definition match_prog (p: RTL.program) (tp: program) :=
+ match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
+
+Lemma transf_program_match:
+ forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
+Proof.
+ intros. eapply match_transform_partial_program_contextual; eauto.
+Qed.
+
+Section BTL_SIMULATES_RTL.
+
+Variable prog: RTL.program.
+Variable tprog: program.
+
+Hypothesis TRANSL: match_prog prog tprog.
+
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Local Open Scope nat_scope.
+
+(** * Match relation from a RTL state to a BTL state
+
+The "option iblock" parameter represents the current BTL execution state.
+Thus, each RTL single step is symbolized by a new BTL "option iblock"
+starting at the equivalent PC.
+
+The simulation diagram for match_states_intro is as follows:
+
+<<
+
+ RTL state match_states_intro BTL state
+ [pcR0,rs0,m0] --------------------------- [pcB0,rs0,m0]
+ | |
+ | |
+ RTL_RUN | *E0 | BTL_RUN
+ | |
+ | MIB |
+ [pcR1,rs1,m1] ------------------------------- [ib]
+
+>>
+*)
+
+Inductive match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst: Prop :=
+ | match_strong_state_intro
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_function dupmap f f')
+ (ATpc0: (fn_code f')!pcB0 = Some ib0)
+ (DUPLIC: dupmap!pcB0 = Some pcR0)
+ (MIB: match_iblock dupmap (RTL.fn_code f) isfst pcR1 ib None)
+ (IS_EXPD: is_normRTL ib)
+ (RTL_RUN: star RTL.step ge (RTL.State st f sp pcR0 rs0 m0) E0 (RTL.State st f sp pcR1 rs1 m1))
+ (BTL_RUN: iblock_istep_run tge sp ib0.(entry) rs0 m0 = iblock_istep_run tge sp ib rs1 m1)
+ : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst
+ .
+
+Inductive match_states: (option iblock) -> RTL.state -> state -> Prop :=
+ | match_states_intro
+ dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst
+ (MSTRONG: match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst)
+ (NGOTO: is_goto ib = false)
+ : match_states (Some ib) (RTL.State st f sp pcR1 rs1 m1) (State st' f' sp pcB0 rs0 m0)
+ | match_states_call
+ st st' f f' args m
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: match_fundef f f')
+ : match_states None (RTL.Callstate st f args m) (Callstate st' f' args m)
+ | match_states_return
+ st st' v m
+ (STACKS: list_forall2 match_stackframes st st')
+ : match_states None (RTL.Returnstate st v m) (Returnstate st' v m)
+ .
+
+Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof.
+ rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
+Qed.
+
+Lemma senv_preserved: Senv.equiv ge tge.
+Proof.
+ eapply (Genv.senv_match TRANSL).
+Qed.
+
+Lemma functions_translated (v: val) (f: RTL.fundef):
+ Genv.find_funct ge v = Some f ->
+ exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tge v = Some tf /\ linkorder cunit prog.
+Proof.
+ intros. exploit (Genv.find_funct_match TRANSL); eauto.
+ intros (cu & tf & A & B & C).
+ repeat eexists; intuition eauto.
+ + unfold incl; auto.
+ + eapply linkorder_refl.
+Qed.
+
+Lemma function_ptr_translated v f:
+ Genv.find_funct_ptr ge v = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = OK tf.
+Proof.
+ intros.
+ exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
+Qed.
+
+Lemma function_sig_translated f tf: transf_fundef f = OK tf -> funsig tf = RTL.funsig f.
+Proof.
+ intros H; apply transf_fundef_correct in H; destruct H; simpl; eauto.
+ erewrite preserv_fnsig; eauto.
+Qed.
+
+Lemma transf_initial_states s1:
+ RTL.initial_state prog s1 ->
+ exists ib s2, initial_state tprog s2 /\ match_states ib s1 s2.
+Proof.
+ intros. inv H.
+ exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF).
+ eexists. eexists. split.
+ - econstructor; eauto.
+ + eapply (Genv.init_mem_transf_partial TRANSL); eauto.
+ + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
+ symmetry. eapply match_program_main. eauto.
+ + erewrite function_sig_translated; eauto.
+ - constructor; eauto.
+ constructor.
+ apply transf_fundef_correct; auto.
+Qed.
+
+Lemma transf_final_states ib s1 s2 r:
+ match_states ib s1 s2 -> RTL.final_state s1 r -> final_state s2 r.
+Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+Qed.
+
+Lemma find_function_preserved ri rs0 fd
+ (FIND : RTL.find_function ge ri rs0 = Some fd)
+ : exists fd', find_function tge ri rs0 = Some fd'
+ /\ transf_fundef fd = OK fd'.
+Proof.
+ pose symbols_preserved as SYMPRES.
+ destruct ri.
+ + simpl in FIND; apply functions_translated in FIND.
+ destruct FIND as (tf & cunit & TFUN & GFIND & LO).
+ eexists; split. eauto. assumption.
+ + simpl in FIND. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate.
+ apply function_ptr_translated in FIND. destruct FIND as (tf & GFF & TF).
+ eexists; split. simpl. rewrite symbols_preserved.
+ rewrite GFS. eassumption. assumption.
+Qed.
+
+(** Representing an intermediate BTL state
+
+We keep a measure of code that remains to be executed with the omeasure
+type defined below. Intuitively, each RTL step corresponds to either
+ - a single BTL step if we are on the last instruction of the block
+ - no BTL step (as we use a "big step" semantics) but a change in
+ the measure which represents the new intermediate state of the BTL code
+ *)
+Fixpoint measure ib: nat :=
+ match ib with
+ | Bseq ib1 ib2
+ | Bcond _ _ ib1 ib2 _ => measure ib1 + measure ib2
+ | ib => 1
+ end.
+
+Definition omeasure (oib: option iblock): nat :=
+ match oib with
+ | None => 0
+ | Some ib => measure ib
+ end.
+
+Remark measure_pos: forall ib,
+ measure ib > 0.
+Proof.
+ induction ib; simpl; auto; lia.
+Qed.
+
+Lemma match_iblock_true_isnt_goto dupmap cfg pc ib opc:
+ match_iblock dupmap cfg true pc ib opc ->
+ is_goto ib = false.
+Proof.
+ intros MIB; inversion MIB as [d1 d2 d3 d4 d5 H H0| | | | | | | |]; subst; simpl; try congruence.
+ inv H0; congruence.
+Qed.
+
+Local Hint Resolve match_iblock_true_isnt_goto normRTL_preserves_iblock_istep_run star_refl star_right: core.
+Local Hint Constructors match_strong_state RTL.step: core.
+
+(** At entry in a block: we init [match_states] on [normRTL] to normalize the block *)
+Lemma match_states_entry dupmap st f sp pc ib rs m st' f' pc'
+ (STACKS : list_forall2 match_stackframes st st')
+ (TRANSF : match_function dupmap f f')
+ (FN : (fn_code f') ! pc' = Some ib)
+ (MI : match_iblock dupmap (RTL.fn_code f) true pc (entry ib) None)
+ (DUP : dupmap ! pc' = Some pc):
+ match_states (Some (normRTL (entry ib))) (RTL.State st f sp pc rs m) (State st' f' sp pc' rs m).
+Proof.
+ exploit is_normRTL_correct; eauto.
+ econstructor; eauto; apply normRTL_matchiblock_correct in MI; eauto.
+Qed.
+Local Hint Resolve match_states_entry: core.
+
+Lemma list_nth_z_rev_dupmap:
+ forall dupmap ln ln' (pc pc': node) val,
+ list_nth_z ln val = Some pc ->
+ list_forall2 (fun n' n => dupmap!n' = Some n) ln' ln ->
+ exists (pc': node),
+ list_nth_z ln' val = Some pc'
+ /\ dupmap!pc' = Some pc.
+Proof.
+ induction ln; intros until val; intros LNZ LFA.
+ - inv LNZ.
+ - inv LNZ. destruct (zeq val 0) eqn:ZEQ.
+ + inv H0. destruct ln'; inv LFA.
+ simpl. exists p. split; auto.
+ + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto.
+ intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence.
+Qed.
+
+
+(** * Match strong state property
+
+Used when executing non-atomic instructions such as Bseq/Bcond(ib1,ib2).
+Two possible executions:
+
+<<
+
+ **ib2 is a Bgoto (left side):**
+
+ RTL state MSS1 BTL state
+ [pcR1,rs1,m1] -------------------------- [ib1,pcB0,rs0,m0]
+ | |
+ | |
+ | | BTL_STEP
+ | |
+ | |
+ RTL_STEP | *E0 [ib2,pc=(Bgoto succ),rs2,m2]
+ | / |
+ | MSS2 / |
+ | _________________/ | BTL_GOTO
+ | / |
+ | / GOAL: match_states |
+ [pcR2,rs2,m2] ------------------------ [ib?,pc=succ,rs2,m2]
+
+
+ **ib2 is any other instruction (right side):**
+
+See explanations of opt_simu below.
+
+>>
+*)
+
+Lemma match_strong_state_simu
+ dupmap st st' f f' sp rs2 m2 rs1 m1 rs0 m0 pcB0 pcR0 pcR1 pcR2 isfst ib1 ib2 ib0 n t s1'
+ (EQt: t=E0)
+ (EQs1': s1'=(RTL.State st f sp pcR2 rs2 m2))
+ (STEP : RTL.step ge (RTL.State st f sp pcR1 rs1 m1) t s1')
+ (MSS1 : match_strong_state dupmap st st' f f' sp rs1 m1 rs0 m0 pcB0 pcR0 pcR1 ib1 ib0 isfst)
+ (MSS2 : match_strong_state dupmap st st' f f' sp rs2 m2 rs0 m0 pcB0 pcR0 pcR2 ib2 ib0 false)
+ (MES : measure ib2 < n)
+ : exists (oib' : option iblock),
+ (exists s2', step tid tge (State st' f' sp pcB0 rs0 m0) E0 s2'
+ /\ match_states oib' s1' s2')
+ \/ (omeasure oib' < n /\ t=E0
+ /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)).
+Proof.
+ subst.
+ destruct (is_goto ib2) eqn:GT.
+ destruct ib2; try destruct fi; try discriminate.
+ - (* Bgoto *)
+ inv MSS2. inversion MIB; subst; try inv H4.
+ remember H2 as ODUPLIC; clear HeqODUPLIC.
+ exploit dupmap_correct; eauto.
+ intros [ib [FNC MI]].
+ eexists; left; eexists; split; eauto.
+ repeat econstructor; eauto.
+ apply iblock_istep_run_equiv in BTL_RUN; eauto.
+ econstructor.
+ - (* Others *)
+ exists (Some ib2); right; split.
+ simpl; auto.
+ split; auto. econstructor; eauto.
+Qed.
+
+Lemma opt_simu_intro
+ dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst s1' t
+ (STEP : RTL.step ge (RTL.State st f sp pcR1 rs m) t s1')
+ (MSTRONG : match_strong_state dupmap st st' f f' sp rs m rs0 m0 pcB0 pcR0 pcR1 ib ib0 isfst)
+ (NGOTO : is_goto ib = false)
+ : exists (oib' : option iblock),
+ (exists s2', step tid tge (State st' f' sp pcB0 rs0 m0) t s2' /\ match_states oib' s1' s2')
+ \/ (omeasure oib' < omeasure (Some ib) /\ t=E0 /\ match_states oib' s1' (State st' f' sp pcB0 rs0 m0)).
+Proof.
+ inv MSTRONG; subst. inv MIB.
+ - (* mib_BF *)
+ inv H0;
+ inversion STEP; subst; try_simplify_someHyps; intros.
+ + (* Breturn *)
+ eexists; left; eexists; split.
+ * econstructor; eauto. econstructor.
+ eexists; eexists; split.
+ eapply iblock_istep_run_equiv in BTL_RUN.
+ eapply BTL_RUN. econstructor; eauto.
+ erewrite preserv_fnstacksize; eauto.
+ * econstructor; eauto.
+ + (* Bcall *)
+ rename H10 into FIND.
+ eapply find_function_preserved in FIND.
+ destruct FIND as (fd' & FF & TRANSFUN).
+ eexists; left; eexists; split.
+ * econstructor; eauto. econstructor.
+ eexists; eexists; split.
+ eapply iblock_istep_run_equiv in BTL_RUN.
+ eapply BTL_RUN. econstructor; eauto.
+ eapply function_sig_translated; eauto.
+ * repeat (econstructor; eauto).
+ eapply transf_fundef_correct; eauto.
+ + (* Btailcall *)
+ rename H9 into FIND.
+ eapply find_function_preserved in FIND.
+ destruct FIND as (fd' & FF & TRANSFUN).
+ eexists; left; eexists; split.
+ * econstructor; eauto. econstructor.
+ eexists; eexists; split.
+ eapply iblock_istep_run_equiv in BTL_RUN.
+ eapply BTL_RUN. econstructor; eauto.
+ eapply function_sig_translated; eauto.
+ erewrite preserv_fnstacksize; eauto.
+ * repeat (econstructor; eauto).
+ eapply transf_fundef_correct; eauto.
+ + (* Bbuiltin *)
+ exploit dupmap_correct; eauto.
+ intros [ib [FNC MI]].
+ exists (Some (normRTL (entry ib))); left; eexists; split; eauto.
+ econstructor; eauto. econstructor.
+ eexists; eexists; split.
+ eapply iblock_istep_run_equiv in BTL_RUN.
+ eapply BTL_RUN. econstructor; eauto.
+ pose symbols_preserved as SYMPRES.
+ eapply eval_builtin_args_preserved; eauto.
+ eapply external_call_symbols_preserved; eauto. eapply senv_preserved.
+ + (* Bjumptable *)
+ exploit list_nth_z_rev_dupmap; eauto.
+ intros (pc'0 & LNZ & DM).
+ exploit dupmap_correct; eauto.
+ intros [ib [FNC MI]].
+ exists (Some (normRTL (entry ib))); left; eexists; split; eauto.
+ econstructor; eauto. econstructor.
+ eexists; eexists; split.
+ eapply iblock_istep_run_equiv in BTL_RUN.
+ eapply BTL_RUN. econstructor; eauto.
+ - (* mib_exit *)
+ discriminate.
+ - (* mib_seq *)
+ inv IS_EXPD; try discriminate.
+ inv H; simpl in *; try congruence;
+ inv STEP; try_simplify_someHyps;
+ intros; eapply match_strong_state_simu; eauto;
+ econstructor; eauto.
+ { (* Bop *)
+ erewrite eval_operation_preserved in H12.
+ erewrite H12 in BTL_RUN; simpl in BTL_RUN; auto.
+ intros; rewrite <- symbols_preserved; trivial. }
+ all: (* Bload / Bstore *)
+ erewrite eval_addressing_preserved in H12;
+ try erewrite H12 in BTL_RUN; try erewrite H13 in BTL_RUN;
+ simpl in BTL_RUN; try destruct trap; auto;
+ intros; rewrite <- symbols_preserved; trivial.
+ - (* mib_cond *)
+ inv IS_EXPD; try discriminate.
+ inversion STEP; subst; try_simplify_someHyps; intros.
+ destruct (is_join_opt_None opc1 opc2); eauto. subst.
+ eapply match_strong_state_simu with (ib1:=Bcond c lr bso bnot iinfo) (ib2:=(if b then bso else bnot)); eauto.
+ + intros; rewrite H14 in BTL_RUN; destruct b; econstructor; eauto.
+ + assert (measure (if b then bnot else bso) > 0) by apply measure_pos; destruct b; simpl; lia.
+ Unshelve.
+ all: eauto.
+Qed.
+
+(** * Main RTL to BTL simulation theorem
+
+Two possible executions:
+
+<<
+
+ **Last instruction (left side):**
+
+ RTL state match_states BTL state
+ s1 ------------------------------------ s2
+ | |
+ STEP | Classical lockstep simu |
+ | |
+ s1' ----------------------------------- s2'
+
+
+ **Middle instruction (right side):**
+
+ RTL state match_states [oib] BTL state
+ s1 ------------------------------------ s2
+ | _______/
+ STEP | *E0 ___________________/
+ | / match_states [oib']
+ s1' ______/
+ Where omeasure oib' < omeasure oib
+
+>>
+*)
+
+Theorem opt_simu s1 t s1' oib s2:
+ RTL.step ge s1 t s1' ->
+ match_states oib s1 s2 ->
+ exists (oib' : option iblock),
+ (exists s2', step tid tge s2 t s2' /\ match_states oib' s1' s2')
+ \/ (omeasure oib' < omeasure oib /\ t=E0 /\ match_states oib' s1' s2)
+ .
+Proof.
+ inversion 2; subst; clear H0.
+ - (* State *)
+ exploit opt_simu_intro; eauto.
+ - (* Callstate *)
+ inv H.
+ + (* Internal function *)
+ inv TRANSF.
+ rename H0 into TRANSF.
+ exploit dupmap_entrypoint; eauto. intros ENTRY.
+ exploit dupmap_correct; eauto.
+ intros [ib [CENTRY MI]].
+ exists (Some (normRTL (entry ib))); left; eexists; split.
+ * eapply exec_function_internal.
+ erewrite preserv_fnstacksize; eauto.
+ * erewrite preserv_fnparams; eauto.
+ + (* External function *)
+ inv TRANSF.
+ eexists; left; eexists; split.
+ * eapply exec_function_external.
+ eapply external_call_symbols_preserved.
+ eapply senv_preserved. eauto.
+ * econstructor; eauto.
+ - (* Returnstate *)
+ inv H. inv STACKS. inv H1.
+ exploit dupmap_correct; eauto.
+ intros [ib [FNC MI]].
+ eexists; left; eexists; split; eauto.
+ eapply exec_return.
+Qed.
+
+Local Hint Resolve plus_one star_refl: core.
+
+Theorem transf_program_correct_cfg:
+ forward_simulation (RTL.semantics prog) (BTLmatchRTL.cfgsem tprog).
+Proof.
+ eapply (Forward_simulation (L1:=RTL.semantics prog) (L2:=cfgsem tprog) (ltof _ omeasure) match_states).
+ constructor 1; simpl.
+ - apply well_founded_ltof.
+ - eapply transf_initial_states.
+ - eapply transf_final_states.
+ - intros s1 t s1' STEP i s2 MATCH. exploit opt_simu; eauto. clear MATCH STEP.
+ destruct 1 as (oib' & [ (s2' & STEP & MATCH) | (MEASURE & TRACE & MATCH) ]).
+ + repeat eexists; eauto.
+ + subst. repeat eexists; eauto.
+ - eapply senv_preserved.
+Qed.
+
+Theorem all_fundef_liveness_ok b f:
+ Genv.find_funct_ptr tge b = Some f -> liveness_ok_fundef f.
+Proof.
+ unfold match_prog, match_program in TRANSL.
+ unfold Genv.find_funct_ptr, tge; simpl; intro X.
+ destruct (Genv.find_def_match_2 TRANSL b) as [|f0 y H]; try congruence.
+ destruct y as [tf0|]; try congruence.
+ inversion X as [H1]. subst. clear X.
+ remember (@Gfun fundef unit f) as f2.
+ destruct H as [ctx' f1 f2 H0|]; try congruence.
+ inversion Heqf2 as [H2]. subst; clear Heqf2.
+ exploit transf_fundef_correct; eauto.
+ destruct f; econstructor.
+ inv H1; eapply liveness_ok; eauto.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (RTL.semantics prog) (BTL.fsem tprog).
+Proof.
+ eapply compose_forward_simulations.
+ - eapply transf_program_correct_cfg.
+ - eapply cfgsem2fsem. apply all_fundef_liveness_ok.
+Qed.
+
+End BTL_SIMULATES_RTL.