aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling
diff options
context:
space:
mode:
Diffstat (limited to 'scheduling')
-rw-r--r--scheduling/BTL.v743
-rw-r--r--scheduling/BTLRenumber.ml112
-rw-r--r--scheduling/BTLScheduleraux.ml344
-rw-r--r--scheduling/BTL_Livecheck.v690
-rw-r--r--scheduling/BTL_SEimpl.v1528
-rw-r--r--scheduling/BTL_SEsimuref.v849
-rw-r--r--scheduling/BTL_SEtheory.v1334
-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.v413
-rw-r--r--scheduling/BTLtypes.ml31
-rw-r--r--scheduling/InstructionScheduler.ml1753
-rw-r--r--scheduling/InstructionScheduler.mli135
-rw-r--r--scheduling/PrintBTL.ml118
-rw-r--r--scheduling/RTLtoBTL.v27
-rw-r--r--scheduling/RTLtoBTLaux.ml119
-rw-r--r--scheduling/RTLtoBTLproof.v762
-rw-r--r--scheduling/abstractbb/AbstractBasicBlocksDef.v471
-rw-r--r--scheduling/abstractbb/ImpSimuTest.v1286
-rw-r--r--scheduling/abstractbb/Parallelizability.v792
-rw-r--r--scheduling/abstractbb/README.md12
-rw-r--r--scheduling/abstractbb/SeqSimuTheory.v397
-rw-r--r--scheduling/postpass_lib/ForwardSimulationBlock.v387
-rw-r--r--scheduling/postpass_lib/Machblock.v388
-rw-r--r--scheduling/postpass_lib/Machblockgen.v221
-rw-r--r--scheduling/postpass_lib/Machblockgenproof.v810
31 files changed, 15593 insertions, 0 deletions
diff --git a/scheduling/BTL.v b/scheduling/BTL.v
new file mode 100644
index 00000000..b3895768
--- /dev/null
+++ b/scheduling/BTL.v
@@ -0,0 +1,743 @@
+(** 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.
+
+Local Hint Constructors has_loaded: core.
+
+(** 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 ge 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 ge 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.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..38bc685c
--- /dev/null
+++ b/scheduling/BTLScheduleraux.ml
@@ -0,0 +1,344 @@
+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.opt_info 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
+ if get_some @@ e.opt_info 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 || not !Clflags.option_fnontrap_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) ->
+ let _opt_info =
+ match t with
+ | AST.TRAP -> Some true
+ | AST.NOTRAP -> Some false in
+ e.opt_info <- _opt_info;
+ 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.opt_info 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..0cbc46e6
--- /dev/null
+++ b/scheduling/BTL_SEimpl.v
@@ -0,0 +1,1528 @@
+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.
+
+(** 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 (ris_sreg 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.
+ simplify_option.
+ - 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_option. 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.
+ simplify_option.
+ + 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' (ris_sreg hrs)
+ else PTree.set r sv (ris_sreg hrs)
+ | false, Sundef _ =>
+ PTree.remove r (ris_sreg hrs)
+ | _, _ => PTree.set r sv (ris_sreg 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 (ris_sreg 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.
+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.
+ simplify_option.
+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. 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; try (eapply CONT; eauto); fail).
+ - (* BF *)
+ wlp_simplify; exploit tr_hrs_correct; eauto.
+ intros; constructor; auto.
+ intros X; apply H0 in X.
+ exploit hrexec_final_sfv_correct; 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;
+ simplify_option.
+ + 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; simplify_option.
+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; simplify_option.
+ }
+ 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;
+ simplify_option.
+ + 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.tree' A): ?? unit :=
+ match d1, d2 with
+ | Node001 r1, Node001 r2 => PTree_eq_check' r1 r2
+ | Node010 x1, Node010 x2 =>
+ phys_check x1 x2 "PTree_eq_check': data physically differ"
+ | Node011 x1 r1, Node011 x2 r2 =>
+ phys_check x1 x2 "PTree_eq_check': data physically differ" ;;
+ PTree_eq_check' r1 r2
+ | Node100 l1, Node100 l2 => PTree_eq_check' l1 l2
+ | Node101 l1 r1, Node101 l2 r2 =>
+ PTree_eq_check' l1 l2 ;;
+ PTree_eq_check' r1 r2
+ | Node110 l1 x1, Node110 l2 x2 =>
+ phys_check x1 x2 "PTree_eq_check': data physically differ" ;;
+ PTree_eq_check' l1 l2
+ | Node111 l1 x1 r1, Node111 l2 x2 r2 =>
+ phys_check x1 x2 "PTree_eq_check': data physically differ" ;;
+ PTree_eq_check' l1 l2 ;;
+ PTree_eq_check' r1 r2
+ | _, _ => FAILWITH "PTree_eq_check': different shapes"
+ end.
+
+Lemma PTree_eq_check'_correct A : forall (d1 d2: PTree.tree' A),
+ WHEN PTree_eq_check' d1 d2 ~> _ THEN d1 = d2.
+Proof.
+ induction d1; destruct d2; cbn; wlp_simplify; try congruence.
+Qed.
+
+Definition PTree_eq_check {A} (d1 d2: PTree.t A): ?? unit :=
+ match d1, d2 with
+ | Empty, Empty => RET tt
+ | Nodes m1, Nodes m2 => PTree_eq_check' m1 m2
+ | _, _ => FAILWITH "PTree_eq_check: empty vs nonempty"
+ end.
+
+Lemma PTree_eq_check_correct' A : forall (d1 d2: PTree.tree A),
+ WHEN PTree_eq_check d1 d2 ~> _ THEN d1 = d2.
+Proof.
+ Local Hint Resolve PTree_eq_check'_correct: wlp.
+ destruct d1 as [ | m1]; destruct d2 as [ | m2]; cbn; wlp_simplify.
+ congruence.
+Qed.
+
+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.
+ Local Hint Resolve PTree_eq_check_correct': wlp.
+ intros.
+ wlp_simplify.
+ congruence.
+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..51b562c7
--- /dev/null
+++ b/scheduling/BTL_SEsimuref.v
@@ -0,0 +1,849 @@
+(** 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_sreg 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.
+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.
+ + 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..55afc19a
--- /dev/null
+++ b/scheduling/BTL_SEtheory.v
@@ -0,0 +1,1334 @@
+(** 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 _ =>
+ SOME args <- eval_list_sval ctx lsv IN
+ SOME m <- eval_smem ctx sm IN
+ match trap with
+ | TRAP =>
+ SOME a <- eval_addressing (cge ctx) (csp ctx) addr args IN
+ Mem.loadv chunk m a
+ | NOTRAP =>
+ match eval_addressing (cge ctx) (csp ctx) addr args with
+ | None => Some Vundef
+ | Some a =>
+ 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.
+ rewrite MEM; simpl. autodestruct. rewrite 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.
+ rewrite MEM; simpl; 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.
+ autodestruct; intros.
+ erewrite IHsv; do 2 autodestruct;
+ erewrite eval_addressing_preserved; eauto.
+ + 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..11a7ee7f
--- /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 _opt_info = { inumb = _inumb; opt_info = _opt_info; visited = false; liveins = Regset.empty }
+
+let def_iinfo () = { inumb = undef_node; opt_info = 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..47262b3e
--- /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.opt_info),
+ 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..c342d076
--- /dev/null
+++ b/scheduling/BTLtoRTLproof.v
@@ -0,0 +1,413 @@
+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. inversion LOAD; subst.
+ + try (eapply exec_Iload; eauto; eapply has_loaded_normal; eauto;
+ rewrite <- EVAL; erewrite <- eval_addressing_preserved; eauto;
+ intros; rewrite symbols_preserved; trivial).
+ + destruct (eval_addressing) eqn:EVAL in LOAD0.
+ * specialize (LOAD0 v).
+ eapply exec_Iload; eauto. eapply has_loaded_default; eauto.
+ rewrite eval_addressing_preserved with (ge1:=ge).
+ intros a EVAL'; rewrite EVAL in EVAL'; inv EVAL'. apply LOAD0; auto.
+ intros; rewrite symbols_preserved; trivial.
+ * eapply exec_Iload; eauto. eapply has_loaded_default; eauto.
+ rewrite eval_addressing_preserved with (ge1:=ge).
+ intros a EVAL'; rewrite EVAL in EVAL'; inv EVAL'.
+ 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..64001adf
--- /dev/null
+++ b/scheduling/BTLtypes.ml
@@ -0,0 +1,31 @@
+open Registers
+
+type inst_info = {
+ mutable inumb : int;
+ mutable opt_info : bool option;
+ mutable visited : bool;
+ mutable liveins : Regset.t;
+}
+(** This struct is a ghost field attached to each BTL instruction
+ * inumb: int field used for remembering the original numbering of CFG
+ * opt_info: option bool used for various tests:
+ * - On Bcond, stores the prediction information
+ * - On Bload, stores the trapping information (if Some false, the load was
+ * initially non-trapping, and the opposite if some true)
+ * visited: boolean used to mark nodes
+ * liveins: set of lives registers
+ *)
+
+type block_info = {
+ mutable bnumb : int;
+ mutable visited : bool;
+ s_output_regs : Regset.t;
+ typing : RTLtyping.regenv;
+}
+(** Struct attached to each BTL iblock_info type
+ * bnumb: int representing the block id in the BTL CFG
+ * visited: boolean used to mark blocks
+ * s_output_regs: set of output registers
+ * typing: field transferring RTL typing information in BTL (used in regpres
+ * scheduling)
+ *)
diff --git a/scheduling/InstructionScheduler.ml b/scheduling/InstructionScheduler.ml
new file mode 100644
index 00000000..4c9bde64
--- /dev/null
+++ b/scheduling/InstructionScheduler.ml
@@ -0,0 +1,1753 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+let with_destructor dtor stuff f =
+ try let ret = f stuff in
+ dtor stuff;
+ ret
+ with exn -> dtor stuff;
+ raise exn;;
+
+let with_out_channel chan f = with_destructor close_out chan f;;
+let with_in_channel chan f = with_destructor close_in chan f;;
+
+(** Schedule instructions on a synchronized pipeline
+@author David Monniaux, CNRS, VERIMAG *)
+
+type latency_constraint = {
+ instr_from : int;
+ instr_to : int;
+ latency : int };;
+
+type problem = {
+ max_latency : int;
+ resource_bounds : int array;
+ live_regs_entry : Registers.Regset.t;
+ typing : RTLtyping.regenv;
+ reference_counting : ((Registers.reg, int * int) Hashtbl.t
+ * ((Registers.reg * bool) list array)) option;
+ instruction_usages : int array array;
+ latency_constraints : latency_constraint list;
+ };;
+
+let print_problem channel problem =
+ (if problem.max_latency >= 0
+ then Printf.fprintf channel "max makespan: %d\n" problem.max_latency);
+ output_string channel "resource bounds:";
+ (Array.iter (fun b -> Printf.fprintf channel " %d" b) problem.resource_bounds);
+ output_string channel ";\n";
+ (Array.iteri (fun i v ->
+ Printf.fprintf channel "instr%d:" i;
+ (Array.iter (fun b -> Printf.fprintf channel " %d" b) v);
+ output_string channel ";\n") problem.instruction_usages);
+ List.iter (fun instr ->
+ Printf.printf "t%d - t%d >= %d;\n"
+ instr.instr_to instr.instr_from instr.latency)
+ problem.latency_constraints;;
+
+let get_nr_instructions problem = Array.length problem.instruction_usages;;
+let get_nr_resources problem = Array.length problem.resource_bounds;;
+
+type solution = int array
+type scheduler = problem -> solution option
+
+(* DISABLED
+(** Schedule the problem optimally by constraint solving using the Gecode solver. *)
+external gecode_scheduler : problem -> solution option =
+ "caml_gecode_schedule_instr";;
+ *)
+
+let maximum_slot_used times =
+ let maxi = ref (-1) in
+ for i=0 to (Array.length times)-2
+ do
+ maxi := max !maxi times.(i)
+ done;
+ !maxi;;
+
+let check_schedule (problem : problem) (times : solution) =
+ let nr_instructions = get_nr_instructions problem in
+ (if Array.length times <> nr_instructions+1
+ then failwith
+ (Printf.sprintf "check_schedule: %d times expected, got %d"
+ (nr_instructions+1) (Array.length times)));
+ (if problem.max_latency >= 0 && times.(nr_instructions)> problem.max_latency
+ then failwith "check_schedule: max_latency exceeded");
+ (Array.iteri (fun i time ->
+ (if time < 0
+ then failwith (Printf.sprintf "time[%d] < 0" i))) times);
+ let slot_resources = Array.init ((maximum_slot_used times)+1)
+ (fun _ -> Array.copy problem.resource_bounds) in
+ for i=0 to nr_instructions -1
+ do
+ let remaining_resources = slot_resources.(times.(i))
+ and used_resources = problem.instruction_usages.(i) in
+ for resource=0 to (Array.length used_resources)-1
+ do
+ let after = remaining_resources.(resource) - used_resources.(resource) in
+ (if after < 0
+ then failwith (Printf.sprintf "check_schedule: instruction %d exceeds resource %d at slot %d" i resource times.(i)));
+ remaining_resources.(resource) <- after
+ done
+ done;
+ List.iter (fun ctr ->
+ if times.(ctr.instr_to) - times.(ctr.instr_from) < ctr.latency
+ then failwith (Printf.sprintf "check_schedule: time[%d]=%d - time[%d]=%d < %d"
+ ctr.instr_to times.(ctr.instr_to)
+ ctr.instr_from times.(ctr.instr_from)
+ ctr.latency)
+ ) problem.latency_constraints;;
+
+let bound_max_time problem =
+ let total = ref(Array.length problem.instruction_usages) in
+ List.iter (fun ctr -> total := !total + ctr.latency) problem.latency_constraints;
+ !total;;
+
+let vector_less_equal a b =
+ try
+ Array.iter2 (fun x y ->
+ if x>y
+ then raise Exit) a b;
+ true
+ with Exit -> false;;
+
+(* let vector_add a b =
+ * assert ((Array.length a) = (Array.length b));
+ * for i=0 to (Array.length a)-1
+ * do
+ * b.(i) <- b.(i) + a.(i)
+ * done;; *)
+
+let vector_subtract a b =
+ assert ((Array.length a) = (Array.length b));
+ for i=0 to (Array.length a)-1
+ do
+ b.(i) <- b.(i) - a.(i)
+ done;;
+
+(* The version with critical path ordering is much better! *)
+type list_scheduler_order =
+ | INSTRUCTION_ORDER
+ | CRITICAL_PATH_ORDER;;
+
+let int_max (x : int) (y : int) =
+ if x > y then x else y;;
+
+let int_min (x : int) (y : int) =
+ if x < y then x else y;;
+
+let get_predecessors problem =
+ let nr_instructions = get_nr_instructions problem in
+ let predecessors = Array.make (nr_instructions+1) [] in
+ List.iter (fun ctr ->
+ predecessors.(ctr.instr_to) <-
+ (ctr.instr_from, ctr.latency)::predecessors.(ctr.instr_to))
+ problem.latency_constraints;
+ predecessors;;
+
+let get_successors problem =
+ let nr_instructions = get_nr_instructions problem in
+ let successors = Array.make nr_instructions [] in
+ List.iter (fun ctr ->
+ successors.(ctr.instr_from) <-
+ (ctr.instr_to, ctr.latency)::successors.(ctr.instr_from))
+ problem.latency_constraints;
+ successors;;
+
+let critical_paths successors =
+ let nr_instructions = Array.length successors in
+ let path_lengths = Array.make nr_instructions (-1) in
+ let rec compute i =
+ if i=nr_instructions then 0 else
+ match path_lengths.(i) with
+ | -2 -> failwith "InstructionScheduler: the dependency graph has cycles"
+ | -1 -> path_lengths.(i) <- -2;
+ let x = List.fold_left
+ (fun cur (j, latency)-> int_max cur (latency+(compute j)))
+ 1 successors.(i)
+ in path_lengths.(i) <- x; x
+ | x -> x
+ in for i = nr_instructions-1 downto 0
+ do
+ ignore (compute i)
+ done;
+ path_lengths;;
+
+let maximum_critical_path problem =
+ let paths = critical_paths (get_successors problem) in
+ Array.fold_left int_max 0 paths;;
+
+let get_earliest_dates predecessors =
+ let nr_instructions = (Array.length predecessors)-1 in
+ let path_lengths = Array.make (nr_instructions+1) (-1) in
+ let rec compute i =
+ match path_lengths.(i) with
+ | -2 -> failwith "InstructionScheduler: the dependency graph has cycles"
+ | -1 -> path_lengths.(i) <- -2;
+ let x = List.fold_left
+ (fun cur (j, latency)-> int_max cur (latency+(compute j)))
+ 0 predecessors.(i)
+ in path_lengths.(i) <- x; x
+ | x -> x
+ in for i = 0 to nr_instructions
+ do
+ ignore (compute i)
+ done;
+ for i = 0 to nr_instructions - 1
+ do
+ path_lengths.(nr_instructions) <- int_max
+ path_lengths.(nr_instructions) (1 + path_lengths.(i))
+ done;
+ path_lengths;;
+
+exception Unschedulable
+
+let get_latest_dates deadline successors =
+ let nr_instructions = Array.length successors
+ and path_lengths = critical_paths successors in
+ Array.init (nr_instructions + 1)
+ (fun i ->
+ if i < nr_instructions then
+ let path_length = path_lengths.(i) in
+ assert (path_length >= 1);
+ (if path_length > deadline
+ then raise Unschedulable);
+ deadline - path_length
+ else deadline);;
+
+let priority_list_scheduler (order : list_scheduler_order)
+ (problem : problem) :
+ solution option =
+ let nr_instructions = get_nr_instructions problem in
+ let successors = get_successors problem
+ and predecessors = get_predecessors problem
+ and times = Array.make (nr_instructions+1) (-1) in
+
+ let priorities = match order with
+ | INSTRUCTION_ORDER -> None
+ | CRITICAL_PATH_ORDER -> Some (critical_paths successors) in
+
+ let module InstrSet =
+ Set.Make (struct type t=int
+ let compare = match priorities with
+ | None -> (fun x y -> x - y)
+ | Some p -> (fun x y ->
+ (match p.(y)-p.(x) with
+ | 0 -> x - y
+ | z -> z))
+ end) in
+
+ let max_time = bound_max_time problem in
+ let ready = Array.make max_time InstrSet.empty in
+ Array.iteri (fun i preds ->
+ if i<nr_instructions && preds=[]
+ then ready.(0) <- InstrSet.add i ready.(0)) predecessors;
+
+ let current_time = ref 0
+ and current_resources = Array.copy problem.resource_bounds
+ and earliest_time i =
+ try
+ let time = ref (-1) in
+ List.iter (fun (j, latency) ->
+ if times.(j) < 0
+ then raise Exit
+ else let t = times.(j) + latency in
+ if t > !time
+ then time := t) predecessors.(i);
+ assert(!time >= 0);
+ !time
+ with Exit -> -1
+ in
+
+ let advance_time() =
+ begin
+ (if !current_time < max_time-1
+ then
+ begin
+ Array.blit problem.resource_bounds 0 current_resources 0
+ (Array.length current_resources);
+ ready.(!current_time + 1) <-
+ InstrSet.union (ready.(!current_time))
+ (ready.(!current_time + 1));
+ ready.(!current_time) <- InstrSet.empty;
+ end);
+ incr current_time
+ end in
+
+ let attempt_scheduling ready usages =
+ let result = ref (-1) in
+ try
+ InstrSet.iter (fun i ->
+ (* Printf.printf "trying scheduling %d\n" i;
+ pr int_vector usages.(i);
+ print _vector current_resources; *)
+ if vector_less_equal usages.(i) current_resources
+ then
+ begin
+ vector_subtract usages.(i) current_resources;
+ result := i;
+ raise Exit
+ end) ready;
+ -1
+ with Exit -> !result in
+
+ while !current_time < max_time
+ do
+ if (InstrSet.is_empty ready.(!current_time))
+ then advance_time()
+ else
+ match attempt_scheduling ready.(!current_time)
+ problem.instruction_usages with
+ | -1 -> advance_time()
+ | i ->
+ begin
+ assert(times.(i) < 0);
+ times.(i) <- !current_time;
+ ready.(!current_time) <- InstrSet.remove i (ready.(!current_time));
+ List.iter (fun (instr_to, latency) ->
+ if instr_to < nr_instructions then
+ match earliest_time instr_to with
+ | -1 -> ()
+ | to_time ->
+ ready.(to_time) <- InstrSet.add instr_to ready.(to_time))
+ successors.(i);
+ successors.(i) <- []
+ end
+ done;
+ try
+ let final_time = ref (-1) in
+ for i=0 to nr_instructions-1
+ do
+ (if times.(i) < 0 then raise Exit);
+ (if !final_time < times.(i)+1 then final_time := times.(i)+1)
+ done;
+ List.iter (fun (i, latency) ->
+ let target_time = latency + times.(i) in
+ if target_time > !final_time
+ then final_time := target_time
+ ) predecessors.(nr_instructions);
+ times.(nr_instructions) <- !final_time;
+ Some times
+ with Exit -> None;;
+
+let list_scheduler = priority_list_scheduler CRITICAL_PATH_ORDER;;
+
+(* dummy code for placating ocaml's warnings *)
+let _ = fun x -> priority_list_scheduler INSTRUCTION_ORDER x;;
+
+
+(* A scheduler sensitive to register pressure *)
+let reg_pres_scheduler (problem : problem) : solution option =
+ (* DebugPrint.debug_flag := true; *)
+ let nr_instructions = get_nr_instructions problem in
+
+ if !Clflags.option_debug_compcert > 6 then
+ (Printf.eprintf "\nSCHEDULING_SUPERBLOCK %d\n" nr_instructions;
+ flush stderr);
+
+ let successors = get_successors problem
+ and predecessors = get_predecessors problem
+ and times = Array.make (nr_instructions+1) (-1) in
+ let live_regs_entry = problem.live_regs_entry in
+
+ let available_regs = Array.copy Machregsaux.nr_regs in
+
+ let nr_types_regs = Array.length available_regs in
+
+ let thres = Array.fold_left (min)
+ (max !(Clflags.option_regpres_threshold) 0)
+ Machregsaux.nr_regs
+ in
+
+
+ let regs_thresholds = Array.make nr_types_regs thres in
+ (* placeholder value *)
+
+ let class_r r =
+ Machregsaux.class_of_type (problem.typing r) in
+
+ let live_regs = Hashtbl.create 42 in
+
+ List.iter (fun r -> let classe = Machregsaux.class_of_type
+ (problem.typing r) in
+ available_regs.(classe)
+ <- available_regs.(classe) - 1;
+ Hashtbl.add live_regs r classe)
+ (Registers.Regset.elements live_regs_entry);
+
+ let csr_b = ref false in
+
+ let counts, mentions =
+ match problem.reference_counting with
+ | Some (l, r) -> l, r
+ | None -> assert false
+ in
+
+ let fold_delta i = (fun a (r, b) ->
+ a +
+ if class_r r <> i then 0 else
+ (if b then
+ if (Hashtbl.find counts r = (i, 1))
+ then 1 else 0
+ else
+ match Hashtbl.find_opt live_regs r with
+ | None -> -1
+ | Some t -> 0
+ )) in
+
+ let priorities = critical_paths successors in
+
+ let current_resources = Array.copy problem.resource_bounds in
+
+ let module InstrSet =
+ struct
+ module MSet =
+ Set.Make (struct
+ type t=int
+ let compare x y =
+ match priorities.(y) - priorities.(x) with
+ | 0 -> x - y
+ | z -> z
+ end)
+
+ let empty = MSet.empty
+ let is_empty = MSet.is_empty
+ let add = MSet.add
+ let remove = MSet.remove
+ let union = MSet.union
+ let iter = MSet.iter
+
+ let compare_regs i x y =
+ let pyi = List.fold_left (fold_delta i) 0 mentions.(y) in
+ (* print_int y;
+ * print_string " ";
+ * print_int pyi;
+ * print_newline ();
+ * flush stdout; *)
+ let pxi = List.fold_left (fold_delta i) 0 mentions.(x) in
+ match pyi - pxi with
+ | 0 -> (match priorities.(y) - priorities.(x) with
+ | 0 -> x - y
+ | z -> z)
+ | z -> z
+
+ (** t is the register class *)
+ let sched_CSR t ready usages =
+ (* print_string "looking for max delta";
+ * print_newline ();
+ * flush stdout; *)
+ let result = ref (-1) in
+ iter (fun i ->
+ if vector_less_equal usages.(i) current_resources
+ then if !result = -1 || (compare_regs t !result i > 0)
+ then result := i) ready;
+ !result
+ end
+ in
+
+ let max_time = bound_max_time problem + 5*nr_instructions in
+ let ready = Array.make max_time InstrSet.empty in
+
+ Array.iteri (fun i preds ->
+ if i < nr_instructions && preds = []
+ then ready.(0) <- InstrSet.add i ready.(0)) predecessors;
+
+ let current_time = ref 0
+ and earliest_time i =
+ try
+ let time = ref (-1) in
+ List.iter (fun (j, latency) ->
+ if times.(j) < 0
+ then raise Exit
+ else let t = times.(j) + latency in
+ if t > !time
+ then time := t) predecessors.(i);
+ assert (!time >= 0);
+ !time
+ with Exit -> -1
+ in
+
+ let advance_time () =
+ (if !current_time < max_time-1
+ then (
+ Array.blit problem.resource_bounds 0 current_resources 0
+ (Array.length current_resources);
+ ready.(!current_time + 1) <-
+ InstrSet.union (ready.(!current_time))
+ (ready.(!current_time +1));
+ ready.(!current_time) <- InstrSet.empty));
+ incr current_time
+ in
+
+ (* ALL MENTIONS TO cnt ARE PLACEHOLDERS *)
+ let cnt = ref 0 in
+
+ let attempt_scheduling ready usages =
+ let result = ref (-1) in
+ DebugPrint.debug "\n\nREADY: ";
+ InstrSet.iter (fun i -> DebugPrint.debug "%d " i) ready;
+ DebugPrint.debug "\n\n";
+ try
+ Array.iteri (fun i avlregs ->
+ DebugPrint.debug "avlregs: %d %d\nlive regs: %d\n"
+ i avlregs (Hashtbl.length live_regs);
+ if !cnt < 5 && avlregs <= regs_thresholds.(i)
+ then (
+ csr_b := true;
+ let maybe = InstrSet.sched_CSR i ready usages in
+ DebugPrint.debug "maybe %d\n" maybe;
+ (if maybe >= 0 &&
+ let delta =
+ List.fold_left (fold_delta i) 0 mentions.(maybe) in
+ DebugPrint.debug "delta %d\n" delta;
+ delta > 0
+ then
+ (vector_subtract usages.(maybe) current_resources;
+ result := maybe)
+ else
+ if not !Clflags.option_regpres_wait_window
+ then
+ (InstrSet.iter (fun ins ->
+ if vector_less_equal usages.(ins) current_resources
+ && List.fold_left (fold_delta i) 0 mentions.(maybe) >= 0
+ then (result := ins;
+ vector_subtract usages.(!result) current_resources;
+ raise Exit)
+ ) ready;
+ if !result <> -1 then
+ vector_subtract usages.(!result) current_resources
+ else incr cnt)
+ else
+ (incr cnt)
+ );
+ raise Exit)) available_regs;
+ InstrSet.iter (fun i ->
+ if vector_less_equal usages.(i) current_resources
+ then (
+ vector_subtract usages.(i) current_resources;
+ result := i;
+ raise Exit)) ready;
+ -1
+ with Exit ->
+ !result in
+
+ while !current_time < max_time
+ do
+ if (InstrSet.is_empty ready.(!current_time))
+ then advance_time ()
+ else
+ match attempt_scheduling ready.(!current_time)
+ problem.instruction_usages with
+ | -1 -> advance_time()
+ | i -> (assert(times.(i) < 0);
+ (DebugPrint.debug "INSTR ISSUED: %d\n" i;
+ if !csr_b && !Clflags.option_debug_compcert > 6 then
+ (Printf.eprintf "REGPRES: high pres class %d\n" i;
+ flush stderr);
+ csr_b := false;
+ (* if !Clflags.option_regpres_wait_window then *)
+ cnt := 0;
+ List.iter (fun (r,b) ->
+ if b then
+ (match Hashtbl.find_opt counts r with
+ | None -> assert false
+ | Some (t, n) ->
+ Hashtbl.remove counts r;
+ if n = 1 then
+ (Hashtbl.remove live_regs r;
+ available_regs.(t)
+ <- available_regs.(t) + 1))
+ else
+ let t = class_r r in
+ match Hashtbl.find_opt live_regs r with
+ | None -> (Hashtbl.add live_regs r t;
+ available_regs.(t)
+ <- available_regs.(t) - 1)
+ | Some i -> ()
+ ) mentions.(i));
+ times.(i) <- !current_time;
+ ready.(!current_time)
+ <- InstrSet.remove i (ready.(!current_time));
+ List.iter (fun (instr_to, latency) ->
+ if instr_to < nr_instructions then
+ match earliest_time instr_to with
+ | -1 -> ()
+ | to_time ->
+ ((* DebugPrint.debug "TO TIME %d : %d\n" to_time
+ * (Array.length ready); *)
+ ready.(to_time)
+ <- InstrSet.add instr_to ready.(to_time))
+ ) successors.(i);
+ successors.(i) <- []
+ )
+ done;
+
+ try
+ let final_time = ref (-1) in
+ for i = 0 to nr_instructions - 1 do
+ DebugPrint.debug "%d " i;
+ (if times.(i) < 0 then raise Exit);
+ (if !final_time < times.(i) + 1 then final_time := times.(i) + 1)
+ done;
+ List.iter (fun (i, latency) ->
+ let target_time = latency + times.(i) in
+ if target_time > !final_time then
+ final_time := target_time) predecessors.(nr_instructions);
+ times.(nr_instructions) <- !final_time;
+ (* DebugPrint.debug_flag := false; *)
+ Some times
+ with Exit ->
+ (* DebugPrint.debug_flag := true; *)
+ DebugPrint.debug "reg_pres_sched failed\n";
+ (* DebugPrint.debug_flag := false; *)
+ None
+
+;;
+
+
+(********************************************************************)
+
+let reg_pres_scheduler_bis (problem : problem) : solution option =
+ (* DebugPrint.debug_flag := true; *)
+ DebugPrint.debug "\nNEW\n\n";
+ let nr_instructions = get_nr_instructions problem in
+ let successors = get_successors problem
+ and predecessors = get_predecessors problem
+ and times = Array.make (nr_instructions+1) (-1) in
+ let live_regs_entry = problem.live_regs_entry in
+
+ (* let available_regs = Array.copy Machregsaux.nr_regs in *)
+
+ let class_r r =
+ Machregsaux.class_of_type (problem.typing r) in
+
+ let live_regs = Hashtbl.create 42 in
+
+ List.iter (fun r -> let classe = Machregsaux.class_of_type
+ (problem.typing r) in
+ (* available_regs.(classe)
+ * <- available_regs.(classe) - 1; *)
+ Hashtbl.add live_regs r classe)
+ (Registers.Regset.elements live_regs_entry);
+
+
+ let counts, mentions =
+ match problem.reference_counting with
+ | Some (l, r) -> l, r
+ | None -> assert false
+ in
+
+ let fold_delta a (r, b) =
+ a + (if b then
+ match Hashtbl.find_opt counts r with
+ | Some (_, 1) -> 1
+ | _ -> 0
+ else
+ match Hashtbl.find_opt live_regs r with
+ | None -> -1
+ | Some t -> 0
+ ) in
+
+ let priorities = critical_paths successors in
+
+ let current_resources = Array.copy problem.resource_bounds in
+
+ let compare_pres x y =
+ let pdy = List.fold_left (fold_delta) 0 mentions.(y) in
+ let pdx = List.fold_left (fold_delta) 0 mentions.(x) in
+ match pdy - pdx with
+ | 0 -> x - y
+ | z -> z
+ in
+
+ let module InstrSet =
+ Set.Make (struct
+ type t = int
+ let compare x y =
+ match priorities.(y) - priorities.(x) with
+ | 0 -> x - y
+ | z -> z
+ end) in
+
+ let max_time = bound_max_time problem (* + 5*nr_instructions *) in
+ let ready = Array.make max_time InstrSet.empty in
+
+ Array.iteri (fun i preds ->
+ if i < nr_instructions && preds = []
+ then ready.(0) <- InstrSet.add i ready.(0)) predecessors;
+
+ let current_time = ref 0
+ and earliest_time i =
+ try
+ let time = ref (-1) in
+ List.iter (fun (j, latency) ->
+ if times.(j) < 0
+ then raise Exit
+ else let t = times.(j) + latency in
+ if t > !time
+ then time := t) predecessors.(i);
+ assert (!time >= 0);
+ !time
+ with Exit -> -1
+ in
+
+ let advance_time () =
+ (* Printf.printf "ADV\n";
+ * flush stdout; *)
+ (if !current_time < max_time-1
+ then (
+ Array.blit problem.resource_bounds 0 current_resources 0
+ (Array.length current_resources);
+ ready.(!current_time + 1) <-
+ InstrSet.union (ready.(!current_time))
+ (ready.(!current_time +1));
+ ready.(!current_time) <- InstrSet.empty));
+ incr current_time
+ in
+
+
+ let attempt_scheduling ready usages =
+ let result = ref [] in
+ try
+ InstrSet.iter (fun i ->
+ if vector_less_equal usages.(i) current_resources
+ then
+ if !result = [] || priorities.(i) = priorities.(List.hd (!result))
+ then
+ result := i::(!result)
+ else raise Exit
+ ) ready;
+ if !result <> [] then raise Exit;
+ -1
+ with
+ Exit ->
+ let mini = List.fold_left (fun a b ->
+ if a = -1 || compare_pres a b > 0
+ then b else a
+ ) (-1) !result in
+ vector_subtract usages.(mini) current_resources;
+ mini
+ in
+
+ while !current_time < max_time
+ do
+ if (InstrSet.is_empty ready.(!current_time))
+ then advance_time ()
+ else
+ match attempt_scheduling ready.(!current_time)
+ problem.instruction_usages with
+ | -1 -> advance_time()
+ | i -> (
+ DebugPrint.debug "ISSUED: %d\nREADY: " i;
+ InstrSet.iter (fun i -> DebugPrint.debug "%d " i)
+ ready.(!current_time);
+ DebugPrint.debug "\nSUCC: ";
+ List.iter (fun (i, l) -> DebugPrint.debug "%d " i)
+ successors.(i);
+ DebugPrint.debug "\n\n";
+ assert(times.(i) < 0);
+ times.(i) <- !current_time;
+ ready.(!current_time)
+ <- InstrSet.remove i (ready.(!current_time));
+ (List.iter (fun (r,b) ->
+ if b then
+ (match Hashtbl.find_opt counts r with
+ | None -> assert false
+ | Some (t, n) ->
+ Hashtbl.remove counts r;
+ if n = 1 then
+ (Hashtbl.remove live_regs r;
+ (* available_regs.(t)
+ * <- available_regs.(t) + 1 *)))
+ else
+ let t = class_r r in
+ match Hashtbl.find_opt live_regs r with
+ | None -> (Hashtbl.add live_regs r t;
+ (* available_regs.(t)
+ * <- available_regs.(t) - 1 *))
+ | Some i -> ()
+ ) mentions.(i));
+ List.iter (fun (instr_to, latency) ->
+ if instr_to < nr_instructions then
+ match earliest_time instr_to with
+ | -1 -> ()
+ | to_time ->
+ ((* DebugPrint.debug "TO TIME %d : %d\n" to_time
+ * (Array.length ready); *)
+ ready.(to_time)
+ <- InstrSet.add instr_to ready.(to_time))
+ ) successors.(i);
+ successors.(i) <- []
+ )
+ done;
+
+ try
+ let final_time = ref (-1) in
+ for i = 0 to nr_instructions - 1 do
+ (* print_int i;
+ * flush stdout; *)
+ (if times.(i) < 0 then raise Exit);
+ (if !final_time < times.(i) + 1 then final_time := times.(i) + 1)
+ done;
+ List.iter (fun (i, latency) ->
+ let target_time = latency + times.(i) in
+ if target_time > !final_time then
+ final_time := target_time) predecessors.(nr_instructions);
+ times.(nr_instructions) <- !final_time;
+ (* DebugPrint.debug_flag := false; *)
+ Some times
+ with Exit ->
+ DebugPrint.debug "reg_pres_sched failed\n";
+ (* DebugPrint.debug_flag := false; *)
+ None
+
+;;
+
+(********************************************************************)
+
+type bundle = int list;;
+
+let rec extract_deps_to index = function
+ | [] -> []
+ | dep :: deps -> let extracts = extract_deps_to index deps in
+ if (dep.instr_to == index) then
+ dep :: extracts
+ else
+ extracts
+
+exception InvalidBundle;;
+
+let dependency_check problem bundle index =
+ let index_deps = extract_deps_to index problem.latency_constraints in
+ List.iter (fun i ->
+ List.iter (fun dep ->
+ if (dep.instr_from == i) then raise InvalidBundle
+ ) index_deps
+ ) bundle;;
+
+let rec make_bundle problem resources bundle index =
+ let resources_copy = Array.copy resources in
+ let nr_instructions = get_nr_instructions problem in
+ if (index >= nr_instructions) then (bundle, index+1) else
+ let inst_usage = problem.instruction_usages.(index) in
+ try match vector_less_equal inst_usage resources with
+ | false -> raise InvalidBundle
+ | true -> (
+ dependency_check problem bundle index;
+ vector_subtract problem.instruction_usages.(index) resources_copy;
+ make_bundle problem resources_copy (index::bundle) (index+1)
+ )
+ with InvalidBundle -> (bundle, index);;
+
+let rec make_bundles problem index : bundle list =
+ if index >= get_nr_instructions problem then
+ []
+ else
+ let (bundle, new_index) = make_bundle problem problem.resource_bounds [] index in
+ bundle :: (make_bundles problem new_index);;
+
+let bundles_to_schedule problem bundles : solution =
+ let nr_instructions = get_nr_instructions problem in
+ let schedule = Array.make (nr_instructions+1) (nr_instructions+4) in
+ let time = ref 0 in
+ List.iter (fun bundle ->
+ begin
+ List.iter (fun i ->
+ schedule.(i) <- !time
+ ) bundle;
+ time := !time + 1
+ end
+ ) bundles; schedule;;
+
+let greedy_scheduler (problem : problem) : solution option =
+ let bundles = make_bundles problem 0 in
+ Some (bundles_to_schedule problem bundles);;
+
+(* alternate implementation
+let swap_array_elements a i j =
+ let x = a.(i) in
+ a.(i) <- a.(j);
+ a.(j) <- x;;
+
+let array_reverse_slice a first last =
+ let i = ref first and j = ref last in
+ while i < j
+ do
+ swap_array_elements a !i !j;
+ incr i;
+ decr j
+ done;;
+
+let array_reverse a =
+ let a' = Array.copy a in
+ array_reverse_slice a' 0 ((Array.length a)-1);
+ a';;
+ *)
+
+(* unneeded
+let array_reverse a =
+ let n=Array.length a in
+ Array.init n (fun i -> a.(n-1-i));;
+ *)
+
+let reverse_constraint nr_instructions ctr =
+ { instr_to = nr_instructions -ctr.instr_from;
+ instr_from = nr_instructions - ctr.instr_to;
+ latency = ctr.latency };;
+
+(* unneeded
+let rec list_map_filter f = function
+ | [] -> []
+ | h::t ->
+ (match f h with
+ | None -> list_map_filter f t
+ | Some x -> x :: (list_map_filter f t));;
+ *)
+
+let reverse_problem problem =
+ let nr_instructions = get_nr_instructions problem in
+ {
+ max_latency = problem.max_latency;
+ resource_bounds = problem.resource_bounds;
+ live_regs_entry = Registers.Regset.empty; (* PLACEHOLDER *)
+ (* Not needed for the revlist sched, and for now we wont bother
+ with creating a reverse scheduler aware of reg press *)
+
+ typing = problem.typing;
+ reference_counting = problem.reference_counting;
+ instruction_usages = Array.init (nr_instructions + 1)
+ (fun i ->
+ if i=0
+ then Array.map (fun _ -> 0) problem.resource_bounds else problem.instruction_usages.(nr_instructions - i));
+ latency_constraints = List.map (reverse_constraint nr_instructions)
+ problem.latency_constraints
+ };;
+
+let max_scheduled_time solution =
+ let time = ref (-1) in
+ for i = 0 to ((Array.length solution) - 2)
+ do
+ time := max !time solution.(i)
+ done;
+ !time;;
+
+(*
+let recompute_makespan problem solution =
+ let n = (Array.length solution) - 1 and ms = ref 0 in
+ List.iter (fun cstr ->
+ if cstr.instr_to = n
+ then ms := max !ms (solution.(cstr.instr_from) + cstr.latency)
+ ) problem.latency_constraints;
+ !ms;;
+ *)
+
+let schedule_reversed (scheduler : problem -> solution option)
+ (problem : problem) =
+ match scheduler (reverse_problem problem) with
+ | None -> None
+ | Some solution ->
+ let nr_instructions = get_nr_instructions problem in
+ let makespan = max_scheduled_time solution in
+ let ret = Array.init (nr_instructions + 1)
+ (fun i -> makespan-solution.(nr_instructions-i)) in
+ ret.(nr_instructions) <- max ((max_scheduled_time ret) + 1)
+ (ret.(nr_instructions));
+ Some ret;;
+
+(** Schedule the problem using a greedy list scheduling algorithm, from the end. *)
+let reverse_list_scheduler = schedule_reversed list_scheduler;;
+
+let check_problem problem =
+ (if (Array.length problem.instruction_usages) < 1
+ then failwith "length(problem.instruction_usages) < 1");;
+
+let validated_scheduler (scheduler : problem -> solution option)
+ (problem : problem) =
+ check_problem problem;
+ match scheduler problem with
+ | None -> None
+ | (Some solution) as ret -> check_schedule problem solution; ret;;
+
+let get_max_latency solution =
+ solution.((Array.length solution)-1);;
+
+let show_date_ranges problem =
+ let deadline = problem.max_latency in
+ assert(deadline >= 0);
+ let successors = get_successors problem
+ and predecessors = get_predecessors problem in
+ let earliest_dates : int array = get_earliest_dates predecessors
+ and latest_dates : int array = get_latest_dates deadline successors in
+ assert ((Array.length earliest_dates) =
+ (Array.length latest_dates));
+ Array.iteri (fun i early ->
+ let late = latest_dates.(i) in
+ Printf.printf "t[%d] in %d..%d\n" i early late)
+ earliest_dates;;
+
+type pseudo_boolean_problem_type =
+ | SATISFIABILITY
+ | OPTIMIZATION;;
+
+type pseudo_boolean_mapper = {
+ mapper_pb_type : pseudo_boolean_problem_type;
+ mapper_nr_instructions : int;
+ mapper_nr_pb_variables : int;
+ mapper_earliest_dates : int array;
+ mapper_latest_dates : int array;
+ mapper_var_offsets : int array;
+ mapper_final_predecessors : (int * int) list
+};;
+
+(* Latency constraints are:
+ presence of instr-to at each t <= sum of presences of instr-from at compatible times
+
+ if reverse_encoding
+ presence of instr-from at each t <= sum of presences of instr-to at compatible times *)
+
+(* Experiments show reverse_encoding=true multiplies time by 2 in sat4j
+ without making hard instances easier *)
+let direct_encoding = false
+and reverse_encoding = false
+and delta_encoding = true
+
+let pseudo_boolean_print_problem channel problem pb_type =
+ let deadline = problem.max_latency in
+ assert (deadline > 0);
+ let nr_instructions = get_nr_instructions problem
+ and nr_resources = get_nr_resources problem
+ and successors = get_successors problem
+ and predecessors = get_predecessors problem in
+ let earliest_dates = get_earliest_dates predecessors
+ and latest_dates = get_latest_dates deadline successors in
+ let var_offsets = Array.make
+ (match pb_type with
+ | OPTIMIZATION -> nr_instructions+1
+ | SATISFIABILITY -> nr_instructions) 0 in
+ let nr_pb_variables =
+ (let nr = ref 0 in
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ var_offsets.(i) <- !nr;
+ nr := !nr + latest_dates.(i) - earliest_dates.(i) + 1
+ done;
+ !nr)
+ and nr_pb_constraints =
+ (match pb_type with
+ | OPTIMIZATION -> nr_instructions+1
+ | SATISFIABILITY -> nr_instructions) +
+
+ (let count = ref 0 in
+ for t=0 to deadline-1
+ do
+ for j=0 to nr_resources-1
+ do
+ try
+ for i=0 to nr_instructions-1
+ do
+ let usage = problem.instruction_usages.(i).(j) in
+ if t >= earliest_dates.(i) && t <= latest_dates.(i)
+ && usage > 0 then raise Exit
+ done
+ with Exit -> incr count
+ done
+ done;
+ !count) +
+
+ (let count=ref 0 in
+ List.iter
+ (fun ctr ->
+ if ctr.instr_to < nr_instructions
+ then count := !count + 1 + latest_dates.(ctr.instr_to)
+ - earliest_dates.(ctr.instr_to)
+ + (if reverse_encoding
+ then 1 + latest_dates.(ctr.instr_from)
+ - earliest_dates.(ctr.instr_from)
+ else 0)
+ )
+ problem.latency_constraints;
+ !count) +
+
+ (match pb_type with
+ | OPTIMIZATION -> (1 + deadline - earliest_dates.(nr_instructions)) * nr_instructions
+ | SATISFIABILITY -> 0)
+ and measured_nr_constraints = ref 0 in
+
+ let pb_var i t =
+ assert(t >= earliest_dates.(i));
+ assert(t <= latest_dates.(i));
+ let v = 1+var_offsets.(i)+t-earliest_dates.(i) in
+ assert(v <= nr_pb_variables);
+ Printf.sprintf "x%d" v in
+
+ let end_constraint () =
+ begin
+ output_string channel ";\n";
+ incr measured_nr_constraints
+ end in
+
+ let gen_latency_constraint i_to i_from latency t_to =
+ Printf.fprintf channel "* t[%d] - t[%d] >= %d when t[%d]=%d\n"
+ i_to i_from latency i_to t_to;
+ for t_from=earliest_dates.(i_from) to
+ int_min latest_dates.(i_from) (t_to - latency)
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i_from t_from)
+ done;
+ Printf.fprintf channel "-1 %s " (pb_var i_to t_to);
+ Printf.fprintf channel ">= 0";
+ end_constraint()
+
+ and gen_dual_latency_constraint i_to i_from latency t_from =
+ Printf.fprintf channel "* t[%d] - t[%d] >= %d when t[%d]=%d\n"
+ i_to i_from latency i_to t_from;
+ for t_to=int_max earliest_dates.(i_to) (t_from + latency)
+ to latest_dates.(i_to)
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i_to t_to)
+ done;
+ Printf.fprintf channel "-1 %s " (pb_var i_from t_from);
+ Printf.fprintf channel ">= 0";
+ end_constraint()
+ in
+
+ Printf.fprintf channel "* #variable= %d #constraint= %d\n" nr_pb_variables nr_pb_constraints;
+ Printf.fprintf channel "* nr_instructions=%d deadline=%d\n" nr_instructions deadline;
+ begin
+ match pb_type with
+ | SATISFIABILITY -> ()
+ | OPTIMIZATION ->
+ output_string channel "min:";
+ for t=earliest_dates.(nr_instructions) to deadline
+ do
+ Printf.fprintf channel " %+d %s" t (pb_var nr_instructions t)
+ done;
+ output_string channel ";\n";
+ end;
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ let early = earliest_dates.(i) and late= latest_dates.(i) in
+ Printf.fprintf channel "* t[%d] in %d..%d\n" i early late;
+ for t=early to late
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i t)
+ done;
+ Printf.fprintf channel "= 1";
+ end_constraint()
+ done;
+
+ for t=0 to deadline-1
+ do
+ for j=0 to nr_resources-1
+ do
+ let bound = problem.resource_bounds.(j)
+ and coeffs = ref [] in
+ for i=0 to nr_instructions-1
+ do
+ let usage = problem.instruction_usages.(i).(j) in
+ if t >= earliest_dates.(i) && t <= latest_dates.(i)
+ && usage > 0
+ then coeffs := (i, usage) :: !coeffs
+ done;
+ if !coeffs <> [] then
+ begin
+ Printf.fprintf channel "* resource #%d at t=%d <= %d\n" j t bound;
+ List.iter (fun (i, usage) ->
+ Printf.fprintf channel "%+d %s " (-usage) (pb_var i t)) !coeffs;
+ Printf.fprintf channel ">= %d" (-bound);
+ end_constraint();
+ end
+ done
+ done;
+
+ List.iter
+ (fun ctr ->
+ if ctr.instr_to < nr_instructions then
+ begin
+ for t_to=earliest_dates.(ctr.instr_to) to latest_dates.(ctr.instr_to)
+ do
+ gen_latency_constraint ctr.instr_to ctr.instr_from ctr.latency t_to
+ done;
+ if reverse_encoding
+ then
+ for t_from=earliest_dates.(ctr.instr_from) to latest_dates.(ctr.instr_from)
+ do
+ gen_dual_latency_constraint ctr.instr_to ctr.instr_from ctr.latency t_from
+ done
+ end
+ ) problem.latency_constraints;
+
+ begin
+ match pb_type with
+ | SATISFIABILITY -> ()
+ | OPTIMIZATION ->
+ let final_latencies = Array.make nr_instructions 1 in
+ List.iter (fun (i, latency) ->
+ final_latencies.(i) <- int_max final_latencies.(i) latency)
+ predecessors.(nr_instructions);
+ for t_to=earliest_dates.(nr_instructions) to deadline
+ do
+ for i_from = 0 to nr_instructions -1
+ do
+ gen_latency_constraint nr_instructions i_from final_latencies.(i_from) t_to
+ done
+ done
+ end;
+ assert (!measured_nr_constraints = nr_pb_constraints);
+ {
+ mapper_pb_type = pb_type;
+ mapper_nr_instructions = nr_instructions;
+ mapper_nr_pb_variables = nr_pb_variables;
+ mapper_earliest_dates = earliest_dates;
+ mapper_latest_dates = latest_dates;
+ mapper_var_offsets = var_offsets;
+ mapper_final_predecessors = predecessors.(nr_instructions)
+ };;
+
+type pb_answer =
+ | Positive
+ | Negative
+ | Unknown
+
+let line_to_pb_solution sol line nr_pb_variables =
+ let assign s v =
+ begin
+ let i = int_of_string s in
+ sol.(i-1) <- v
+ end in
+ List.iter
+ begin
+ function "" -> ()
+ | item ->
+ (match String.get item 0 with
+ | '+' ->
+ assert ((String.length item) >= 3);
+ assert ((String.get item 1) = 'x');
+ assign (String.sub item 2 ((String.length item)-2)) Positive
+ | '-' ->
+ assert ((String.length item) >= 3);
+ assert ((String.get item 1) = 'x');
+ assign (String.sub item 2 ((String.length item)-2)) Negative
+ | 'x' ->
+ assert ((String.length item) >= 2);
+ assign (String.sub item 1 ((String.length item)-1)) Positive
+ | _ -> failwith "syntax error in pseudo Boolean solution: epected + - or x"
+ )
+ end
+ (String.split_on_char ' ' (String.sub line 2 ((String.length line)-2)));;
+
+let pb_solution_to_schedule mapper pb_solution =
+ Array.mapi (fun i offset ->
+ let first = mapper.mapper_earliest_dates.(i)
+ and last = mapper.mapper_latest_dates.(i)
+ and time = ref (-1) in
+ for t=first to last
+ do
+ match pb_solution.(t - first + offset) with
+ | Positive ->
+ (if !time = -1
+ then time:=t
+ else failwith "duplicate time in pseudo boolean solution")
+ | Negative -> ()
+ | Unknown -> failwith "unknown value in pseudo boolean solution"
+ done;
+ (if !time = -1
+ then failwith "no time in pseudo boolean solution");
+ !time
+ ) mapper.mapper_var_offsets;;
+
+let pseudo_boolean_read_solution mapper channel =
+ let optimum = ref (-1)
+ and optimum_found = ref false
+ and solution = Array.make mapper.mapper_nr_pb_variables Unknown in
+ try
+ while true do
+ match input_line channel with
+ | "" -> ()
+ | line ->
+ begin
+ match String.get line 0 with
+ | 'c' -> ()
+ | 'o' ->
+ assert ((String.length line) >= 2);
+ assert ((String.get line 1) = ' ');
+ optimum := int_of_string (String.sub line 2 ((String.length line)-2))
+ | 's' -> (match line with
+ | "s OPTIMUM FOUND" -> optimum_found := true
+ | "s SATISFIABLE" -> ()
+ | "s UNSATISFIABLE" -> close_in channel;
+ raise Unschedulable
+ | _ -> failwith line)
+ | 'v' -> line_to_pb_solution solution line mapper.mapper_nr_pb_variables
+ | x -> Printf.printf "unknown: %s\n" line
+ end
+ done;
+ assert false
+ with End_of_file ->
+ close_in channel;
+ begin
+ let sol = pb_solution_to_schedule mapper solution in
+ sol
+ end;;
+
+let recompute_max_latency mapper solution =
+ let maxi = ref (-1) in
+ for i=0 to (mapper.mapper_nr_instructions-1)
+ do
+ maxi := int_max !maxi (1+solution.(i))
+ done;
+ List.iter (fun (i, latency) ->
+ maxi := int_max !maxi (solution.(i) + latency)) mapper.mapper_final_predecessors;
+ !maxi;;
+
+let adjust_check_solution mapper solution =
+ match mapper.mapper_pb_type with
+ | OPTIMIZATION ->
+ let max_latency = recompute_max_latency mapper solution in
+ assert (max_latency = solution.(mapper.mapper_nr_instructions));
+ solution
+ | SATISFIABILITY ->
+ let max_latency = recompute_max_latency mapper solution in
+ Array.init (mapper.mapper_nr_instructions+1)
+ (fun i -> if i < mapper.mapper_nr_instructions
+ then solution.(i)
+ else max_latency);;
+
+(* let pseudo_boolean_solver = ref "/local/monniaux/progs/naps/naps" *)
+(* let pseudo_boolean_solver = ref "/local/monniaux/packages/sat4j/org.sat4j.pb.jar CuttingPlanes" *)
+
+(* let pseudo_boolean_solver = ref "java -jar /usr/share/java/org.sat4j.pb.jar CuttingPlanes" *)
+(* let pseudo_boolean_solver = ref "java -jar /usr/share/java/org.sat4j.pb.jar" *)
+(* let pseudo_boolean_solver = ref "clasp" *)
+(* let pseudo_boolean_solver = ref "/home/monniaux/progs/CP/open-wbo/open-wbo_static -formula=1" *)
+(* let pseudo_boolean_solver = ref "/home/monniaux/progs/CP/naps/naps" *)
+(* let pseudo_boolean_solver = ref "/home/monniaux/progs/CP/minisatp/build/release/bin/minisatp" *)
+(* let pseudo_boolean_solver = ref "java -jar sat4j-pb.jar CuttingPlanesStar" *)
+let pseudo_boolean_solver = ref "pb_solver"
+
+let pseudo_boolean_scheduler pb_type problem =
+ try
+ let filename_in = "problem.opb" in
+ (* needed only if not using stdout and filename_out = "problem.sol" *)
+ let mapper =
+ with_out_channel (open_out filename_in)
+ (fun opb_problem ->
+ pseudo_boolean_print_problem opb_problem problem pb_type) in
+ Some (with_in_channel
+ (Unix.open_process_in (!pseudo_boolean_solver ^ " " ^ filename_in))
+ (fun opb_solution -> adjust_check_solution mapper (pseudo_boolean_read_solution mapper opb_solution)))
+ with
+ | Unschedulable -> None;;
+
+let rec reoptimizing_scheduler (scheduler : scheduler) (previous_solution : solution) (problem : problem) =
+ if (get_max_latency previous_solution)>1 then
+ begin
+ Printf.printf "reoptimizing < %d\n" (get_max_latency previous_solution);
+ flush stdout;
+ match scheduler
+ { problem with max_latency = (get_max_latency previous_solution)-1 }
+ with
+ | None -> previous_solution
+ | Some solution -> reoptimizing_scheduler scheduler solution problem
+ end
+ else previous_solution;;
+
+let smt_var i = Printf.sprintf "t%d" i
+
+let is_resource_used problem j =
+ try
+ Array.iter (fun usages ->
+ if usages.(j) > 0
+ then raise Exit) problem.instruction_usages;
+ false
+ with Exit -> true;;
+
+let smt_use_quantifiers = false
+
+let smt_print_problem channel problem =
+ let nr_instructions = get_nr_instructions problem in
+ let gen_smt_resource_constraint time j =
+ output_string channel "(<= (+";
+ Array.iteri
+ (fun i usages ->
+ let usage=usages.(j) in
+ if usage > 0
+ then Printf.fprintf channel " (ite (= %s %s) %d 0)"
+ time (smt_var i) usage)
+ problem.instruction_usages;
+ Printf.fprintf channel ") %d)" problem.resource_bounds.(j)
+ in
+ output_string channel "(set-option :produce-models true)\n";
+ for i=0 to nr_instructions
+ do
+ Printf.fprintf channel "(declare-const %s Int)\n" (smt_var i);
+ Printf.fprintf channel "(assert (>= %s 0))\n" (smt_var i)
+ done;
+ for i=0 to nr_instructions-1
+ do
+ Printf.fprintf channel "(assert (< %s %s))\n"
+ (smt_var i) (smt_var nr_instructions)
+ done;
+ (if problem.max_latency > 0
+ then Printf.fprintf channel "(assert (<= %s %d))\n"
+ (smt_var nr_instructions) problem.max_latency);
+ List.iter (fun ctr ->
+ Printf.fprintf channel "(assert (>= (- %s %s) %d))\n"
+ (smt_var ctr.instr_to)
+ (smt_var ctr.instr_from)
+ ctr.latency) problem.latency_constraints;
+ for j=0 to (Array.length problem.resource_bounds)-1
+ do
+ if is_resource_used problem j
+ then
+ begin
+ if smt_use_quantifiers
+ then
+ begin
+ Printf.fprintf channel
+ "; resource #%d <= %d\n(assert (forall ((t Int)) "
+ j problem.resource_bounds.(j);
+ gen_smt_resource_constraint "t" j;
+ output_string channel "))\n"
+ end
+ else
+ begin
+ (if problem.max_latency < 0
+ then failwith "quantifier explosion needs max latency");
+ for t=0 to problem.max_latency
+ do
+ Printf.fprintf channel
+ "; resource #%d <= %d at t=%d\n(assert "
+ j problem.resource_bounds.(j) t;
+ gen_smt_resource_constraint (string_of_int t) j;
+ output_string channel ")\n"
+ done
+ end
+ end
+ done;
+ output_string channel "(check-sat)(get-model)\n";;
+
+
+let ilp_print_problem channel problem pb_type =
+ let deadline = problem.max_latency in
+ assert (deadline > 0);
+ let nr_instructions = get_nr_instructions problem
+ and nr_resources = get_nr_resources problem
+ and successors = get_successors problem
+ and predecessors = get_predecessors problem in
+ let earliest_dates = get_earliest_dates predecessors
+ and latest_dates = get_latest_dates deadline successors in
+
+ let pb_var i t =
+ Printf.sprintf "x%d_%d" i t in
+
+ let gen_latency_constraint i_to i_from latency t_to =
+ Printf.fprintf channel "\\ t[%d] - t[%d] >= %d when t[%d]=%d\n"
+ i_to i_from latency i_to t_to;
+ Printf.fprintf channel "c_%d_%d_%d_%d: "
+ i_to i_from latency t_to;
+ for t_from=earliest_dates.(i_from) to
+ int_min latest_dates.(i_from) (t_to - latency)
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i_from t_from)
+ done;
+ Printf.fprintf channel "-1 %s " (pb_var i_to t_to);
+ output_string channel ">= 0\n"
+
+ and gen_dual_latency_constraint i_to i_from latency t_from =
+ Printf.fprintf channel "\\ t[%d] - t[%d] >= %d when t[%d]=%d\n"
+ i_to i_from latency i_to t_from;
+ Printf.fprintf channel "d_%d_%d_%d_%d: "
+ i_to i_from latency t_from;
+ for t_to=int_max earliest_dates.(i_to) (t_from + latency)
+ to latest_dates.(i_to)
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i_to t_to)
+ done;
+ Printf.fprintf channel "-1 %s " (pb_var i_from t_from);
+ Printf.fprintf channel ">= 0\n"
+
+ and gen_delta_constraint i_from i_to latency =
+ if delta_encoding
+ then Printf.fprintf channel "l_%d_%d_%d: +1 t%d -1 t%d >= %d\n"
+ i_from i_to latency i_to i_from latency
+
+ in
+
+ Printf.fprintf channel "\\ nr_instructions=%d deadline=%d\n" nr_instructions deadline;
+ begin
+ match pb_type with
+ | SATISFIABILITY -> output_string channel "Minimize dummy: 0\n"
+ | OPTIMIZATION ->
+ Printf.fprintf channel "Minimize\nmakespan: t%d\n" nr_instructions
+ end;
+ output_string channel "Subject To\n";
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ let early = earliest_dates.(i) and late= latest_dates.(i) in
+ Printf.fprintf channel "\\ t[%d] in %d..%d\ntimes%d: " i early late i;
+ for t=early to late
+ do
+ Printf.fprintf channel "+1 %s " (pb_var i t)
+ done;
+ Printf.fprintf channel "= 1\n"
+ done;
+
+ for t=0 to deadline-1
+ do
+ for j=0 to nr_resources-1
+ do
+ let bound = problem.resource_bounds.(j)
+ and coeffs = ref [] in
+ for i=0 to nr_instructions-1
+ do
+ let usage = problem.instruction_usages.(i).(j) in
+ if t >= earliest_dates.(i) && t <= latest_dates.(i)
+ && usage > 0
+ then coeffs := (i, usage) :: !coeffs
+ done;
+ if !coeffs <> [] then
+ begin
+ Printf.fprintf channel "\\ resource #%d at t=%d <= %d\nr%d_%d: " j t bound j t;
+ List.iter (fun (i, usage) ->
+ Printf.fprintf channel "%+d %s " (-usage) (pb_var i t)) !coeffs;
+ Printf.fprintf channel ">= %d\n" (-bound)
+ end
+ done
+ done;
+
+ List.iter
+ (fun ctr ->
+ if ctr.instr_to < nr_instructions then
+ begin
+ gen_delta_constraint ctr.instr_from ctr.instr_to ctr.latency;
+ begin
+ if direct_encoding
+ then
+ for t_to=earliest_dates.(ctr.instr_to) to latest_dates.(ctr.instr_to)
+ do
+ gen_latency_constraint ctr.instr_to ctr.instr_from ctr.latency t_to
+ done
+ end;
+ begin
+ if reverse_encoding
+ then
+ for t_from=earliest_dates.(ctr.instr_from) to latest_dates.(ctr.instr_from)
+ do
+ gen_dual_latency_constraint ctr.instr_to ctr.instr_from ctr.latency t_from
+ done
+ end
+ end
+ ) problem.latency_constraints;
+
+ begin
+ match pb_type with
+ | SATISFIABILITY -> ()
+ | OPTIMIZATION ->
+ let final_latencies = Array.make nr_instructions 1 in
+ List.iter (fun (i, latency) ->
+ final_latencies.(i) <- int_max final_latencies.(i) latency)
+ predecessors.(nr_instructions);
+ for i_from = 0 to nr_instructions -1
+ do
+ gen_delta_constraint i_from nr_instructions final_latencies.(i_from)
+ done;
+ for t_to=earliest_dates.(nr_instructions) to deadline
+ do
+ for i_from = 0 to nr_instructions -1
+ do
+ gen_latency_constraint nr_instructions i_from final_latencies.(i_from) t_to
+ done
+ done
+ end;
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ Printf.fprintf channel "ct%d : -1 t%d" i i;
+ let early = earliest_dates.(i) and late= latest_dates.(i) in
+ for t=early to late do
+ Printf.fprintf channel " +%d %s" t (pb_var i t)
+ done;
+ output_string channel " = 0\n"
+ done;
+ output_string channel "Bounds\n";
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ let early = earliest_dates.(i) and late= latest_dates.(i) in
+ begin
+ Printf.fprintf channel "%d <= t%d <= %d\n" early i late;
+ if true then
+ for t=early to late do
+ Printf.fprintf channel "0 <= %s <= 1\n" (pb_var i t)
+ done
+ end
+ done;
+ output_string channel "Integer\n";
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ Printf.fprintf channel "t%d " i
+ done;
+ output_string channel "\nBinary\n";
+ for i=0 to (match pb_type with
+ | OPTIMIZATION -> nr_instructions
+ | SATISFIABILITY -> nr_instructions-1)
+ do
+ let early = earliest_dates.(i) and late= latest_dates.(i) in
+ for t=early to late do
+ output_string channel (pb_var i t);
+ output_string channel " "
+ done;
+ output_string channel "\n"
+ done;
+ output_string channel "End\n";
+ {
+ mapper_pb_type = pb_type;
+ mapper_nr_instructions = nr_instructions;
+ mapper_nr_pb_variables = 0;
+ mapper_earliest_dates = earliest_dates;
+ mapper_latest_dates = latest_dates;
+ mapper_var_offsets = [| |];
+ mapper_final_predecessors = predecessors.(nr_instructions)
+ };;
+
+(* Guess what? Cplex sometimes outputs 11.000000004 instead of integer 11 *)
+
+let positive_float_round x = truncate (x +. 0.5)
+
+let float_round (x : float) : int =
+ if x > 0.0
+ then positive_float_round x
+ else - (positive_float_round (-. x))
+
+let rounded_int_of_string x = float_round (float_of_string x)
+
+let ilp_read_solution mapper channel =
+ let times = Array.make
+ (match mapper.mapper_pb_type with
+ | OPTIMIZATION -> 1+mapper.mapper_nr_instructions
+ | SATISFIABILITY -> mapper.mapper_nr_instructions) (-1) in
+ try
+ while true do
+ let line = input_line channel in
+ ( if (String.length line) < 3
+ then failwith (Printf.sprintf "bad ilp output: length(line) < 3: %s" line));
+ match String.get line 0 with
+ | 'x' -> ()
+ | 't' -> let space =
+ try String.index line ' '
+ with Not_found ->
+ failwith "bad ilp output: no t variable number"
+ in
+ let tnumber =
+ try int_of_string (String.sub line 1 (space-1))
+ with Failure _ ->
+ failwith "bad ilp output: not a variable number"
+ in
+ (if tnumber < 0 || tnumber >= (Array.length times)
+ then failwith (Printf.sprintf "bad ilp output: not a correct variable number: %d (%d)" tnumber (Array.length times)));
+ let value =
+ let s = String.sub line (space+1) ((String.length line)-space-1) in
+ try rounded_int_of_string s
+ with Failure _ ->
+ failwith (Printf.sprintf "bad ilp output: not a time number (%s)" s)
+ in
+ (if value < 0
+ then failwith "bad ilp output: negative time");
+ times.(tnumber) <- value
+ | '#' -> ()
+ | '0' -> ()
+ | _ -> failwith (Printf.sprintf "bad ilp output: bad variable initial, line = %s" line)
+ done;
+ assert false
+ with End_of_file ->
+ Array.iteri (fun i x ->
+ if i<(Array.length times)-1
+ && x<0 then raise Unschedulable) times;
+ times;;
+
+let ilp_solver = ref "ilp_solver"
+
+let problem_nr = ref 0
+
+let ilp_scheduler pb_type problem =
+ try
+ let filename_in = Printf.sprintf "problem%05d.lp" !problem_nr
+ and filename_out = Printf.sprintf "problem%05d.sol" !problem_nr in
+ incr problem_nr;
+ let mapper = with_out_channel (open_out filename_in)
+ (fun opb_problem -> ilp_print_problem opb_problem problem pb_type) in
+
+ begin
+ match Unix.system (!ilp_solver ^ " " ^ filename_in ^ " " ^ filename_out) with
+ | Unix.WEXITED 0 ->
+ Some (with_in_channel
+ (open_in filename_out)
+ (fun opb_solution ->
+ adjust_check_solution mapper
+ (ilp_read_solution mapper opb_solution)))
+ | Unix.WEXITED _ -> failwith "failed to start ilp solver"
+ | _ -> None
+ end
+ with
+ | Unschedulable -> None;;
+
+let current_utime_all () =
+ let t = Unix.times() in
+ t.Unix.tms_cutime +. t.Unix.tms_utime;;
+
+let utime_all_fn fn arg =
+ let utime_start = current_utime_all () in
+ let output = fn arg in
+ let utime_end = current_utime_all () in
+ (output, utime_end -. utime_start);;
+
+let cascaded_scheduler (problem : problem) =
+ let (some_initial_solution, list_scheduler_time) =
+ utime_all_fn (validated_scheduler list_scheduler) problem in
+ match some_initial_solution with
+ | None -> None
+ | Some initial_solution ->
+ let (solution, reoptimizing_time) = utime_all_fn (reoptimizing_scheduler (validated_scheduler (ilp_scheduler SATISFIABILITY)) initial_solution) problem in
+ begin
+ let latency2 = get_max_latency solution
+ and latency1 = get_max_latency initial_solution in
+ Printf.printf "postpass %s: %d, %d, %d, %g, %g\n"
+ (if latency2 < latency1 then "REOPTIMIZED" else "unchanged")
+ (get_nr_instructions problem)
+ latency1 latency2
+ list_scheduler_time reoptimizing_time;
+ flush stdout
+ end;
+ Some solution;;
+
+let scheduler_by_name name =
+ match name with
+ | "ilp" -> validated_scheduler cascaded_scheduler
+ | "list" -> validated_scheduler list_scheduler
+ | "revlist" -> validated_scheduler reverse_list_scheduler
+ | "regpres" -> validated_scheduler reg_pres_scheduler
+ | "regpres_bis" -> validated_scheduler reg_pres_scheduler_bis
+ | "greedy" -> greedy_scheduler
+ | s -> failwith ("unknown scheduler: " ^ s);;
diff --git a/scheduling/InstructionScheduler.mli b/scheduling/InstructionScheduler.mli
new file mode 100644
index 00000000..29b05b6c
--- /dev/null
+++ b/scheduling/InstructionScheduler.mli
@@ -0,0 +1,135 @@
+(** Schedule instructions on a synchronized pipeline
+by David Monniaux, CNRS, VERIMAG *)
+
+(** A latency constraint: instruction number [instr_to] should be scheduled at least [latency] clock ticks before [instr_from].
+
+It is possible to specify [latency]=0, meaning that [instr_to] can be scheduled at the same clock tick as [instr_from], but not before.
+
+[instr_to] can be the special value equal to the number of instructions, meaning that it refers to the final output latency. *)
+type latency_constraint = {
+ instr_from : int;
+ instr_to : int;
+ latency : int;
+ }
+
+(** A scheduling problem.
+
+In addition to the latency constraints, the resource constraints should be satisfied: at every clock tick, the sum of vectors of resources used by the instructions scheduled at that tick does not exceed the resource bounds.
+*)
+type problem = {
+ max_latency : int;
+ (** An optional maximal total latency of the problem, after which the problem is deemed not schedulable. -1 means there should be no maximum. *)
+
+ resource_bounds : int array;
+ (** An array of number of units available indexed by the kind of resources to be allocated. It can be empty, in which case the problem is scheduling without resource constraints. *)
+
+ live_regs_entry : Registers.Regset.t;
+ (** The set of live pseudo-registers at entry. *)
+
+ typing : RTLtyping.regenv;
+ (** Register type map. *)
+
+ reference_counting : ((Registers.reg, int * int) Hashtbl.t
+ * ((Registers.reg * bool) list array)) option;
+ (** 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] *)
+
+ latency_constraints : latency_constraint list
+ (** The latency constraints that must be satisfied *)
+ };;
+
+(** Print problem for human readability. *)
+val print_problem : out_channel -> problem -> unit;;
+
+(** Get the number of instructions in a problem *)
+val get_nr_instructions : problem -> int;;
+
+(** Get the number of resources in a problem *)
+val get_nr_resources : problem -> int;;
+
+(** Scheduling solution. For {i n} instructions to schedule, and 0≤{i i}<{i n}, position {i i} contains the time to which instruction {i i} should be scheduled. Position {i n} contains the final output latency. *)
+type solution = int array
+
+(** A scheduling algorithm.
+The return value [Some x] is a solution [x].
+[None] means that scheduling failed. *)
+type scheduler = problem -> solution option;;
+
+(* DISABLED
+(** Schedule the problem optimally by constraint solving using the Gecode solver. *)
+external gecode_scheduler : problem -> solution option
+ = "caml_gecode_schedule_instr"
+ *)
+
+(** Get the number the last scheduling time used for an instruction in a solution.
+@return The last clock tick used *)
+val maximum_slot_used : solution -> int
+
+(** Validate that a solution is truly a solution of a scheduling problem.
+@raise Failure if validation fails *)
+val check_schedule : problem -> solution -> unit
+
+(** Schedule the problem using a greedy list scheduling algorithm, from the start.
+The first (according to instruction ordering) instruction that is ready (according to the latency constraints) is scheduled at the current clock tick.
+Once a clock tick is full go to the next.
+
+@return [Some solution] when a solution is found, [None] if not. *)
+val list_scheduler : problem -> solution option
+
+(** WIP : Same as list_scheduler, but schedules instructions which decrease
+register pressure when it gets too high. *)
+val reg_pres_scheduler : problem -> solution option
+
+val reg_pres_scheduler_bis : problem -> solution option
+
+(** Schedule the problem using the order of instructions without any reordering *)
+val greedy_scheduler : problem -> solution option
+
+(** Schedule a problem using a scheduler applied in the opposite direction, e.g. for list scheduling from the end instead of the start. BUGGY *)
+val schedule_reversed : scheduler -> problem -> int array option
+
+(** Schedule a problem from the end using a list scheduler. BUGGY *)
+val reverse_list_scheduler : problem -> int array option
+
+(** Check that a problem is well-formed.
+@raise Failure if validation fails *)
+val check_problem : problem -> unit
+
+(** Apply a scheduler and validate the result against the input problem.
+@return The solution found
+@raise Failure if validation fails *)
+val validated_scheduler : scheduler -> problem -> solution option;;
+
+(** Get max latency from solution
+@return Max latency *)
+val get_max_latency : solution -> int;;
+
+(** Get the length of a maximal critical path
+@return Max length *)
+val maximum_critical_path : problem -> int;;
+
+(** Apply line scheduler then advanced solver
+@return A solution if found *)
+val cascaded_scheduler : problem -> solution option;;
+
+val show_date_ranges : problem -> unit;;
+
+type pseudo_boolean_problem_type =
+ | SATISFIABILITY
+ | OPTIMIZATION;;
+
+type pseudo_boolean_mapper
+val pseudo_boolean_print_problem : out_channel -> problem -> pseudo_boolean_problem_type -> pseudo_boolean_mapper;;
+val pseudo_boolean_read_solution : pseudo_boolean_mapper -> in_channel -> solution;;
+val pseudo_boolean_scheduler : pseudo_boolean_problem_type -> problem -> solution option;;
+
+val smt_print_problem : out_channel -> problem -> unit;;
+
+val ilp_print_problem : out_channel -> problem -> pseudo_boolean_problem_type -> pseudo_boolean_mapper;;
+
+val ilp_scheduler : pseudo_boolean_problem_type -> problem -> solution option;;
+
+(** Schedule a problem using a scheduler given by a string name *)
+val scheduler_by_name : string -> problem -> int array option;;
diff --git a/scheduling/PrintBTL.ml b/scheduling/PrintBTL.ml
new file mode 100644
index 00000000..64e83651
--- /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.opt_info with
+ | None -> "none"
+ | Some true -> "branch"
+ | Some false -> "fallthrough");
+ let pref' = pref ^ " " in
+ fprintf pp "%sifso = [\n" pref;
+ if is_rec then print_iblock pp is_rec pref' ib1 else fprintf pp "...\n";
+ fprintf pp "%s]\n" pref;
+ fprintf pp "%sifnot = [\n" pref;
+ if is_rec then print_iblock pp is_rec pref' ib2 else fprintf pp "...\n";
+ fprintf pp "%s]\n" pref
+ | BF (Bjumptable (arg, tbl), 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/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..0ed6ec61
--- /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.opt_info <- 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.opt_info <- 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..6ec32ffc
--- /dev/null
+++ b/scheduling/RTLtoBTLproof.v
@@ -0,0 +1,762 @@
+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. }
+ (* Bload/Bstore *)
+ inv H12; [ idtac | destruct (eval_addressing) eqn:EVAL in LOAD;[ specialize (LOAD v) |] ];
+ rename LOAD into MEMT.
+ 4: rename H12 into EVAL; rename H13 into MEMT.
+ all:
+ erewrite eval_addressing_preserved in EVAL;
+ try erewrite EVAL in BTL_RUN; try erewrite MEMT 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.
diff --git a/scheduling/abstractbb/AbstractBasicBlocksDef.v b/scheduling/abstractbb/AbstractBasicBlocksDef.v
new file mode 100644
index 00000000..fec716c4
--- /dev/null
+++ b/scheduling/abstractbb/AbstractBasicBlocksDef.v
@@ -0,0 +1,471 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** Syntax and Sequential Semantics of Abstract Basic Blocks.
+*)
+Require Import Setoid.
+Require Import ImpPrelude.
+
+Module Type PseudoRegisters.
+
+Parameter t: Type.
+
+Parameter eq_dec: forall (x y: t), { x = y } + { x<>y }.
+
+End PseudoRegisters.
+
+
+(** * Parameters of the language of Basic Blocks *)
+Module Type LangParam.
+
+Declare Module R: PseudoRegisters.
+
+Parameter value: Type.
+
+(** Declare the type of operations *)
+
+Parameter op: Type. (* type of operations *)
+
+Parameter genv: Type. (* environment to be used for evaluating an op *)
+
+Parameter op_eval: genv -> op -> list value -> option value.
+
+End LangParam.
+
+
+
+(** * Syntax and (sequential) semantics of "abstract basic blocks" *)
+Module MkSeqLanguage(P: LangParam).
+
+Export P.
+
+Local Open Scope list.
+
+Section SEQLANG.
+
+Variable ge: genv.
+
+Definition mem := R.t -> value.
+
+Definition assign (m: mem) (x:R.t) (v: value): mem
+ := fun y => if R.eq_dec x y then v else m y.
+
+
+(** Expressions *)
+
+Inductive exp :=
+ | PReg (x:R.t) (**r pseudo-register *)
+ | Op (o:op) (le: list_exp) (**r operation *)
+ | Old (e: exp) (**r evaluation of [e] in the initial state of the instruction (see [inst] below) *)
+with list_exp :=
+ | Enil
+ | Econs (e:exp) (le:list_exp)
+ | LOld (le: list_exp)
+.
+
+Fixpoint exp_eval (e: exp) (m old: mem): option value :=
+ match e with
+ | PReg x => Some (m x)
+ | Op o le =>
+ match list_exp_eval le m old with
+ | Some lv => op_eval ge o lv
+ | _ => None
+ end
+ | Old e => exp_eval e old old
+ end
+with list_exp_eval (le: list_exp) (m old: mem): option (list value) :=
+ match le with
+ | Enil => Some nil
+ | Econs e le' =>
+ match exp_eval e m old, list_exp_eval le' m old with
+ | Some v, Some lv => Some (v::lv)
+ | _, _ => None
+ end
+ | LOld le => list_exp_eval le old old
+ end.
+
+(** An instruction represents a sequence of assignments where [Old] refers to the initial state of the sequence. *)
+Definition inst := list (R.t * exp).
+
+Fixpoint inst_run (i: inst) (m old: mem): option mem :=
+ match i with
+ | nil => Some m
+ | (x, e)::i' =>
+ match exp_eval e m old with
+ | Some v' => inst_run i' (assign m x v') old
+ | None => None
+ end
+ end.
+
+(** A basic block is a sequence of instructions. *)
+Definition bblock := list inst.
+
+Fixpoint run (p: bblock) (m: mem): option mem :=
+ match p with
+ | nil => Some m
+ | i::p' =>
+ match inst_run i m m with
+ | Some m' => run p' m'
+ | None => None
+ end
+ end.
+
+(* A few useful lemma *)
+Lemma assign_eq m x v:
+ (assign m x v) x = v.
+Proof.
+ unfold assign. destruct (R.eq_dec x x); try congruence.
+Qed.
+
+Lemma assign_diff m x y v:
+ x<>y -> (assign m x v) y = m y.
+Proof.
+ unfold assign. destruct (R.eq_dec x y); try congruence.
+Qed.
+
+Lemma assign_skips m x y:
+ (assign m x (m x)) y = m y.
+Proof.
+ unfold assign. destruct (R.eq_dec x y); try congruence.
+Qed.
+
+Lemma assign_swap m x1 v1 x2 v2 y:
+ x1 <> x2 -> (assign (assign m x1 v1) x2 v2) y = (assign (assign m x2 v2) x1 v1) y.
+Proof.
+ intros; destruct (R.eq_dec x2 y).
+ - subst. rewrite assign_eq, assign_diff; auto. rewrite assign_eq; auto.
+ - rewrite assign_diff; auto.
+ destruct (R.eq_dec x1 y).
+ + subst; rewrite! assign_eq. auto.
+ + rewrite! assign_diff; auto.
+Qed.
+
+
+(** A small theory of bblock simulation *)
+
+(* equalities on bblock outputs *)
+Definition res_eq (om1 om2: option mem): Prop :=
+ match om1 with
+ | Some m1 => exists m2, om2 = Some m2 /\ forall x, m1 x = m2 x
+ | None => om2 = None
+ end.
+
+Scheme exp_mut := Induction for exp Sort Prop
+with list_exp_mut := Induction for list_exp Sort Prop.
+
+Lemma exp_equiv e old1 old2:
+ (forall x, old1 x = old2 x) ->
+ forall m1 m2, (forall x, m1 x = m2 x) ->
+ (exp_eval e m1 old1) = (exp_eval e m2 old2).
+Proof.
+ intros H1.
+ induction e using exp_mut with (P0:=fun l => forall m1 m2, (forall x, m1 x = m2 x) -> list_exp_eval l m1 old1 = list_exp_eval l m2 old2); cbn; try congruence; auto.
+ - intros; erewrite IHe; eauto.
+ - intros; erewrite IHe, IHe0; auto.
+Qed.
+
+Definition bblock_simu (p1 p2: bblock): Prop
+ := forall m, (run p1 m) <> None -> res_eq (run p1 m) (run p2 m).
+
+Lemma inst_equiv_refl i old1 old2:
+ (forall x, old1 x = old2 x) ->
+ forall m1 m2, (forall x, m1 x = m2 x) ->
+ res_eq (inst_run i m1 old1) (inst_run i m2 old2).
+Proof.
+ intro H; induction i as [ | [x e]]; cbn; eauto.
+ intros m1 m2 H1. erewrite exp_equiv; eauto.
+ destruct (exp_eval e m2 old2); cbn; auto.
+ apply IHi.
+ unfold assign; intro y. destruct (R.eq_dec x y); auto.
+Qed.
+
+Lemma bblock_equiv_refl p: forall m1 m2, (forall x, m1 x = m2 x) -> res_eq (run p m1) (run p m2).
+Proof.
+ induction p as [ | i p']; cbn; eauto.
+ intros m1 m2 H; lapply (inst_equiv_refl i m1 m2); auto.
+ intros X; lapply (X m1 m2); auto; clear X.
+ destruct (inst_run i m1 m1); cbn.
+ - intros [m3 [H1 H2]]; rewrite H1; cbn; auto.
+ - intros H1; rewrite H1; cbn; auto.
+Qed.
+
+Lemma res_eq_sym om1 om2: res_eq om1 om2 -> res_eq om2 om1.
+Proof.
+ destruct om1; cbn.
+ - intros [m2 [H1 H2]]; subst; cbn. eauto.
+ - intros; subst; cbn; eauto.
+Qed.
+
+Lemma res_eq_trans (om1 om2 om3: option mem):
+ (res_eq om1 om2) -> (res_eq om2 om3) -> (res_eq om1 om3).
+Proof.
+ destruct om1; cbn.
+ - intros [m2 [H1 H2]]; subst; cbn.
+ intros [m3 [H3 H4]]; subst; cbn.
+ eapply ex_intro; intuition eauto. rewrite H2; auto.
+ - intro; subst; cbn; auto.
+Qed.
+
+Lemma bblock_simu_alt p1 p2: bblock_simu p1 p2 <-> (forall m1 m2, (forall x, m1 x = m2 x) -> (run p1 m1)<>None -> res_eq (run p1 m1) (run p2 m2)).
+Proof.
+ unfold bblock_simu; intuition.
+ intros; eapply res_eq_trans. eauto.
+ eapply bblock_equiv_refl; eauto.
+Qed.
+
+
+Lemma run_app p1: forall m1 p2,
+ run (p1++p2) m1 =
+ match run p1 m1 with
+ | Some m2 => run p2 m2
+ | None => None
+ end.
+Proof.
+ induction p1; cbn; try congruence.
+ intros; destruct (inst_run _ _ _); cbn; auto.
+Qed.
+
+Lemma run_app_None p1 m1 p2:
+ run p1 m1 = None ->
+ run (p1++p2) m1 = None.
+Proof.
+ intro H; rewrite run_app. rewrite H; auto.
+Qed.
+
+Lemma run_app_Some p1 m1 m2 p2:
+ run p1 m1 = Some m2 ->
+ run (p1++p2) m1 = run p2 m2.
+Proof.
+ intros H; rewrite run_app. rewrite H; auto.
+Qed.
+
+End SEQLANG.
+
+
+(** * Terms in the symbolic execution *)
+
+(** Such a term represents the successive computations in one given pseudo-register.
+The [hid] has no formal semantics: it is only used by the hash-consing oracle (itself dynamically checked to behave like an identity function).
+
+*)
+
+Module Terms.
+
+Inductive term :=
+ | Input (x:R.t) (hid:hashcode)
+ | App (o: op) (l: list_term) (hid:hashcode)
+with list_term :=
+ | LTnil (hid:hashcode)
+ | LTcons (t:term) (l:list_term) (hid:hashcode)
+ .
+
+Scheme term_mut := Induction for term Sort Prop
+with list_term_mut := Induction for list_term Sort Prop.
+
+Bind Scope pattern_scope with term.
+Delimit Scope term_scope with term.
+Delimit Scope pattern_scope with pattern.
+
+Local Open Scope pattern_scope.
+
+Notation "[ ]" := (LTnil _) (format "[ ]"): pattern_scope.
+Notation "[ x ]" := (LTcons x [ ] _): pattern_scope.
+Notation "[ x ; y ; .. ; z ]" := (LTcons x (LTcons y .. (LTcons z (LTnil _) _) .. _) _): pattern_scope.
+Notation "o @ l" := (App o l _) (at level 50, no associativity): pattern_scope.
+
+Import HConsingDefs.
+
+Notation "[ ]" := (LTnil unknown_hid) (format "[ ]"): term_scope.
+Notation "[ x ]" := (LTcons x []%term unknown_hid): term_scope.
+Notation "[ x ; y ; .. ; z ]" := (LTcons x (LTcons y .. (LTcons z (LTnil unknown_hid) unknown_hid) .. unknown_hid) unknown_hid): term_scope.
+Notation "o @ l" := (App o l unknown_hid) (at level 50, no associativity): term_scope.
+
+
+Fixpoint term_eval (ge: genv) (t: term) (m: mem): option value :=
+ match t with
+ | Input x _ => Some (m x)
+ | o @ l =>
+ match list_term_eval ge l m with
+ | Some v => op_eval ge o v
+ | _ => None
+ end
+ end
+with list_term_eval ge (l: list_term) (m: mem) {struct l}: option (list value) :=
+ match l with
+ | [] => Some nil
+ | LTcons t l' _ =>
+ match term_eval ge t m, list_term_eval ge l' m with
+ | Some v, Some lv => Some (v::lv)
+ | _, _ => None
+ end
+ end.
+
+
+Definition term_get_hid (t: term): hashcode :=
+ match t with
+ | Input _ hid => hid
+ | App _ _ hid => hid
+ end.
+
+Definition list_term_get_hid (l: list_term): hashcode :=
+ match l with
+ | LTnil hid => hid
+ | LTcons _ _ hid => hid
+ end.
+
+
+Fixpoint allvalid ge (l: list term) m : Prop :=
+ match l with
+ | nil => True
+ | t::nil => term_eval ge t m <> None
+ | t::l' => term_eval ge t m <> None /\ allvalid ge l' m
+ end.
+
+Lemma allvalid_extensionality ge (l: list term) m:
+ allvalid ge l m <-> (forall t, List.In t l -> term_eval ge t m <> None).
+Proof.
+ induction l as [|t l]; cbn; try (tauto).
+ destruct l.
+ - intuition (congruence || eauto).
+ - rewrite IHl; clear IHl. intuition (congruence || eauto).
+Qed.
+
+(** * Rewriting rules in the symbolic execution *)
+
+(** The symbolic execution is parametrized by rewriting rules on pseudo-terms. *)
+
+Record pseudo_term: Type := intro_fail {
+ mayfail: list term;
+ effect: term
+}.
+
+(** Simulation relation between a term and a pseudo-term *)
+
+Definition match_pt (t: term) (pt: pseudo_term) :=
+ (forall ge m, term_eval ge t m <> None <-> allvalid ge pt.(mayfail) m)
+ /\ (forall ge m0 m1, term_eval ge t m0 = Some m1 -> term_eval ge pt.(effect) m0 = Some m1).
+
+Lemma inf_option_equivalence (A:Type) (o1 o2: option A):
+ (o1 <> None -> o1 = o2) <-> (forall m1, o1 = Some m1 -> o2 = Some m1).
+Proof.
+ destruct o1; intuition (congruence || eauto).
+ symmetry; eauto.
+Qed.
+
+Lemma intro_fail_correct (l: list term) (t: term) :
+ (forall ge m, term_eval ge t m <> None <-> allvalid ge l m) -> match_pt t (intro_fail l t).
+Proof.
+ unfold match_pt; cbn; intros; intuition congruence.
+Qed.
+Hint Resolve intro_fail_correct: wlp.
+
+(** The default reduction of a term to a pseudo-term *)
+Definition identity_fail (t: term):= intro_fail (t::nil) t.
+
+Lemma identity_fail_correct (t: term): match_pt t (identity_fail t).
+Proof.
+ eapply intro_fail_correct; cbn; tauto.
+Qed.
+Global Opaque identity_fail.
+Hint Resolve identity_fail_correct: wlp.
+
+(** The reduction for constant term *)
+Definition nofail (is_constant: op -> bool) (t: term):=
+ match t with
+ | Input x _ => intro_fail nil t
+ | o @ [] => if is_constant o then (intro_fail nil t) else (identity_fail t)
+ | _ => identity_fail t
+ end.
+
+Lemma nofail_correct (is_constant: op -> bool) t:
+ (forall ge o, is_constant o = true -> op_eval ge o nil <> None) -> match_pt t (nofail is_constant t).
+Proof.
+ destruct t; cbn.
+ + intros; eapply intro_fail_correct; cbn; intuition congruence.
+ + intros; destruct l; cbn; auto with wlp.
+ destruct (is_constant o) eqn:Heqo; cbn; intuition eauto with wlp.
+ eapply intro_fail_correct; cbn; intuition eauto with wlp.
+Qed.
+Global Opaque nofail.
+Hint Resolve nofail_correct: wlp.
+
+(** Term equivalence preserve the simulation by pseudo-terms *)
+Definition term_equiv t1 t2:= forall ge m, term_eval ge t1 m = term_eval ge t2 m.
+
+Global Instance term_equiv_Equivalence : Equivalence term_equiv.
+Proof.
+ split; intro x; unfold term_equiv; intros; eauto.
+ eapply eq_trans; eauto.
+Qed.
+
+Lemma match_pt_term_equiv t1 t2 pt: term_equiv t1 t2 -> match_pt t1 pt -> match_pt t2 pt.
+Proof.
+ unfold match_pt, term_equiv.
+ intros H. intuition; try (erewrite <- H1 in * |- *; congruence).
+ erewrite <- H2; eauto; congruence.
+Qed.
+Hint Resolve match_pt_term_equiv: wlp.
+
+(** Other generic reductions *)
+Definition app_fail (l: list term) (pt: pseudo_term): pseudo_term :=
+ {| mayfail := List.rev_append l pt.(mayfail); effect := pt.(effect) |}.
+
+Lemma app_fail_allvalid_correct l pt t1 t2: forall
+ (V1: forall (ge : genv) (m : mem), term_eval ge t1 m <> None <-> allvalid ge (mayfail pt) m)
+ (V2: forall (ge : genv) (m : mem), term_eval ge t2 m <> None <-> allvalid ge (mayfail {| mayfail := t1 :: l; effect := t1 |}) m)
+ (ge : genv) (m : mem), term_eval ge t2 m <> None <-> allvalid ge (mayfail (app_fail l pt)) m.
+Proof.
+ intros; generalize (V1 ge m) (V2 ge m); rewrite !allvalid_extensionality; cbn. clear V1 V2.
+ intuition subst.
+ + rewrite rev_append_rev, in_app_iff, <- in_rev in H3. destruct H3; eauto.
+ + eapply H3; eauto.
+ intros. intuition subst.
+ * eapply H2; eauto. intros; eapply H0; eauto. rewrite rev_append_rev, in_app_iff; auto.
+ * intros; eapply H0; eauto. rewrite rev_append_rev, in_app_iff, <- in_rev; auto.
+Qed.
+Local Hint Resolve app_fail_allvalid_correct: core.
+
+Lemma app_fail_correct l pt t1 t2:
+ match_pt t1 pt ->
+ match_pt t2 {| mayfail:=t1::l; effect:=t1 |} ->
+ match_pt t2 (app_fail l pt).
+Proof.
+ unfold match_pt in * |- *; intros (V1 & E1) (V2 & E2); split; intros ge m; try (eauto; fail).
+Qed.
+Extraction Inline app_fail.
+
+Import ImpCore.Notations.
+Local Open Scope impure_scope.
+
+(** Specification of rewriting functions in parameter of the symbolic execution: in the impure monad, because the rewriting functions produce hash-consed terms (wrapped in pseudo-terms).
+*)
+Record reduction:= {
+ result:> term -> ?? pseudo_term;
+ result_correct: forall t, WHEN result t ~> pt THEN match_pt t pt;
+}.
+Hint Resolve result_correct: wlp.
+
+End Terms.
+
+End MkSeqLanguage.
+
+
+Module Type SeqLanguage.
+
+Declare Module LP: LangParam.
+
+Include MkSeqLanguage LP.
+
+End SeqLanguage.
+
diff --git a/scheduling/abstractbb/ImpSimuTest.v b/scheduling/abstractbb/ImpSimuTest.v
new file mode 100644
index 00000000..6b64e1d8
--- /dev/null
+++ b/scheduling/abstractbb/ImpSimuTest.v
@@ -0,0 +1,1286 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** Implementation of a simulation test (ie a "scheduling verifier") for the sequential semantics of Abstract Basic Blocks.
+
+It is based on a symbolic execution procedure of Abstract Basic Blocks with imperative hash-consing and rewriting.
+
+It also provides debugging information when the test fails.
+
+
+*)
+
+Require Export Impure.ImpHCons. (**r Import the Impure library. See https://github.com/boulme/ImpureDemo *)
+Export Notations.
+Import HConsing.
+
+Require Import Coq.Bool.Bool.
+Require Export SeqSimuTheory.
+Require Import PArith.
+
+
+Local Open Scope impure.
+
+Import ListNotations.
+Local Open Scope list_scope.
+
+
+Definition FULL_DEBUG_DUMP : bool := false. (* print debug traces, even if the verifier succeeds. *)
+
+(** * Interface of (impure) equality tests for operators *)
+Module Type ImpParam.
+
+Include LangParam.
+
+Parameter op_eq: op -> op -> ?? bool.
+
+Parameter op_eq_correct: forall o1 o2,
+ WHEN op_eq o1 o2 ~> b THEN
+ b=true -> o1 = o2.
+
+End ImpParam.
+
+
+Module Type ISeqLanguage.
+
+Declare Module LP: ImpParam.
+
+Include MkSeqLanguage LP.
+
+End ISeqLanguage.
+
+
+(** * A generic dictinary on PseudoRegisters with an impure equality test *)
+
+Module Type ImpDict.
+
+Declare Module R: PseudoRegisters.
+
+Parameter t: Type -> Type.
+
+Parameter get: forall {A}, t A -> R.t -> option A.
+
+Parameter set: forall {A}, t A -> R.t -> A -> t A.
+
+Parameter set_spec_eq: forall A d x (v: A),
+ get (set d x v) x = Some v.
+
+Parameter set_spec_diff: forall A d x y (v: A),
+ x <> y -> get (set d x v) y = get d y.
+
+Parameter rem: forall {A}, t A -> R.t -> t A.
+
+Parameter rem_spec_eq: forall A (d: t A) x,
+ get (rem d x) x = None.
+
+Parameter rem_spec_diff: forall A (d: t A) x y,
+ x <> y -> get (rem d x) y = get d y.
+
+Parameter empty: forall {A}, t A.
+
+Parameter empty_spec: forall A x,
+ get (empty (A:=A)) x = None.
+
+Parameter eq_test: forall {A}, t A -> t A -> ?? bool.
+
+Parameter eq_test_correct: forall A (d1 d2: t A),
+ WHEN eq_test d1 d2 ~> b THEN
+ b=true -> forall x, get d1 x = get d2 x.
+
+(* NB: we could also take an eq_test on R.t (but not really useful with "pure" dictionaries *)
+
+(** only for debugging *)
+Parameter not_eq_witness: forall {A}, t A -> t A -> ?? option R.t.
+
+End ImpDict.
+
+
+(** * Specification of the provided tests *)
+Module Type ImpSimuInterface.
+
+Declare Module CoreL: ISeqLanguage.
+Import CoreL.
+Import Terms.
+
+(** the silent test (without debugging informations) *)
+Parameter bblock_simu_test: reduction -> bblock -> bblock -> ?? bool.
+
+Parameter bblock_simu_test_correct: forall reduce (p1 p2 : bblock),
+ WHEN bblock_simu_test reduce p1 p2 ~> b
+ THEN b = true -> forall ge : genv, bblock_simu ge p1 p2.
+
+(** the verbose test extended with debugging informations *)
+Parameter verb_bblock_simu_test
+ : reduction ->
+ (R.t -> ?? pstring) ->
+ (op -> ?? pstring) -> bblock -> bblock -> ?? bool.
+
+Parameter verb_bblock_simu_test_correct:
+ forall reduce
+ (string_of_name : R.t -> ?? pstring)
+ (string_of_op : op -> ?? pstring)
+ (p1 p2 : bblock),
+ WHEN verb_bblock_simu_test reduce string_of_name string_of_op p1 p2 ~> b
+ THEN b = true -> forall ge : genv, bblock_simu ge p1 p2.
+
+End ImpSimuInterface.
+
+
+(** * Implementation of the provided tests *)
+
+Module ImpSimu (L: ISeqLanguage) (Dict: ImpDict with Module R:=L.LP.R): ImpSimuInterface with Module CoreL := L.
+
+Module CoreL:=L.
+
+Module ST := SimuTheory L.
+
+Import ST.
+Import Terms.
+
+Definition term_set_hid (t: term) (hid: hashcode): term :=
+ match t with
+ | Input x _ => Input x hid
+ | App op l _ => App op l hid
+ end.
+
+Definition list_term_set_hid (l: list_term) (hid: hashcode): list_term :=
+ match l with
+ | LTnil _ => LTnil hid
+ | LTcons t l' _ => LTcons t l' hid
+ end.
+
+Lemma term_eval_set_hid ge t hid m:
+ term_eval ge (term_set_hid t hid) m = term_eval ge t m.
+Proof.
+ destruct t; cbn; auto.
+Qed.
+
+Lemma list_term_eval_set_hid ge l hid m:
+ list_term_eval ge (list_term_set_hid l hid) m = list_term_eval ge l m.
+Proof.
+ destruct l; cbn; auto.
+Qed.
+
+(* Local nickname *)
+Module D:=ImpPrelude.Dict.
+
+Section SimuWithReduce.
+
+Variable reduce: reduction.
+
+Section CanonBuilding. (** Implementation of the symbolic execution (ie a "canonical form" representing the semantics of an abstract basic block) *)
+
+Variable hC_term: hashinfo term -> ?? term.
+Hypothesis hC_term_correct: forall t, WHEN hC_term t ~> t' THEN forall ge m, term_eval ge (hdata t) m = term_eval ge t' m.
+
+Variable hC_list_term: hashinfo list_term -> ?? list_term.
+Hypothesis hC_list_term_correct: forall t, WHEN hC_list_term t ~> t' THEN forall ge m, list_term_eval ge (hdata t) m = list_term_eval ge t' m.
+
+(* First, we wrap constructors for hashed values !*)
+
+Local Open Scope positive.
+Local Open Scope list_scope.
+
+Definition hInput_hcodes (x:R.t) :=
+ DO hc <~ hash 1;;
+ DO hv <~ hash x;;
+ RET [hc;hv].
+Extraction Inline hInput_hcodes.
+
+Definition hInput (x:R.t): ?? term :=
+ DO hv <~ hInput_hcodes x;;
+ hC_term {| hdata:=Input x unknown_hid; hcodes :=hv; |}.
+
+Lemma hInput_correct x:
+ WHEN hInput x ~> t THEN forall ge m, term_eval ge t m = Some (m x).
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hInput.
+Hint Resolve hInput_correct: wlp.
+
+Definition hApp_hcodes (o:op) (l: list_term) :=
+ DO hc <~ hash 2;;
+ DO hv <~ hash o;;
+ RET [hc;hv;list_term_get_hid l].
+Extraction Inline hApp_hcodes.
+
+Definition hApp (o:op) (l: list_term) : ?? term :=
+ DO hv <~ hApp_hcodes o l;;
+ hC_term {| hdata:=App o l unknown_hid; hcodes:=hv |}.
+
+Lemma hApp_correct o l:
+ WHEN hApp o l ~> t THEN forall ge m,
+ term_eval ge t m = match list_term_eval ge l m with
+ | Some v => op_eval ge o v
+ | None => None
+ end.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hApp.
+Hint Resolve hApp_correct: wlp.
+
+Definition hLTnil (_: unit): ?? list_term :=
+ hC_list_term {| hdata:=LTnil unknown_hid; hcodes := nil; |} .
+
+Lemma hLTnil_correct x:
+ WHEN hLTnil x ~> l THEN forall ge m, list_term_eval ge l m = Some nil.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hLTnil.
+Hint Resolve hLTnil_correct: wlp.
+
+
+Definition hLTcons (t: term) (l: list_term): ?? list_term :=
+ hC_list_term {| hdata:=LTcons t l unknown_hid; hcodes := [term_get_hid t; list_term_get_hid l]; |}.
+
+Lemma hLTcons_correct t l:
+ WHEN hLTcons t l ~> l' THEN forall ge m,
+ list_term_eval ge l' m = match term_eval ge t m, list_term_eval ge l m with
+ | Some v, Some lv => Some (v::lv)
+ | _, _ => None
+ end.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque hLTcons.
+Hint Resolve hLTcons_correct: wlp.
+
+(* Second, we use these hashed constructors ! *)
+
+Record hsmem:= {hpre: list term; hpost:> Dict.t term}.
+
+(** evaluation of the post-condition *)
+Definition hsmem_post_eval ge (hd: Dict.t term) x (m:mem) :=
+ match Dict.get hd x with
+ | None => Some (m x)
+ | Some ht => term_eval ge ht m
+ end.
+
+Definition hsmem_get (d:hsmem) x: ?? term :=
+ match Dict.get d x with
+ | None => hInput x
+ | Some t => RET t
+ end.
+
+Lemma hsmem_get_correct (d:hsmem) x:
+ WHEN hsmem_get d x ~> t THEN forall ge m, term_eval ge t m = hsmem_post_eval ge d x m.
+Proof.
+ unfold hsmem_get, hsmem_post_eval; destruct (Dict.get d x); wlp_simplify.
+Qed.
+Global Opaque hsmem_get.
+Hint Resolve hsmem_get_correct: wlp.
+
+Local Opaque allvalid.
+
+Definition smem_model ge (d: smem) (hd:hsmem): Prop :=
+ (forall m, allvalid ge hd.(hpre) m <-> smem_valid ge d m)
+ /\ (forall m x, smem_valid ge d m -> hsmem_post_eval ge hd x m = (ST.term_eval ge (d x) m)).
+
+Lemma smem_model_smem_valid_alt ge d hd: smem_model ge d hd ->
+ forall m x, smem_valid ge d m -> hsmem_post_eval ge hd x m <> None.
+Proof.
+ intros (H1 & H2) m x H. rewrite H2; auto.
+ unfold smem_valid in H. intuition eauto.
+Qed.
+
+Lemma smem_model_allvalid_alt ge d hd: smem_model ge d hd ->
+ forall m x, allvalid ge hd.(hpre) m -> hsmem_post_eval ge hd x m <> None.
+Proof.
+ intros (H1 & H2) m x H. eapply smem_model_smem_valid_alt.
+ - split; eauto.
+ - rewrite <- H1; auto.
+Qed.
+
+Definition naive_set (hd:hsmem) x (t:term) :=
+ {| hpre:= t::hd.(hpre); hpost:=Dict.set hd x t |}.
+
+Lemma naive_set_correct hd x ht ge d t:
+ smem_model ge d hd ->
+ (forall m, smem_valid ge d m -> term_eval ge ht m = ST.term_eval ge t m) ->
+ smem_model ge (smem_set d x t) (naive_set hd x ht).
+Proof.
+ unfold naive_set; intros (DM0 & DM1) EQT; split.
+ - intros m.
+ destruct (DM0 m) as (PRE & VALID0); clear DM0.
+ assert (VALID1: allvalid ge hd.(hpre) m -> pre d ge m). { unfold smem_valid in PRE; tauto. }
+ assert (VALID2: allvalid ge hd.(hpre) m -> forall x : Dict.R.t, ST.term_eval ge (d x) m <> None). { unfold smem_valid in PRE; tauto. }
+ rewrite !allvalid_extensionality in * |- *; cbn.
+ intuition (subst; eauto).
+ + eapply smem_valid_set_proof; eauto.
+ erewrite <- EQT; eauto.
+ + exploit smem_valid_set_decompose_1; eauto.
+ intros X1; exploit smem_valid_set_decompose_2; eauto.
+ rewrite <- EQT; eauto.
+ + exploit smem_valid_set_decompose_1; eauto.
+ - clear DM0. unfold hsmem_post_eval, hsmem_post_eval in * |- *; cbn.
+ Local Hint Resolve smem_valid_set_decompose_1: core.
+ intros; case (R.eq_dec x x0).
+ + intros; subst; rewrite !Dict.set_spec_eq; cbn; eauto.
+ + intros; rewrite !Dict.set_spec_diff; cbn; eauto.
+Qed.
+Local Hint Resolve naive_set_correct: core.
+
+Definition equiv_hsmem ge (hd1 hd2: hsmem) :=
+ (forall m, allvalid ge hd1.(hpre) m <-> allvalid ge hd2.(hpre) m)
+ /\ (forall m x, allvalid ge hd1.(hpre) m -> hsmem_post_eval ge hd1 x m = hsmem_post_eval ge hd2 x m).
+
+Lemma equiv_smem_symmetry ge hd1 hd2:
+ equiv_hsmem ge hd1 hd2 -> equiv_hsmem ge hd2 hd1.
+Proof.
+ intros (V1 & P1); split.
+ - intros; symmetry; auto.
+ - intros; symmetry; eapply P1. rewrite V1; auto.
+Qed.
+
+Lemma equiv_hsmem_models ge hd1 hd2 d:
+ smem_model ge d hd1 -> equiv_hsmem ge hd1 hd2 -> smem_model ge d hd2.
+Proof.
+ intros (VALID & EQUIV) (HEQUIV & PEQUIV); split.
+ - intros m; rewrite <- VALID; auto. symmetry; auto.
+ - intros m x H. rewrite <- EQUIV; auto.
+ rewrite PEQUIV; auto.
+ rewrite VALID; auto.
+Qed.
+
+Variable log_assign: R.t -> term -> ?? unit.
+
+Definition lift {A B} hid (x:A) (k: B -> ?? A) (y:B): ?? A :=
+ DO b <~ phys_eq hid unknown_hid;;
+ if b then k y else RET x.
+
+Fixpoint hterm_lift (t: term): ?? term :=
+ match t with
+ | Input x hid => lift hid t hInput x
+ | App o l hid =>
+ lift hid t
+ (fun l => DO lt <~ hlist_term_lift l;;
+ hApp o lt) l
+ end
+with hlist_term_lift (l: list_term) {struct l}: ?? list_term :=
+ match l with
+ | LTnil hid => lift hid l hLTnil ()
+ | LTcons t l' hid =>
+ lift hid l
+ (fun t => DO t <~ hterm_lift t;;
+ DO lt <~ hlist_term_lift l';;
+ hLTcons t lt) t
+ end.
+
+Lemma hterm_lift_correct t:
+ WHEN hterm_lift t ~> ht THEN forall ge m, term_eval ge ht m = term_eval ge t m.
+Proof.
+ induction t using term_mut with (P0:=fun lt =>
+ WHEN hlist_term_lift lt ~> hlt THEN forall ge m, list_term_eval ge hlt m = list_term_eval ge lt m);
+ wlp_simplify.
+ - rewrite H0, H; auto.
+ - rewrite H1, H0, H; auto.
+Qed.
+Local Hint Resolve hterm_lift_correct: wlp.
+Global Opaque hterm_lift.
+
+Variable log_new_hterm: term -> ?? unit.
+
+Fixpoint hterm_append (l: list term) (lh: list term): ?? list term :=
+ match l with
+ | nil => RET lh
+ | t::l' =>
+ DO ht <~ hterm_lift t;;
+ log_new_hterm ht;;
+ hterm_append l' (ht::lh)
+ end.
+
+Lemma hterm_append_correct l: forall lh,
+ WHEN hterm_append l lh ~> lh' THEN (forall ge m, allvalid ge lh' m <-> (allvalid ge l m /\ allvalid ge lh m)).
+Proof.
+ Local Hint Resolve eq_trans: localhint.
+ induction l as [|t l']; cbn; wlp_xsimplify ltac:(eauto with wlp).
+ - intros; rewrite! allvalid_extensionality; intuition eauto.
+ - intros REC ge m; rewrite REC; clear IHl' REC. rewrite !allvalid_extensionality.
+ cbn; intuition (subst; eauto with wlp localhint).
+Qed.
+(*Local Hint Resolve hterm_append_correct: wlp.*)
+Global Opaque hterm_append.
+
+Definition smart_set (hd:hsmem) x (ht:term) :=
+ match ht with
+ | Input y _ =>
+ if R.eq_dec x y then
+ RET (Dict.rem hd x)
+ else (
+ log_assign x ht;;
+ RET (Dict.set hd x ht)
+ )
+ | _ =>
+ log_assign x ht;;
+ RET (Dict.set hd x ht)
+ end.
+
+Lemma smart_set_correct hd x ht:
+ WHEN smart_set hd x ht ~> d THEN
+ forall ge m y, hsmem_post_eval ge d y m = hsmem_post_eval ge (Dict.set hd x ht) y m.
+Proof.
+ destruct ht; wlp_simplify.
+ unfold hsmem_post_eval; cbn. case (R.eq_dec x0 y).
+ - intros; subst. rewrite Dict.set_spec_eq, Dict.rem_spec_eq. cbn; congruence.
+ - intros; rewrite Dict.set_spec_diff, Dict.rem_spec_diff; auto.
+Qed.
+(*Local Hint Resolve smart_set_correct: wlp.*)
+Global Opaque smart_set.
+
+Definition hsmem_set (hd:hsmem) x (t:term) :=
+ DO pt <~ reduce t;;
+ DO lht <~ hterm_append pt.(mayfail) hd.(hpre);;
+ DO ht <~ hterm_lift pt.(effect);;
+ log_new_hterm ht;;
+ DO nd <~ smart_set hd x ht;;
+ RET {| hpre := lht; hpost := nd |}.
+
+Lemma hsmem_set_correct hd x ht:
+ WHEN hsmem_set hd x ht ~> nhd THEN
+ forall ge d t, smem_model ge d hd ->
+ (forall m, smem_valid ge d m -> term_eval ge ht m = ST.term_eval ge t m) ->
+ smem_model ge (smem_set d x t) nhd.
+Proof.
+ intros; wlp_simplify.
+ generalize (hterm_append_correct _ _ _ Hexta0); intro APPEND.
+ generalize (hterm_lift_correct _ _ Hexta1); intro LIFT.
+ generalize (smart_set_correct _ _ _ _ Hexta3); intro SMART.
+ eapply equiv_hsmem_models; eauto; unfold equiv_hsmem; cbn.
+ destruct H as (VALID & EFFECT); split.
+ - intros; rewrite APPEND, <- VALID.
+ rewrite !allvalid_extensionality in * |- *; cbn; intuition (subst; eauto).
+ - intros m x0 ALLVALID; rewrite SMART.
+ destruct (term_eval ge ht m) eqn: Hht.
+ * case (R.eq_dec x x0).
+ + intros; subst. unfold hsmem_post_eval; cbn. rewrite !Dict.set_spec_eq.
+ erewrite LIFT, EFFECT; eauto.
+ + intros; unfold hsmem_post_eval; cbn. rewrite !Dict.set_spec_diff; auto.
+ * rewrite allvalid_extensionality in ALLVALID; destruct (ALLVALID ht); cbn; auto.
+Qed.
+Local Hint Resolve hsmem_set_correct: wlp.
+Global Opaque hsmem_set.
+
+(* VARIANTE: we do not hash-cons the term from the expression
+Lemma exp_hterm_correct ge e hod od:
+ smem_model ge od hod ->
+ forall hd d,
+ smem_model ge d hd ->
+ forall m, smem_valid ge d m -> smem_valid ge od m -> term_eval ge (exp_term e hd hod) m = term_eval ge (exp_term e d od) m.
+Proof.
+ intro H.
+ induction e using exp_mut with (P0:=fun le => forall d hd,
+ smem_model ge d hd -> forall m, smem_valid ge d m -> smem_valid ge od m -> list_term_eval ge (list_exp_term le hd hod) m = list_term_eval ge (list_exp_term le d od) m);
+ unfold smem_model in * |- * ; cbn; intuition eauto.
+ - erewrite IHe; eauto.
+ - erewrite IHe0, IHe; eauto.
+Qed.
+Local Hint Resolve exp_hterm_correct: wlp.
+*)
+
+Fixpoint exp_hterm (e: exp) (hd hod: hsmem): ?? term :=
+ match e with
+ | PReg x => hsmem_get hd x
+ | Op o le =>
+ DO lt <~ list_exp_hterm le hd hod;;
+ hApp o lt
+ | Old e => exp_hterm e hod hod
+ end
+with list_exp_hterm (le: list_exp) (hd hod: hsmem): ?? list_term :=
+ match le with
+ | Enil => hLTnil tt
+ | Econs e le' =>
+ DO t <~ exp_hterm e hd hod;;
+ DO lt <~ list_exp_hterm le' hd hod;;
+ hLTcons t lt
+ | LOld le => list_exp_hterm le hod hod
+ end.
+
+Lemma exp_hterm_correct_x ge e hod od:
+ smem_model ge od hod ->
+ forall hd d,
+ smem_model ge d hd ->
+ WHEN exp_hterm e hd hod ~> t THEN forall m, smem_valid ge d m -> smem_valid ge od m -> term_eval ge t m = ST.term_eval ge (exp_term e d od) m.
+ Proof.
+ intro H.
+ induction e using exp_mut with (P0:=fun le => forall d hd,
+ smem_model ge d hd ->
+ WHEN list_exp_hterm le hd hod ~> lt THEN forall m, smem_valid ge d m -> smem_valid ge od m -> list_term_eval ge lt m = ST.list_term_eval ge (list_exp_term le d od) m);
+ unfold smem_model, hsmem_post_eval in * |- * ; cbn; wlp_simplify.
+ - rewrite H1, <- H4; auto.
+ - rewrite H4, <- H0; cbn; auto.
+ - rewrite H5, <- H0, <- H4; cbn; auto.
+Qed.
+Global Opaque exp_hterm.
+
+Lemma exp_hterm_correct e hd hod:
+ WHEN exp_hterm e hd hod ~> t THEN forall ge od d m, smem_model ge od hod -> smem_model ge d hd -> smem_valid ge d m -> smem_valid ge od m -> term_eval ge t m = ST.term_eval ge (exp_term e d od) m.
+Proof.
+ unfold wlp; intros; eapply exp_hterm_correct_x; eauto.
+Qed.
+Hint Resolve exp_hterm_correct: wlp.
+
+Fixpoint hinst_smem (i: inst) (hd hod: hsmem): ?? hsmem :=
+ match i with
+ | nil => RET hd
+ | (x, e)::i' =>
+ DO ht <~ exp_hterm e hd hod;;
+ DO nd <~ hsmem_set hd x ht;;
+ hinst_smem i' nd hod
+ end.
+
+Lemma hinst_smem_correct i: forall hd hod,
+ WHEN hinst_smem i hd hod ~> hd' THEN
+ forall ge od d, smem_model ge od hod -> smem_model ge d hd -> (forall m, smem_valid ge d m -> smem_valid ge od m) -> smem_model ge (inst_smem i d od) hd'.
+Proof.
+ Local Hint Resolve smem_valid_set_proof: core.
+ induction i; cbn; wlp_simplify; eauto 15 with wlp.
+Qed.
+Global Opaque hinst_smem.
+Local Hint Resolve hinst_smem_correct: wlp.
+
+(* logging info: we log the number of inst-instructions passed ! *)
+Variable log_new_inst: unit -> ?? unit.
+
+Fixpoint bblock_hsmem_rec (p: bblock) (d: hsmem): ?? hsmem :=
+ match p with
+ | nil => RET d
+ | i::p' =>
+ log_new_inst tt;;
+ DO d' <~ hinst_smem i d d;;
+ bblock_hsmem_rec p' d'
+ end.
+
+Lemma bblock_hsmem_rec_correct p: forall hd,
+ WHEN bblock_hsmem_rec p hd ~> hd' THEN forall ge d, smem_model ge d hd -> smem_model ge (bblock_smem_rec p d) hd'.
+Proof.
+ induction p; cbn; wlp_simplify.
+Qed.
+Global Opaque bblock_hsmem_rec.
+Local Hint Resolve bblock_hsmem_rec_correct: wlp.
+
+Definition hsmem_empty: hsmem := {| hpre:= nil ; hpost := Dict.empty |}.
+
+Lemma hsmem_empty_correct ge: smem_model ge smem_empty hsmem_empty.
+Proof.
+ unfold smem_model, smem_valid, hsmem_post_eval; cbn; intuition try congruence.
+ rewrite !Dict.empty_spec; cbn; auto.
+Qed.
+
+Definition bblock_hsmem: bblock -> ?? hsmem
+ := fun p => bblock_hsmem_rec p hsmem_empty.
+
+Lemma bblock_hsmem_correct p:
+ WHEN bblock_hsmem p ~> hd THEN forall ge, smem_model ge (bblock_smem p) hd.
+Proof.
+ Local Hint Resolve hsmem_empty_correct: core.
+ wlp_simplify.
+Qed.
+Global Opaque bblock_hsmem.
+
+End CanonBuilding.
+
+(* 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 term_hash_eq (ta tb: term): ?? bool :=
+ match ta, tb with
+ | Input xa _, Input xb _ =>
+ if R.eq_dec xa xb (* Inefficient in some cases ? *)
+ then RET true
+ else RET false
+ | App oa lta _, App ob ltb _ =>
+ DO b <~ op_eq oa ob ;;
+ if b then phys_eq lta ltb
+ else RET false
+ | _,_ => RET false
+ end.
+
+Lemma term_hash_eq_correct: forall ta tb, WHEN term_hash_eq ta tb ~> b THEN b=true -> term_set_hid ta unknown_hid=term_set_hid tb unknown_hid.
+Proof.
+ Local Hint Resolve op_eq_correct: wlp.
+ destruct ta, tb; wlp_simplify; (discriminate || (subst; auto)).
+Qed.
+Global Opaque term_hash_eq.
+Hint Resolve term_hash_eq_correct: wlp.
+
+Definition list_term_hash_eq (lta ltb: list_term): ?? bool :=
+ match lta, ltb with
+ | LTnil _, LTnil _ => RET true
+ | LTcons ta lta _, LTcons tb ltb _ =>
+ DO b <~ phys_eq ta tb ;;
+ if b then phys_eq lta ltb
+ else RET false
+ | _,_ => RET false
+ end.
+
+Lemma list_term_hash_eq_correct: forall lta ltb, WHEN list_term_hash_eq lta ltb ~> b THEN b=true -> list_term_set_hid lta unknown_hid=list_term_set_hid ltb unknown_hid.
+Proof.
+ destruct lta, ltb; wlp_simplify; (discriminate || (subst; auto)).
+Qed.
+Global Opaque list_term_hash_eq.
+Hint Resolve list_term_hash_eq_correct: wlp.
+
+Lemma hsmem_post_eval_intro (d1 d2: hsmem):
+ (forall x, Dict.get d1 x = Dict.get d2 x) -> (forall ge x m, hsmem_post_eval ge d1 x m = hsmem_post_eval ge d2 x m).
+Proof.
+ unfold hsmem_post_eval; intros H ge x m; rewrite H. destruct (Dict.get d2 x); auto.
+Qed.
+
+Local Hint Resolve bblock_hsmem_correct Dict.eq_test_correct: wlp.
+
+Program Definition mk_hash_params (log: term -> ?? unit): Dict.hash_params term :=
+ {|
+ Dict.test_eq := phys_eq;
+ Dict.hashing := fun (ht: term) => RET (term_get_hid ht);
+ Dict.log := log |}.
+Obligation 1.
+ eauto with wlp.
+Qed.
+
+(*** A GENERIC EQ_TEST: IN ORDER TO SUPPORT SEVERAL DEBUGGING MODE !!! ***)
+Definition no_log_assign (x:R.t) (t:term): ?? unit := RET tt.
+Definition no_log_new_term (t:term): ?? unit := RET tt.
+
+Section Prog_Eq_Gen.
+
+Variable log_assign: R.t -> term -> ?? unit.
+Variable log_new_term: hashConsing term -> hashConsing list_term -> ??(term -> ?? unit).
+Variable log_inst1: unit -> ?? unit. (* log of p1 insts *)
+Variable log_inst2: unit -> ?? unit. (* log of p2 insts *)
+
+Variable hco_term: hashConsing term.
+Hypothesis hco_term_correct: forall t, WHEN hco_term.(hC) t ~> t' THEN forall ge m, term_eval ge (hdata t) m = term_eval ge t' m.
+
+Variable hco_list: hashConsing list_term.
+Hypothesis hco_list_correct: forall t, WHEN hco_list.(hC) t ~> t' THEN forall ge m, list_term_eval ge (hdata t) m = list_term_eval ge t' m.
+
+Variable print_end_error: hsmem -> hsmem -> ?? unit.
+Variable print_dump: (option pstring) -> ?? unit.
+
+Variable check_failpreserv: bool.
+Variable dbg_failpreserv: term -> ?? unit. (* info of additional failure of the output bbloc p2 wrt the input bbloc p1 *)
+
+Program Definition g_bblock_simu_test (p1 p2: bblock): ?? bool :=
+ DO failure_in_failpreserv <~ make_cref false;;
+ DO r <~ (TRY
+ DO d1 <~ bblock_hsmem hco_term.(hC) hco_list.(hC) log_assign no_log_new_term log_inst1 p1;;
+ DO log_new_term <~ log_new_term hco_term hco_list;;
+ DO d2 <~ bblock_hsmem hco_term.(hC) hco_list.(hC) no_log_assign log_new_term log_inst2 p2;;
+ DO b <~ Dict.eq_test d1 d2 ;;
+ if b then (
+ (if check_failpreserv then
+ let hp := mk_hash_params dbg_failpreserv in
+ failure_in_failpreserv.(set)(true);;
+ Sets.assert_list_incl hp d2.(hpre) d1.(hpre)
+ else RET tt);;
+ (if FULL_DEBUG_DUMP then
+ print_dump None (* not an error... *)
+ else RET tt);;
+ RET check_failpreserv
+ ) else (
+ print_end_error d1 d2 ;;
+ RET false
+ )
+ CATCH_FAIL s, _ =>
+ DO b <~ failure_in_failpreserv.(get)();;
+ if b then RET false
+ else print_dump (Some s);; RET false
+ ENSURE (fun b => b=true -> forall ge, bblock_simu ge p1 p2));;
+ RET (`r).
+Obligation 1.
+ constructor 1; wlp_simplify; try congruence.
+ destruct (H ge) as (EQPRE1&EQPOST1); destruct (H0 ge) as (EQPRE2&EQPOST2); clear H H0.
+ apply bblock_smem_simu; auto. split.
+ + intros m; rewrite <- EQPRE1, <- EQPRE2.
+ rewrite ! allvalid_extensionality.
+ unfold incl in * |- *; intuition eauto.
+ + intros m0 x VALID; rewrite <- EQPOST1, <- EQPOST2; auto.
+ erewrite hsmem_post_eval_intro; eauto.
+ erewrite <- EQPRE2; auto.
+ erewrite <- EQPRE1 in VALID.
+ rewrite ! allvalid_extensionality in * |- *.
+ unfold incl in * |- *; intuition eauto.
+Qed.
+
+Theorem g_bblock_simu_test_correct p1 p2:
+ WHEN g_bblock_simu_test p1 p2 ~> b THEN b=true -> forall ge, bblock_simu ge p1 p2.
+Proof.
+ wlp_simplify.
+ destruct exta0; cbn in * |- *; auto.
+Qed.
+Global Opaque g_bblock_simu_test.
+
+End Prog_Eq_Gen.
+
+
+
+Definition hpt: hashP term := {| hash_eq := term_hash_eq; get_hid:=term_get_hid; set_hid:=term_set_hid |}.
+Definition hplt: hashP list_term := {| hash_eq := list_term_hash_eq; get_hid:=list_term_get_hid; set_hid:=list_term_set_hid |}.
+
+Definition recover_hcodes (t:term): ??(hashinfo term) :=
+ match t with
+ | Input x _ =>
+ DO hv <~ hInput_hcodes x ;;
+ RET {| hdata := t; hcodes := hv |}
+ | App o l _ =>
+ DO hv <~ hApp_hcodes o l ;;
+ RET {| hdata := t; hcodes := hv |}
+ end.
+
+
+Definition msg_end_of_bblock: pstring :="--- unknown subterms in the graph".
+
+Definition log_new_term
+ (unknownHash_msg: term -> ?? pstring)
+ (hct:hashConsing term)
+ (hcl:hashConsing list_term)
+ : ?? (term -> ?? unit) :=
+ DO clock <~ hct.(next_hid)();;
+ hct.(next_log) msg_end_of_bblock;;
+ hcl.(next_log) msg_end_of_bblock;;
+ RET (fun t =>
+ DO ok <~ hash_older (term_get_hid t) clock;;
+ if ok
+ then
+ RET tt
+ else
+ DO ht <~ recover_hcodes t;;
+ hct.(remove) ht;;
+ DO msg <~ unknownHash_msg t;;
+ FAILWITH msg).
+
+Definition skip (_:unit): ?? unit := RET tt.
+
+Definition msg_prefix: pstring := "*** ERROR INFO from bblock_simu_test: ".
+Definition msg_error_on_end: pstring := "mismatch in final assignments !".
+Definition msg_unknow_term: pstring := "unknown term".
+Definition msg_number: pstring := "on 2nd bblock -- on inst num ".
+Definition msg_notfailpreserv: pstring := "a possible failure of 2nd bblock is absent in 1st bblock (INTERNAL ERROR: this error is expected to be detected before!!!)".
+
+Definition print_end_error (_ _: hsmem): ?? unit
+ := println (msg_prefix +; msg_error_on_end).
+
+Definition print_error (log: logger unit) (os: option pstring): ?? unit
+ := match os with
+ | Some s =>
+ DO n <~ log_info log ();;
+ println (msg_prefix +; msg_number +; n +; " -- " +; s)
+ | None => RET tt
+ end.
+
+Definition failpreserv_error (_: term): ?? unit
+ := println (msg_prefix +; msg_notfailpreserv).
+
+Lemma term_eval_set_hid_equiv ge t1 t2 hid1 hid2 m:
+ term_set_hid t1 hid1 = term_set_hid t2 hid2 -> term_eval ge t1 m = term_eval ge t2 m.
+Proof.
+ intro H; erewrite <- term_eval_set_hid; rewrite H. apply term_eval_set_hid.
+Qed.
+
+Lemma list_term_eval_set_hid_equiv ge t1 t2 hid1 hid2 m:
+ list_term_set_hid t1 hid1 = list_term_set_hid t2 hid2 -> list_term_eval ge t1 m = list_term_eval ge t2 m.
+Proof.
+ intro H; erewrite <- list_term_eval_set_hid; rewrite H. apply list_term_eval_set_hid.
+Qed.
+
+Local Hint Resolve term_eval_set_hid_equiv list_term_eval_set_hid_equiv: core.
+
+Program Definition bblock_simu_test (p1 p2: bblock): ?? bool :=
+ DO log <~ count_logger ();;
+ DO hco_term <~ mk_annot (hCons hpt);;
+ DO hco_list <~ mk_annot (hCons hplt);;
+ g_bblock_simu_test
+ no_log_assign
+ (log_new_term (fun _ => RET msg_unknow_term))
+ skip
+ (log_insert log)
+ hco_term _
+ hco_list _
+ print_end_error
+ (print_error log)
+ true (* check_failpreserv *)
+ failpreserv_error
+ p1 p2.
+Obligation 1.
+ generalize (hCons_correct _ _ _ H0); clear H0.
+ wlp_simplify.
+Qed.
+Obligation 2.
+ generalize (hCons_correct _ _ _ H); clear H.
+ wlp_simplify.
+Qed.
+
+Local Hint Resolve g_bblock_simu_test_correct: core.
+
+Theorem bblock_simu_test_correct p1 p2:
+ WHEN bblock_simu_test p1 p2 ~> b THEN b=true -> forall ge, bblock_simu ge p1 p2.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque bblock_simu_test.
+
+(** This is only to print info on each bblock_simu_test run **)
+Section Verbose_version.
+
+Variable string_of_name: R.t -> ?? pstring.
+Variable string_of_op: op -> ?? pstring.
+
+
+Local Open Scope string_scope.
+
+Definition string_term_hid (t: term): ?? pstring :=
+ DO id <~ string_of_hashcode (term_get_hid t);;
+ RET ("E" +; (CamlStr id)).
+
+Definition string_list_hid (lt: list_term): ?? pstring :=
+ DO id <~ string_of_hashcode (list_term_get_hid lt);;
+ RET ("L" +; (CamlStr id)).
+
+Definition print_raw_term (t: term): ?? unit :=
+ match t with
+ | Input x _ =>
+ DO s <~ string_of_name x;;
+ println( "init_access " +; s)
+ | App o (LTnil _) _ =>
+ DO so <~ string_of_op o;;
+ println so
+ | App o l _ =>
+ DO so <~ string_of_op o;;
+ DO sl <~ string_list_hid l;;
+ println (so +; " " +; sl)
+ end.
+
+(*
+Definition print_raw_list(lt: list_term): ?? unit :=
+ match lt with
+ | LTnil _=> println ""
+ | LTcons t l _ =>
+ DO st <~ string_term_hid t;;
+ DO sl <~ string_list_hid l;;
+ println(st +; " " +; sl)
+ end.
+*)
+
+Section PrettryPrint.
+
+Variable get_debug_info: term -> ?? option pstring.
+
+Fixpoint string_of_term (t: term): ?? pstring :=
+ match t with
+ | Input x _ => string_of_name x
+ | App o (LTnil _) _ => string_of_op o
+ | App o l _ =>
+ DO so <~ string_of_op o;;
+ DO sl <~ string_of_list_term l;;
+ RET (so +; "[" +; sl +; "]")
+ end
+with string_of_list_term (l: list_term): ?? pstring :=
+ match l with
+ | LTnil _ => RET (Str "")
+ | LTcons t (LTnil _) _ =>
+ DO dbg <~ get_debug_info t;;
+ match dbg with
+ | Some x => RET x
+ | None => string_of_term t
+ end
+ | LTcons t l' _ =>
+ DO st <~ (DO dbg <~ get_debug_info t;;
+ match dbg with
+ | Some x => RET x
+ | None => string_of_term t
+ end);;
+ DO sl <~ string_of_list_term l';;
+ RET (st +; ";" +; sl)
+ end.
+
+End PrettryPrint.
+
+
+Definition pretty_term gdi t :=
+ DO r <~ string_of_term gdi t;;
+ println(r).
+
+Fixpoint print_head (head: list pstring): ?? unit :=
+ match head with
+ | i::head' => println (i);; print_head head'
+ | _ => RET tt
+ end.
+
+Definition print_term gdi (head: list pstring) (t: term): ?? unit :=
+ print_head head;;
+ DO s <~ string_term_hid t;;
+ print (s +; ": ");;
+ print_raw_term t;;
+ DO dbg <~ gdi t;;
+ match dbg with
+ | Some x =>
+ print("// " +; x +; " <- ");;
+ pretty_term gdi t
+ | None => RET tt
+ end.
+
+Definition print_list gdi (head: list pstring) (lt: list_term): ?? unit :=
+ print_head head;;
+ DO s <~ string_list_hid lt ;;
+ print (s +; ": ");;
+ (* print_raw_list lt;; *)
+ DO ps <~ string_of_list_term gdi lt;;
+ println("[" +; ps +; "]").
+
+
+Definition print_tables gdi ext exl: ?? unit :=
+ println "-- term table --" ;;
+ iterall ext (fun head _ pt => print_term gdi head pt.(hdata));;
+ println "-- list table --" ;;
+ iterall exl (fun head _ pl => print_list gdi head pl.(hdata));;
+ println "----------------".
+
+Definition print_final_debug gdi (d1 d2: hsmem): ?? unit
+ := DO b <~ Dict.not_eq_witness d1 d2 ;;
+ match b with
+ | Some x =>
+ DO s <~ string_of_name x;;
+ println("mismatch on: " +; s);;
+ match Dict.get d1 x with
+ | None => println("=> unassigned in 1st bblock")
+ | Some t1 =>
+ print("=> node expected from 1st bblock: ");;
+ pretty_term gdi t1
+ end;;
+ match Dict.get d2 x with
+ | None => println("=> unassigned in 2nd bblock")
+ | Some t2 =>
+ print("=> node found from 2nd bblock: ");;
+ pretty_term gdi t2
+ end
+ | None => FAILWITH "bug in Dict.not_eq_witness ?"
+ end.
+
+Definition witness:= option term.
+
+Definition msg_term (cr: cref witness) t :=
+ set cr (Some t);;
+ RET msg_unknow_term.
+
+Definition print_witness gdi cr (*msg*) :=
+ DO wit <~ get cr ();;
+ match wit with
+ | Some t =>
+ println("=> unknown term node: ");;
+ pretty_term gdi t (*;;
+ println("=> encoded on " +; msg +; " graph as: ");;
+ print_raw_term t *)
+ | None => println "Unexpected failure: no witness info (hint: hash-consing bug ?)"
+ end.
+
+Definition print_end_error1 gdi hct hcl (d1 d2:hsmem): ?? unit
+ := println "- GRAPH of 1st bblock";;
+ DO ext <~ export hct ();;
+ DO exl <~ export hcl ();;
+ print_tables gdi ext exl;;
+ print_end_error d1 d2;;
+ print_final_debug gdi d1 d2.
+
+Definition print_dump1 gdi hct hcl cr log os : ?? unit
+ := println "- GRAPH of 1st bblock";;
+ DO ext <~ export hct ();;
+ DO exl <~ export hcl ();;
+ print_tables gdi ext exl;;
+ print_error log os;;
+ match os with
+ | Some _ => print_witness gdi cr (*"1st"*)
+ | None => RET tt
+ end.
+
+Definition xmsg_number: pstring := "on 1st bblock -- on inst num ".
+
+Definition print_end_error2 gdi hct hcl (d1 d2:hsmem): ?? unit
+ := println (msg_prefix +; msg_error_on_end);;
+ println "- GRAPH of 2nd bblock";;
+ DO ext <~ export hct ();;
+ DO exl <~ export hcl ();;
+ print_tables gdi ext exl.
+
+Definition print_dump2 gdi hct hcl cr (log: logger unit) (os:option pstring): ?? unit
+ := DO n <~ log_info log ();;
+ DO ext <~ export hct ();;
+ DO exl <~ export hcl ();;
+ match os with
+ | Some s =>
+ println (msg_prefix +; xmsg_number +; n +; " -- " +; s);;
+ print_witness gdi cr (*"2nd"*)
+ | None => RET tt
+ end;;
+ println "- GRAPH of 2nd bblock";;
+ print_tables gdi ext exl.
+
+(* USELESS
+Definition simple_log_assign (d: D.t term pstring) (x: R.t) (t: term): ?? unit :=
+ DO s <~ string_of_name x;;
+ d.(D.set) (t,s).
+*)
+
+Definition log_assign (d: D.t term pstring) (log: logger unit) (x: R.t) (t: term): ?? unit :=
+ DO i <~ log_info log ();;
+ DO sx <~ string_of_name x;;
+ d.(D.set) (t,(sx +; "@" +; i)).
+
+Definition msg_new_inst : pstring := "--- inst ".
+
+Definition hlog (log: logger unit) (hct: hashConsing term) (hcl: hashConsing list_term): unit -> ?? unit :=
+ (fun _ =>
+ log_insert log tt ;;
+ DO s <~ log_info log tt;;
+ let s:= msg_new_inst +; s in
+ next_log hct s;;
+ next_log hcl s
+ ).
+
+Program Definition verb_bblock_simu_test (p1 p2: bblock): ?? bool :=
+ DO dict_info <~ make_dict (mk_hash_params (fun _ => RET tt));;
+ DO log1 <~ count_logger ();;
+ DO log2 <~ count_logger ();;
+ DO cr <~ make_cref None;;
+ DO hco_term <~ mk_annot (hCons hpt);;
+ DO hco_list <~ mk_annot (hCons hplt);;
+ (if FULL_DEBUG_DUMP then
+ println("");;
+ println("-- START simu checker --")
+ else RET tt);;
+ DO result1 <~ (g_bblock_simu_test
+ (log_assign dict_info log1)
+ (log_new_term (msg_term cr))
+ (hlog log1 hco_term hco_list)
+ (log_insert log2)
+ hco_term _
+ hco_list _
+ (print_end_error1 dict_info.(D.get) hco_term hco_list)
+ (print_dump1 dict_info.(D.get) hco_term hco_list cr log2)
+ true
+ failpreserv_error
+ p1 p2);;
+ if (if FULL_DEBUG_DUMP then false else result1)
+ then RET true
+ else (
+ DO dict_info <~ make_dict (mk_hash_params (fun _ => RET tt));;
+ DO log1 <~ count_logger ();;
+ DO log2 <~ count_logger ();;
+ DO cr <~ make_cref None;;
+ DO hco_term <~ mk_annot (hCons hpt);;
+ DO hco_list <~ mk_annot (hCons hplt);;
+ DO result2 <~ g_bblock_simu_test
+ (log_assign dict_info log1)
+ (*fun _ _ => RET no_log_new_term*) (* REM: too weak !! *)
+ (log_new_term (msg_term cr)) (* REM: too strong ?? *)
+ (hlog log1 hco_term hco_list)
+ (log_insert log2)
+ hco_term _
+ hco_list _
+ (print_end_error2 dict_info.(D.get) hco_term hco_list)
+ (print_dump2 dict_info.(D.get) hco_term hco_list cr log2)
+ false
+ (fun _ => RET tt)
+ p2 p1;;
+ if FULL_DEBUG_DUMP then
+ println("-- END simu checker --");;
+ println("");;
+ RET result1
+ else if result2
+ then (
+ println (msg_prefix +; " OOops - symmetry violation in bblock_simu_test => this is a bug of bblock_simu_test ??");;
+ RET false
+ ) else RET false
+ ).
+Obligation 1.
+ generalize (hCons_correct _ _ _ H1); clear H1.
+ wlp_simplify.
+Qed.
+Obligation 2.
+ generalize (hCons_correct _ _ _ H0); clear H0.
+ wlp_simplify.
+Qed.
+Obligation 3.
+ generalize (hCons_correct _ _ _ H1); clear H1.
+ wlp_simplify.
+Qed.
+Obligation 4.
+ generalize (hCons_correct _ _ _ H0); clear H0.
+ wlp_simplify.
+Qed.
+
+Theorem verb_bblock_simu_test_correct p1 p2:
+ WHEN verb_bblock_simu_test p1 p2 ~> b THEN b=true -> forall ge, bblock_simu ge p1 p2.
+Proof.
+ wlp_simplify.
+Qed.
+Global Opaque verb_bblock_simu_test.
+
+End Verbose_version.
+
+End SimuWithReduce.
+
+(* TODO: why inlining fails here ? *)
+Transparent hterm_lift.
+Extraction Inline lift.
+
+End ImpSimu.
+
+
+(** * Implementation of the Dictionary (based on PositiveMap) *)
+Require Import FMapPositive.
+Require Import PArith.
+Require Import FMapPositive.
+
+Module ImpPosDict <: ImpDict with Module R:=Pos.
+
+Module R:=Pos.
+
+Definition t:=PositiveMap.t.
+
+Definition get {A} (d:t A) (x:R.t): option A
+ := PositiveMap.find x d.
+
+Definition set {A} (d:t A) (x:R.t) (v:A): t A
+ := PositiveMap.add x v d.
+
+Local Hint Unfold PositiveMap.E.eq: core.
+
+Lemma set_spec_eq A d x (v: A):
+ get (set d x v) x = Some v.
+Proof.
+ unfold get, set; apply PositiveMap.add_1; auto.
+Qed.
+
+Lemma set_spec_diff A d x y (v: A):
+ x <> y -> get (set d x v) y = get d y.
+Proof.
+ unfold get, set; intros; apply PositiveMap.gso; auto.
+Qed.
+
+Definition rem {A} (d:t A) (x:R.t): t A
+ := PositiveMap.remove x d.
+
+Lemma rem_spec_eq A (d: t A) x:
+ get (rem d x) x = None.
+Proof.
+ unfold get, rem; apply PositiveMap.grs; auto.
+Qed.
+
+Lemma rem_spec_diff A (d: t A) x y:
+ x <> y -> get (rem d x) y = get d y.
+Proof.
+ unfold get, rem; intros; apply PositiveMap.gro; auto.
+Qed.
+
+
+Definition empty {A}: t A := PositiveMap.empty A.
+
+Lemma empty_spec A x:
+ get (empty (A:=A)) x = None.
+Proof.
+ unfold get, empty; apply PositiveMap.gempty; auto.
+Qed.
+
+Import PositiveMap.
+
+Fixpoint eq_test {A} (d1 d2: t A): ?? bool :=
+ match d1, d2 with
+ | Leaf _, Leaf _ => RET true
+ | Node l1 (Some x1) r1, Node l2 (Some x2) r2 =>
+ DO b0 <~ phys_eq x1 x2 ;;
+ if b0 then
+ DO b1 <~ eq_test l1 l2 ;;
+ if b1 then
+ eq_test r1 r2
+ else
+ RET false
+ else
+ RET false
+ | Node l1 None r1, Node l2 None r2 =>
+ DO b1 <~ eq_test l1 l2 ;;
+ if b1 then
+ eq_test r1 r2
+ else
+ RET false
+ | _, _ => RET false
+ end.
+
+Lemma eq_test_correct A d1: forall (d2: t A),
+ WHEN eq_test d1 d2 ~> b THEN
+ b=true -> forall x, get d1 x = get d2 x.
+Proof.
+ unfold get; induction d1 as [|l1 Hl1 [x1|] r1 Hr1]; destruct d2 as [|l2 [x2|] r2]; cbn;
+ wlp_simplify; (discriminate || (subst; destruct x; cbn; auto)).
+Qed.
+Global Opaque eq_test.
+
+(** ONLY FOR DEBUGGING INFO: get some key of a non-empty d *)
+Fixpoint pick {A} (d: t A): ?? R.t :=
+ match d with
+ | Leaf _ => FAILWITH "unexpected empty dictionary"
+ | Node _ (Some _) _ => RET xH
+ | Node (Leaf _) None r =>
+ DO p <~ pick r;;
+ RET (xI p)
+ | Node l None _ =>
+ DO p <~ pick l;;
+ RET (xO p)
+ end.
+
+(** ONLY FOR DEBUGGING INFO: find one variable on which d1 and d2 differs *)
+Fixpoint not_eq_witness {A} (d1 d2: t A): ?? option R.t :=
+ match d1, d2 with
+ | Leaf _, Leaf _ => RET None
+ | Node l1 (Some x1) r1, Node l2 (Some x2) r2 =>
+ DO b0 <~ phys_eq x1 x2 ;;
+ if b0 then
+ DO b1 <~ not_eq_witness l1 l2;;
+ match b1 with
+ | None =>
+ DO b2 <~ not_eq_witness r1 r2;;
+ match b2 with
+ | None => RET None
+ | Some p => RET (Some (xI p))
+ end
+ | Some p => RET (Some (xO p))
+ end
+ else
+ RET (Some xH)
+ | Node l1 None r1, Node l2 None r2 =>
+ DO b1 <~ not_eq_witness l1 l2;;
+ match b1 with
+ | None =>
+ DO b2 <~ not_eq_witness r1 r2;;
+ match b2 with
+ | None => RET None
+ | Some p => RET (Some (xI p))
+ end
+ | Some p => RET (Some (xO p))
+ end
+ | l, Leaf _ => DO p <~ pick l;; RET (Some p)
+ | Leaf _, r => DO p <~ pick r;; RET (Some p)
+ | _, _ => RET (Some xH)
+ end.
+
+End ImpPosDict.
+
diff --git a/scheduling/abstractbb/Parallelizability.v b/scheduling/abstractbb/Parallelizability.v
new file mode 100644
index 00000000..afa4b9fd
--- /dev/null
+++ b/scheduling/abstractbb/Parallelizability.v
@@ -0,0 +1,792 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** Parallel Semantics of Abstract Basic Blocks and parallelizability test.
+*)
+
+Require Setoid. (* in order to rewrite <-> *)
+Require Export AbstractBasicBlocksDef.
+
+Require Import List.
+Import ListNotations.
+Local Open Scope list_scope.
+
+Require Import Sorting.Permutation.
+Require Import Bool.
+Local Open Scope lazy_bool_scope.
+
+(** * Definition of the parallel semantics *)
+Module ParallelSemantics (L: SeqLanguage).
+
+Export L.
+Local Open Scope list.
+
+Section PARALLEL.
+Variable ge: genv.
+
+(* parallel run of a inst *)
+Fixpoint inst_prun (i: inst) (m tmp old: mem): option mem :=
+ match i with
+ | nil => Some m
+ | (x, e)::i' =>
+ match exp_eval ge e tmp old with
+ | Some v' => inst_prun i' (assign m x v') (assign tmp x v') old
+ | None => None
+ end
+ end.
+
+(* [inst_prun] is generalization of [inst_run] *)
+Lemma inst_run_prun i: forall m old,
+ inst_run ge i m old = inst_prun i m m old.
+Proof.
+ induction i as [|[y e] i']; cbn; auto.
+ intros m old; destruct (exp_eval ge e m old); cbn; auto.
+Qed.
+
+
+(* parallel run of a bblock -- with in-order writes *)
+Fixpoint prun_iw (p: bblock) m old: option mem :=
+ match p with
+ | nil => Some m
+ | i::p' =>
+ match inst_prun i m old old with
+ | Some m1 => prun_iw p' m1 old
+ | None => None
+ end
+ end.
+
+(* non-deterministic parallel run, due to arbitrary writes order *)
+Definition prun (p: bblock) m (om: option mem) := exists p', res_eq om (prun_iw p' m m) /\ Permutation p p'.
+
+
+(* a few lemma on equality *)
+
+Lemma inst_prun_equiv i old: forall m1 m2 tmp,
+ (forall x, m1 x = m2 x) ->
+ res_eq (inst_prun i m1 tmp old) (inst_prun i m2 tmp old).
+Proof.
+ induction i as [|[x e] i']; cbn; eauto.
+ intros m1 m2 tmp H; destruct (exp_eval ge e tmp old); cbn; auto.
+ eapply IHi'; unfold assign. intros; destruct (R.eq_dec x x0); auto.
+Qed.
+
+Lemma prun_iw_equiv p: forall m1 m2 old,
+ (forall x, m1 x = m2 x) ->
+ res_eq (prun_iw p m1 old) (prun_iw p m2 old).
+Proof.
+ induction p as [|i p']; cbn; eauto.
+ - intros m1 m2 old H.
+ generalize (inst_prun_equiv i old m1 m2 old H);
+ destruct (inst_prun i m1 old old); cbn.
+ + intros (m3 & H3 & H4); rewrite H3; cbn; eauto.
+ + intros H1; rewrite H1; cbn; auto.
+Qed.
+
+
+Lemma prun_iw_app p1: forall m1 old p2,
+ prun_iw (p1++p2) m1 old =
+ match prun_iw p1 m1 old with
+ | Some m2 => prun_iw p2 m2 old
+ | None => None
+ end.
+Proof.
+ induction p1; cbn; try congruence.
+ intros; destruct (inst_prun _ _ _); cbn; auto.
+Qed.
+
+Lemma prun_iw_app_None p1: forall m1 old p2,
+ prun_iw p1 m1 old = None ->
+ prun_iw (p1++p2) m1 old = None.
+Proof.
+ intros m1 old p2 H; rewrite prun_iw_app. rewrite H; auto.
+Qed.
+
+Lemma prun_iw_app_Some p1: forall m1 old m2 p2,
+ prun_iw p1 m1 old = Some m2 ->
+ prun_iw (p1++p2) m1 old = prun_iw p2 m2 old.
+Proof.
+ intros m1 old m2 p2 H; rewrite prun_iw_app. rewrite H; auto.
+Qed.
+
+End PARALLEL.
+End ParallelSemantics.
+
+
+
+Fixpoint notIn {A} (x: A) (l:list A): Prop :=
+ match l with
+ | nil => True
+ | a::l' => x <> a /\ notIn x l'
+ end.
+
+Lemma notIn_iff A (x:A) l: (~List.In x l) <-> notIn x l.
+Proof.
+ induction l; cbn; intuition.
+Qed.
+
+Lemma notIn_app A (x:A) l1: forall l2, notIn x (l1++l2) <-> (notIn x l1 /\ notIn x l2).
+Proof.
+ induction l1; cbn.
+ - intuition.
+ - intros; rewrite IHl1. intuition.
+Qed.
+
+
+Lemma In_Permutation A (l1 l2: list A): Permutation l1 l2 -> forall x, In x l1 -> In x l2.
+Proof.
+ induction 1; cbn; intuition.
+Qed.
+
+Lemma Permutation_incl A (l1 l2: list A): Permutation l1 l2 -> incl l1 l2.
+Proof.
+ unfold incl; intros; eapply In_Permutation; eauto.
+Qed.
+
+Lemma notIn_incl A (l1 l2: list A) x: incl l1 l2 -> notIn x l2 -> notIn x l1.
+Proof.
+ unfold incl; rewrite <- ! notIn_iff; intuition.
+Qed.
+
+
+Definition disjoint {A: Type} (l l':list A) : Prop := forall x, In x l -> notIn x l'.
+
+Lemma disjoint_sym_imp A (l1 l2: list A): disjoint l1 l2 -> disjoint l2 l1.
+Proof.
+ unfold disjoint. intros H x H1. generalize (H x). rewrite <- !notIn_iff. intuition.
+Qed.
+
+Lemma disjoint_sym A (l1 l2: list A): disjoint l1 l2 <-> disjoint l2 l1.
+Proof.
+ constructor 1; apply disjoint_sym_imp; auto.
+Qed.
+
+
+Lemma disjoint_cons_l A (x:A) (l1 l2: list A): disjoint (x::l1) l2 <-> (notIn x l2) /\ (disjoint l1 l2).
+Proof.
+ unfold disjoint. cbn; intuition subst; auto.
+Qed.
+
+Lemma disjoint_cons_r A (x:A) (l1 l2: list A): disjoint l1 (x::l2) <-> (notIn x l1) /\ (disjoint l1 l2).
+Proof.
+ rewrite disjoint_sym, disjoint_cons_l, disjoint_sym; intuition.
+Qed.
+
+Lemma disjoint_app_r A (l l1 l2: list A): disjoint l (l1++l2) <-> (disjoint l l1 /\ disjoint l l2).
+Proof.
+ unfold disjoint. intuition.
+ - generalize (H x H0). rewrite notIn_app; intuition.
+ - generalize (H x H0). rewrite notIn_app; intuition.
+ - rewrite notIn_app; intuition.
+Qed.
+
+Lemma disjoint_app_l A (l l1 l2: list A): disjoint (l1++l2) l <-> (disjoint l1 l /\ disjoint l2 l).
+Proof.
+ rewrite disjoint_sym, disjoint_app_r; intuition; rewrite disjoint_sym; auto.
+Qed.
+
+Lemma disjoint_incl_r A (l1 l2: list A): incl l1 l2 -> forall l, disjoint l l2 -> disjoint l l1.
+Proof.
+ unfold disjoint. intros; eapply notIn_incl; eauto.
+Qed.
+
+Lemma disjoint_incl_l A (l1 l2: list A): incl l1 l2 -> forall l, disjoint l2 l -> disjoint l1 l.
+Proof.
+ intros; rewrite disjoint_sym. eapply disjoint_incl_r; eauto. rewrite disjoint_sym; auto.
+Qed.
+
+
+Module ParallelizablityChecking (L: SeqLanguage).
+
+Include ParallelSemantics L.
+
+Section PARALLELI.
+Variable ge: genv.
+
+(** * Preliminary notions on frames *)
+
+Lemma notIn_dec (x: R.t) l : { notIn x l } + { In x l }.
+Proof.
+ destruct (In_dec R.eq_dec x l). constructor 2; auto.
+ constructor 1; rewrite <- notIn_iff. auto.
+Qed.
+
+Fixpoint frame_assign m1 (f: list R.t) m2 :=
+ match f with
+ | nil => m1
+ | x::f' => frame_assign (assign m1 x (m2 x)) f' m2
+ end.
+
+Lemma frame_assign_def f: forall m1 m2 x,
+ frame_assign m1 f m2 x = if notIn_dec x f then m1 x else m2 x.
+Proof.
+ induction f as [|y f] ; cbn; auto.
+ - intros; destruct (notIn_dec x []); cbn in *; tauto.
+ - intros; rewrite IHf; destruct (notIn_dec x (y::f)); cbn in *.
+ + destruct (notIn_dec x f); cbn in *; intuition.
+ rewrite assign_diff; auto.
+ rewrite <- notIn_iff in *; intuition.
+ + destruct (notIn_dec x f); cbn in *; intuition subst.
+ rewrite assign_eq; auto.
+ rewrite <- notIn_iff in *; intuition.
+Qed.
+
+Lemma frame_assign_In m1 f m2 x:
+ In x f -> frame_assign m1 f m2 x = m2 x.
+Proof.
+ intros; rewrite frame_assign_def; destruct (notIn_dec x f); auto.
+ rewrite <- notIn_iff in *; intuition.
+Qed.
+
+Lemma frame_assign_notIn m1 f m2 x:
+ notIn x f -> frame_assign m1 f m2 x = m1 x.
+Proof.
+ intros; rewrite frame_assign_def; destruct (notIn_dec x f); auto.
+ rewrite <- notIn_iff in *; intuition.
+Qed.
+
+Definition frame_eq (frame: R.t -> Prop) (om1 om2: option mem): Prop :=
+ match om1 with
+ | Some m1 => exists m2, om2 = Some m2 /\ forall x, (frame x) -> m1 x = m2 x
+ | None => om2 = None
+ end.
+
+Lemma frame_eq_list_split f1 (f2: R.t -> Prop) om1 om2:
+ frame_eq (fun x => In x f1) om1 om2 ->
+ (forall m1 m2 x, om1 = Some m1 -> om2 = Some m2 -> f2 x -> notIn x f1 -> m1 x = m2 x) ->
+ frame_eq f2 om1 om2.
+Proof.
+ unfold frame_eq; destruct om1 as [ m1 | ]; cbn; auto.
+ intros (m2 & H0 & H1); subst.
+ intros H.
+ eexists; intuition eauto.
+ destruct (notIn_dec x f1); auto.
+Qed.
+
+(*
+Lemma frame_eq_res_eq f om1 om2:
+ frame_eq (fun x => In x f) om1 om2 ->
+ (forall m1 m2 x, om1 = Some m1 -> om2 = Some m2 -> notIn x f -> m1 x = m2 x) ->
+ res_eq om1 om2.
+Proof.
+ intros H H0; lapply (frame_eq_list_split f (fun _ => True) om1 om2 H); eauto.
+ clear H H0; unfold frame_eq, res_eq. destruct om1; cbn; firstorder.
+Qed.
+*)
+
+(** * Writing frames *)
+
+Fixpoint inst_wframe(i:inst): list R.t :=
+ match i with
+ | nil => nil
+ | a::i' => (fst a)::(inst_wframe i')
+ end.
+
+Lemma inst_wframe_correct i m' old: forall m tmp,
+ inst_prun ge i m tmp old = Some m' ->
+ forall x, notIn x (inst_wframe i) -> m' x = m x.
+Proof.
+ induction i as [|[y e] i']; cbn.
+ - intros m tmp H x H0; inversion_clear H; auto.
+ - intros m tmp H x (H1 & H2); destruct (exp_eval ge e tmp old); cbn; try congruence.
+ replace (m x) with (assign m y v x); eauto.
+ rewrite assign_diff; auto.
+Qed.
+
+Lemma inst_prun_fequiv i old: forall m1 m2 tmp,
+ frame_eq (fun x => In x (inst_wframe i)) (inst_prun ge i m1 tmp old) (inst_prun ge i m2 tmp old).
+Proof.
+ induction i as [|[y e] i']; cbn.
+ - intros m1 m2 tmp; eexists; intuition eauto.
+ - intros m1 m2 tmp. destruct (exp_eval ge e tmp old); cbn; auto.
+ eapply frame_eq_list_split; eauto. clear IHi'.
+ intros m1' m2' x H1 H2.
+ lapply (inst_wframe_correct i' m1' old (assign m1 y v) (assign tmp y v)); eauto.
+ lapply (inst_wframe_correct i' m2' old (assign m2 y v) (assign tmp y v)); eauto.
+ intros Xm2 Xm1 H H0. destruct H.
+ + subst. rewrite Xm1, Xm2; auto. rewrite !assign_eq. auto.
+ + rewrite <- notIn_iff in H0; tauto.
+Qed.
+
+Lemma inst_prun_None i m1 m2 tmp old:
+ inst_prun ge i m1 tmp old = None ->
+ inst_prun ge i m2 tmp old = None.
+Proof.
+ intros H; generalize (inst_prun_fequiv i old m1 m2 tmp).
+ rewrite H; cbn; auto.
+Qed.
+
+Lemma inst_prun_Some i m1 m2 tmp old m1':
+ inst_prun ge i m1 tmp old = Some m1' ->
+ res_eq (Some (frame_assign m2 (inst_wframe i) m1')) (inst_prun ge i m2 tmp old).
+Proof.
+ intros H; generalize (inst_prun_fequiv i old m1 m2 tmp).
+ rewrite H; cbn.
+ intros (m2' & H1 & H2).
+ eexists; intuition eauto.
+ rewrite frame_assign_def.
+ lapply (inst_wframe_correct i m2' old m2 tmp); eauto.
+ destruct (notIn_dec x (inst_wframe i)); auto.
+ intros X; rewrite X; auto.
+Qed.
+
+Fixpoint bblock_wframe(p:bblock): list R.t :=
+ match p with
+ | nil => nil
+ | i::p' => (inst_wframe i)++(bblock_wframe p')
+ end.
+
+Local Hint Resolve Permutation_app_head Permutation_app_tail Permutation_app_comm: core.
+
+Lemma bblock_wframe_Permutation p p':
+ Permutation p p' -> Permutation (bblock_wframe p) (bblock_wframe p').
+Proof.
+ induction 1 as [|i p p'|i1 i2 p|p1 p2 p3]; cbn; auto.
+ - rewrite! app_assoc; auto.
+ - eapply Permutation_trans; eauto.
+Qed.
+
+(*
+Lemma bblock_wframe_correct p m' old: forall m,
+ prun_iw p m old = Some m' ->
+ forall x, notIn x (bblock_wframe p) -> m' x = m x.
+Proof.
+ induction p as [|i p']; cbn.
+ - intros m H; inversion_clear H; auto.
+ - intros m H x; rewrite notIn_app; intros (H1 & H2).
+ remember (inst_prun i m old old) as om.
+ destruct om as [m1|]; cbn.
+ + eapply eq_trans.
+ eapply IHp'; eauto.
+ eapply inst_wframe_correct; eauto.
+ + inversion H.
+Qed.
+
+Lemma prun_iw_fequiv p old: forall m1 m2,
+ frame_eq (fun x => In x (bblock_wframe p)) (prun_iw p m1 old) (prun_iw p m2 old).
+Proof.
+ induction p as [|i p']; cbn.
+ - intros m1 m2; eexists; intuition eauto.
+ - intros m1 m2; generalize (inst_prun_fequiv i old m1 m2 old).
+ remember (inst_prun i m1 old old) as om.
+ destruct om as [m1'|]; cbn.
+ + intros (m2' & H1 & H2). rewrite H1; cbn.
+ eapply frame_eq_list_split; eauto. clear IHp'.
+ intros m1'' m2'' x H3 H4. rewrite in_app_iff.
+ intros X X2. assert (X1: In x (inst_wframe i)). { destruct X; auto. rewrite <- notIn_iff in X2; tauto. }
+ clear X.
+ lapply (bblock_wframe_correct p' m1'' old m1'); eauto.
+ lapply (bblock_wframe_correct p' m2'' old m2'); eauto.
+ intros Xm2' Xm1'.
+ rewrite Xm1', Xm2'; auto.
+ + intro H; rewrite H; cbn; auto.
+Qed.
+
+Lemma prun_iw_equiv p m1 m2 old:
+ (forall x, notIn x (bblock_wframe p) -> m1 x = m2 x) ->
+ res_eq (prun_iw p m1 old) (prun_iw p m2 old).
+Proof.
+ intros; eapply frame_eq_res_eq.
+ eapply prun_iw_fequiv.
+ intros m1' m2' x H1 H2 H0.Require
+ lapply (bblock_wframe_correct p m1' old m1); eauto.
+ lapply (bblock_wframe_correct p m2' old m2); eauto.
+ intros X2 X1; rewrite X1, X2; auto.
+Qed.
+*)
+
+(** * Checking that parallel semantics is deterministic *)
+
+Fixpoint is_det (p: bblock): Prop :=
+ match p with
+ | nil => True
+ | i::p' =>
+ disjoint (inst_wframe i) (bblock_wframe p') (* no WRITE-AFTER-WRITE *)
+ /\ is_det p'
+ end.
+
+Lemma is_det_Permutation p p':
+ Permutation p p' -> is_det p -> is_det p'.
+Proof.
+ induction 1; cbn; auto.
+ - intros; intuition. eapply disjoint_incl_r. 2: eauto.
+ eapply Permutation_incl. eapply Permutation_sym.
+ eapply bblock_wframe_Permutation; auto.
+ - rewrite! disjoint_app_r in * |- *. intuition.
+ rewrite disjoint_sym; auto.
+Qed.
+
+Theorem is_det_correct p p':
+ Permutation p p' ->
+ is_det p ->
+ forall m old, res_eq (prun_iw ge p m old) (prun_iw ge p' m old).
+Proof.
+ induction 1 as [ | i p p' | i1 i2 p | p1 p2 p3 ]; cbn; eauto.
+ - intros [H0 H1] m old.
+ remember (inst_prun ge i m old old) as om0.
+ destruct om0 as [ m0 | ]; cbn; auto.
+ - rewrite disjoint_app_r.
+ intros ([Z1 Z2] & Z3 & Z4) m old.
+ remember (inst_prun ge i2 m old old) as om2.
+ destruct om2 as [ m2 | ]; cbn; auto.
+ + remember (inst_prun ge i1 m old old) as om1.
+ destruct om1 as [ m1 | ]; cbn; auto.
+ * lapply (inst_prun_Some i2 m m1 old old m2); cbn; auto.
+ lapply (inst_prun_Some i1 m m2 old old m1); cbn; auto.
+ intros (m1' & Hm1' & Xm1') (m2' & Hm2' & Xm2').
+ rewrite Hm1', Hm2'; cbn.
+ eapply prun_iw_equiv.
+ intros x; rewrite <- Xm1', <- Xm2'. clear Xm2' Xm1' Hm1' Hm2' m1' m2'.
+ rewrite frame_assign_def.
+ rewrite disjoint_sym in Z1; unfold disjoint in Z1.
+ destruct (notIn_dec x (inst_wframe i1)) as [ X1 | X1 ].
+ { rewrite frame_assign_def; destruct (notIn_dec x (inst_wframe i2)) as [ X2 | X2 ]; auto.
+ erewrite (inst_wframe_correct i2 m2 old m old); eauto.
+ erewrite (inst_wframe_correct i1 m1 old m old); eauto.
+ }
+ rewrite frame_assign_notIn; auto.
+ * erewrite inst_prun_None; eauto. cbn; auto.
+ + remember (inst_prun ge i1 m old old) as om1.
+ destruct om1 as [ m1 | ]; cbn; auto.
+ erewrite inst_prun_None; eauto.
+ - intros; eapply res_eq_trans.
+ eapply IHPermutation1; eauto.
+ eapply IHPermutation2; eauto.
+ eapply is_det_Permutation; eauto.
+Qed.
+
+(** * Standard Frames *)
+
+Fixpoint exp_frame (e: exp): list R.t :=
+ match e with
+ | PReg x => x::nil
+ | Op o le => list_exp_frame le
+ | Old e => exp_frame e
+ end
+with list_exp_frame (le: list_exp): list R.t :=
+ match le with
+ | Enil => nil
+ | Econs e le' => exp_frame e ++ list_exp_frame le'
+ | LOld le => list_exp_frame le
+ end.
+
+Lemma exp_frame_correct e old1 old2:
+ (forall x, In x (exp_frame e) -> old1 x = old2 x) ->
+ forall m1 m2, (forall x, In x (exp_frame e) -> m1 x = m2 x) ->
+ (exp_eval ge e m1 old1)=(exp_eval ge e m2 old2).
+Proof.
+ induction e using exp_mut with (P0:=fun l => (forall x, In x (list_exp_frame l) -> old1 x = old2 x) -> forall m1 m2, (forall x, In x (list_exp_frame l) -> m1 x = m2 x) ->
+ (list_exp_eval ge l m1 old1)=(list_exp_eval ge l m2 old2)); cbn; auto.
+ - intros H1 m1 m2 H2; rewrite H2; auto.
+ - intros H1 m1 m2 H2; erewrite IHe; eauto.
+ - intros H1 m1 m2 H2; erewrite IHe, IHe0; eauto;
+ intros; (eapply H1 || eapply H2); rewrite in_app_iff; auto.
+Qed.
+
+Fixpoint inst_frame (i: inst): list R.t :=
+ match i with
+ | nil => nil
+ | a::i' => (fst a)::(exp_frame (snd a) ++ inst_frame i')
+ end.
+
+Lemma inst_wframe_frame i x: In x (inst_wframe i) -> In x (inst_frame i).
+Proof.
+ induction i as [ | [y e] i']; cbn; intuition.
+Qed.
+
+
+Lemma inst_frame_correct i wframe old1 old2: forall m tmp1 tmp2,
+ (disjoint (inst_frame i) wframe) ->
+ (forall x, notIn x wframe -> old1 x = old2 x) ->
+ (forall x, notIn x wframe -> tmp1 x = tmp2 x) ->
+ inst_prun ge i m tmp1 old1 = inst_prun ge i m tmp2 old2.
+Proof.
+ induction i as [|[x e] i']; cbn; auto.
+ intros m tmp1 tmp2; rewrite disjoint_cons_l, disjoint_app_l.
+ intros (H1 & H2 & H3) H6 H7.
+ replace (exp_eval ge e tmp1 old1) with (exp_eval ge e tmp2 old2).
+ - destruct (exp_eval ge e tmp2 old2); auto.
+ eapply IHi'; eauto.
+ cbn; intros x0 H0; unfold assign. destruct (R.eq_dec x x0); cbn; auto.
+ - unfold disjoint in H2; apply exp_frame_correct.
+ intros;rewrite H6; auto.
+ intros;rewrite H7; auto.
+Qed.
+
+(** * Parallelizability *)
+
+Fixpoint pararec (p: bblock) (wframe: list R.t): Prop :=
+ match p with
+ | nil => True
+ | i::p' =>
+ disjoint (inst_frame i) wframe (* no USE-AFTER-WRITE *)
+ /\ pararec p' ((inst_wframe i) ++ wframe)
+ end.
+
+Lemma pararec_disjoint (p: bblock): forall wframe, pararec p wframe -> disjoint (bblock_wframe p) wframe.
+Proof.
+ induction p as [|i p']; cbn.
+ - unfold disjoint; cbn; intuition.
+ - intros wframe [H0 H1]; rewrite disjoint_app_l.
+ generalize (IHp' _ H1).
+ rewrite disjoint_app_r. intuition.
+ eapply disjoint_incl_l. 2: eapply H0.
+ unfold incl. eapply inst_wframe_frame; eauto.
+Qed.
+
+Lemma pararec_det p: forall wframe, pararec p wframe -> is_det p.
+Proof.
+ induction p as [|i p']; cbn; auto.
+ intros wframe [H0 H1]. generalize (pararec_disjoint _ _ H1). rewrite disjoint_app_r.
+ intuition.
+ - apply disjoint_sym; auto.
+ - eapply IHp'. eauto.
+Qed.
+
+Lemma pararec_correct p old: forall wframe m,
+ pararec p wframe ->
+ (forall x, notIn x wframe -> m x = old x) ->
+ run ge p m = prun_iw ge p m old.
+Proof.
+ elim p; clear p; cbn; auto.
+ intros i p' X wframe m [H H0] H1.
+ erewrite inst_run_prun, inst_frame_correct; eauto.
+ remember (inst_prun ge i m old old) as om0.
+ destruct om0 as [m0 | ]; try congruence.
+ eapply X; eauto.
+ intro x; rewrite notIn_app. intros [H3 H4].
+ rewrite <- H1; auto.
+ eapply inst_wframe_correct; eauto.
+Qed.
+
+Definition parallelizable (p: bblock) := pararec p nil.
+
+Theorem parallelizable_correct p m om':
+ parallelizable p -> (prun ge p m om' <-> res_eq om' (run ge p m)).
+Proof.
+ intros H. constructor 1.
+ - intros (p' & H0 & H1). eapply res_eq_trans; eauto.
+ erewrite pararec_correct; eauto.
+ eapply res_eq_sym.
+ eapply is_det_correct; eauto.
+ eapply pararec_det; eauto.
+ - intros; unfold prun.
+ eexists. constructor 1. 2: apply Permutation_refl.
+ erewrite pararec_correct in H0; eauto.
+Qed.
+
+End PARALLELI.
+
+End ParallelizablityChecking.
+
+
+(** * We assume a datatype [PseudoRegSet.t] refining [list R.t] *)
+
+(**
+This data-refinement is given by an abstract "invariant" match_frame below,
+preserved by the following operations.
+*)
+
+Module Type PseudoRegSet.
+
+Declare Module R: PseudoRegisters.
+
+Parameter t: Type.
+Parameter match_frame: t -> (list R.t) -> Prop.
+
+Parameter empty: t.
+Parameter empty_match_frame: match_frame empty nil.
+
+Parameter add: R.t -> t -> t.
+Parameter add_match_frame: forall s x l, match_frame s l -> match_frame (add x s) (x::l).
+
+Parameter union: t -> t -> t.
+Parameter union_match_frame: forall s1 s2 l1 l2, match_frame s1 l1 -> match_frame s2 l2 -> match_frame (union s1 s2) (l1++l2).
+
+Parameter is_disjoint: t -> t -> bool.
+Parameter is_disjoint_match_frame: forall s1 s2 l1 l2, match_frame s1 l1 -> match_frame s2 l2 -> (is_disjoint s1 s2)=true -> disjoint l1 l2.
+
+End PseudoRegSet.
+
+
+Lemma lazy_andb_bool_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true.
+Proof.
+ destruct b1, b2; intuition.
+Qed.
+
+
+
+
+Module ParallelChecks (L: SeqLanguage) (S:PseudoRegSet with Module R:=L.LP.R).
+
+Include ParallelizablityChecking L.
+
+Section PARALLEL2.
+Variable ge: genv.
+
+Local Hint Resolve S.empty_match_frame S.add_match_frame S.union_match_frame S.is_disjoint_match_frame: core.
+
+(** Now, refinement of each operation toward parallelizable *)
+
+Fixpoint inst_wsframe(i:inst): S.t :=
+ match i with
+ | nil => S.empty
+ | a::i' => S.add (fst a) (inst_wsframe i')
+ end.
+
+Lemma inst_wsframe_correct i: S.match_frame (inst_wsframe i) (inst_wframe i).
+Proof.
+ induction i; cbn; auto.
+Qed.
+
+Fixpoint exp_sframe (e: exp): S.t :=
+ match e with
+ | PReg x => S.add x S.empty
+ | Op o le => list_exp_sframe le
+ | Old e => exp_sframe e
+ end
+with list_exp_sframe (le: list_exp): S.t :=
+ match le with
+ | Enil => S.empty
+ | Econs e le' => S.union (exp_sframe e) (list_exp_sframe le')
+ | LOld le => list_exp_sframe le
+ end.
+
+Lemma exp_sframe_correct e: S.match_frame (exp_sframe e) (exp_frame e).
+Proof.
+ induction e using exp_mut with (P0:=fun l => S.match_frame (list_exp_sframe l) (list_exp_frame l)); cbn; auto.
+Qed.
+
+Fixpoint inst_sframe (i: inst): S.t :=
+ match i with
+ | nil => S.empty
+ | a::i' => S.add (fst a) (S.union (exp_sframe (snd a)) (inst_sframe i'))
+ end.
+
+Local Hint Resolve exp_sframe_correct: core.
+
+Lemma inst_sframe_correct i: S.match_frame (inst_sframe i) (inst_frame i).
+Proof.
+ induction i as [|[y e] i']; cbn; auto.
+Qed.
+
+Local Hint Resolve inst_wsframe_correct inst_sframe_correct: core.
+
+Fixpoint is_pararec (p: bblock) (wsframe: S.t): bool :=
+ match p with
+ | nil => true
+ | i::p' =>
+ S.is_disjoint (inst_sframe i) wsframe (* no USE-AFTER-WRITE *)
+ &&& is_pararec p' (S.union (inst_wsframe i) wsframe)
+ end.
+
+Lemma is_pararec_correct (p: bblock): forall s l, S.match_frame s l -> (is_pararec p s)=true -> (pararec p l).
+Proof.
+ induction p; cbn; auto.
+ intros s l H1 H2; rewrite lazy_andb_bool_true in H2. destruct H2 as [H2 H3].
+ constructor 1; eauto.
+Qed.
+
+Definition is_parallelizable (p: bblock) := is_pararec p S.empty.
+
+Lemma is_para_correct_aux p: is_parallelizable p = true -> parallelizable p.
+Proof.
+ unfold is_parallelizable, parallelizable; intros; eapply is_pararec_correct; eauto.
+Qed.
+
+Theorem is_parallelizable_correct p:
+ is_parallelizable p = true -> forall m om', (prun ge p m om' <-> res_eq om' (run ge p m)).
+Proof.
+ intros; apply parallelizable_correct.
+ apply is_para_correct_aux. auto.
+Qed.
+
+End PARALLEL2.
+End ParallelChecks.
+
+
+
+(** * Implementing the datatype [PosPseudoRegSet.t] refining [list R.t] *)
+
+(* This data-refinement is given by an abstract "invariant" match_frame below,
+preserved by the following operations.
+*)
+
+Require Import PArith.
+Require Import MSets.MSetPositive.
+
+Module PosPseudoRegSet <: PseudoRegSet with Module R:=Pos.
+
+Module R:=Pos.
+
+
+Definition t:=PositiveSet.t.
+
+Definition match_frame (s:t) (l:list R.t): Prop
+ := forall x, PositiveSet.In x s <-> In x l.
+
+Definition empty:=PositiveSet.empty.
+
+Lemma empty_match_frame: match_frame empty nil.
+Proof.
+ unfold match_frame, empty, PositiveSet.In; cbn; intuition.
+Qed.
+
+Definition add: R.t -> t -> t := PositiveSet.add.
+
+Lemma add_match_frame: forall s x l, match_frame s l -> match_frame (add x s) (x::l).
+Proof.
+ unfold match_frame, add; cbn.
+ intros s x l H y. rewrite PositiveSet.add_spec, H.
+ intuition.
+Qed.
+
+Definition union: t -> t -> t := PositiveSet.union.
+Lemma union_match_frame: forall s1 s2 l1 l2, match_frame s1 l1 -> match_frame s2 l2 -> match_frame (union s1 s2) (l1++l2).
+Proof.
+ unfold match_frame, union.
+ intros s1 s2 l1 l2 H1 H2 x. rewrite PositiveSet.union_spec, H1, H2.
+ intuition.
+Qed.
+
+Fixpoint is_disjoint (s s': PositiveSet.t) : bool :=
+ match s with
+ | PositiveSet.Leaf => true
+ | PositiveSet.Node l o r =>
+ match s' with
+ | PositiveSet.Leaf => true
+ | PositiveSet.Node l' o' r' =>
+ if (o &&& o') then false else (is_disjoint l l' &&& is_disjoint r r')
+ end
+ end.
+
+Lemma is_disjoint_spec_true s: forall s', is_disjoint s s' = true -> forall x, PositiveSet.In x s -> PositiveSet.In x s' -> False.
+Proof.
+ unfold PositiveSet.In; induction s as [ |l IHl o r IHr]; cbn; try discriminate.
+ destruct s' as [|l' o' r']; cbn; try discriminate.
+ intros X.
+ assert (H: ~(o = true /\ o'=true) /\ is_disjoint l l' = true /\ is_disjoint r r'=true).
+ { destruct o, o', (is_disjoint l l'), (is_disjoint r r'); cbn in X; intuition. }
+ clear X; destruct H as (H & H1 & H2).
+ destruct x as [i|i|]; cbn; eauto.
+Qed.
+
+Lemma is_disjoint_match_frame: forall s1 s2 l1 l2, match_frame s1 l1 -> match_frame s2 l2 -> (is_disjoint s1 s2)=true -> disjoint l1 l2.
+Proof.
+ unfold match_frame, disjoint.
+ intros s1 s2 l1 l2 H1 H2 H3 x.
+ rewrite <- notIn_iff, <- H1, <- H2.
+ intros H4 H5; eapply is_disjoint_spec_true; eauto.
+Qed.
+
+End PosPseudoRegSet.
diff --git a/scheduling/abstractbb/README.md b/scheduling/abstractbb/README.md
new file mode 100644
index 00000000..69e5defc
--- /dev/null
+++ b/scheduling/abstractbb/README.md
@@ -0,0 +1,12 @@
+# Coq sources of AbstractBasicBlocks
+
+- [AbstractBasicBlocksDef](AbstractBasicBlocksDef.v): syntax and sequential semantics of abstract basic blocks (on which we define our analyzes).
+This syntax and semantics is parametrized in order to adapt the language for different concrete basic block languages.
+
+- [Parallelizability](Parallelizability.v): define the parallel semantics and the 'is_parallelizable' function which tests whether the sequential run of a given abstract basic block is the same than a parallel run.
+
+- [DepTreeTheory](DepTreeTheory.v): defines a theory of dependency trees, such that two basic blocks with the same dependency tree have the same sequential semantics. In practice, permuting the instructions inside a basic block while perserving the dependencies of assignments should not change the dependency tree. The idea is to verify list schedulings, following ideas of [Formal verification of translation validators proposed by Tristan and Leroy](https://hal.inria.fr/inria-00289540/).
+
+- [ImpDep](ImpDep.v): adds a hash-consing mechanism to trees of [DepTreeTheory](DepTreeTheory.v), and thus provides an efficient "equality" test (a true answer ensures that the two basic blocks in input have the same sequential semantics) in order to check the correctness of list schedulings.
+
+- [DepExample](DepExample.v) defines a toy language (syntax and semantics); [DepExampleEqTest](DepExampleEqTest.v) defines a compiler of the toy language into abstract basic blocks and derives an equality test for the toy language; [DepExampleParallelTest](DepExampleParallelTest.v) derives a parallelizability test from the previous compiler; [DepExampleDemo](DepExampleDemo.v) is a test-suite for both tetsts.
diff --git a/scheduling/abstractbb/SeqSimuTheory.v b/scheduling/abstractbb/SeqSimuTheory.v
new file mode 100644
index 00000000..df6b9963
--- /dev/null
+++ b/scheduling/abstractbb/SeqSimuTheory.v
@@ -0,0 +1,397 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** A theory for checking/proving simulation by symbolic execution.
+
+*)
+
+
+Require Coq.Logic.FunctionalExtensionality. (* not really necessary -- see lemma at the end *)
+Require Setoid. (* in order to rewrite <-> *)
+Require Export AbstractBasicBlocksDef.
+Require Import List.
+Require Import ImpPrelude.
+Import HConsingDefs.
+
+
+Module SimuTheory (L: SeqLanguage).
+
+Export L.
+Export LP.
+
+Inductive term :=
+ | Input (x:R.t)
+ | App (o: op) (l: list_term)
+with list_term :=
+ | LTnil
+ | LTcons (t:term) (l:list_term)
+ .
+
+Fixpoint term_eval (ge: genv) (t: term) (m: mem): option value :=
+ match t with
+ | Input x => Some (m x)
+ | App o l =>
+ match list_term_eval ge l m with
+ | Some v => op_eval ge o v
+ | _ => None
+ end
+ end
+with list_term_eval ge (l: list_term) (m: mem) {struct l}: option (list value) :=
+ match l with
+ | LTnil => Some nil
+ | LTcons t l' =>
+ match term_eval ge t m, list_term_eval ge l' m with
+ | Some v, Some lv => Some (v::lv)
+ | _, _ => None
+ end
+ end.
+
+(** The (abstract) symbolic memory state *)
+Record smem :=
+{
+ pre: genv -> mem -> Prop; (**r pre-condition expressing that the computation has not yet abort on a None. *)
+ post:> R.t -> term (**r the output term computed on each pseudo-register *)
+}.
+
+(** Initial symbolic memory state *)
+Definition smem_empty := {| pre:=fun _ _ => True; post:=(fun x => Input x) |}.
+
+Fixpoint exp_term (e: exp) (d old: smem) : term :=
+ match e with
+ | PReg x => d x
+ | Op o le => App o (list_exp_term le d old)
+ | Old e => exp_term e old old
+ end
+with list_exp_term (le: list_exp) (d old: smem) : list_term :=
+ match le with
+ | Enil => LTnil
+ | Econs e le' => LTcons (exp_term e d old) (list_exp_term le' d old)
+ | LOld le => list_exp_term le old old
+ end.
+
+
+(** assignment of the symbolic memory state *)
+Definition smem_set (d:smem) x (t:term) :=
+ {| pre:=(fun ge m => (term_eval ge (d x) m) <> None /\ (d.(pre) ge m));
+ post:=fun y => if R.eq_dec x y then t else d y |}.
+
+(** Simulation theory: the theorem [bblock_smem_simu] ensures that the simulation between two abstract basic blocks is implied by the simulation between their symbolic execution. *)
+Section SIMU_THEORY.
+
+Variable ge: genv.
+
+Lemma set_spec_eq d x t m:
+ term_eval ge (smem_set d x t x) m = term_eval ge t m.
+Proof.
+ unfold smem_set; cbn; case (R.eq_dec x x); try congruence.
+Qed.
+
+Lemma set_spec_diff d x y t m:
+ x <> y -> term_eval ge (smem_set d x t y) m = term_eval ge (d y) m.
+Proof.
+ unfold smem_set; cbn; case (R.eq_dec x y); try congruence.
+Qed.
+
+Fixpoint inst_smem (i: inst) (d old: smem): smem :=
+ match i with
+ | nil => d
+ | (x, e)::i' =>
+ let t:=exp_term e d old in
+ inst_smem i' (smem_set d x t) old
+ end.
+
+Fixpoint bblock_smem_rec (p: bblock) (d: smem): smem :=
+ match p with
+ | nil => d
+ | i::p' =>
+ let d':=inst_smem i d d in
+ bblock_smem_rec p' d'
+ end.
+
+Definition bblock_smem: bblock -> smem
+ := fun p => bblock_smem_rec p smem_empty.
+
+Lemma inst_smem_pre_monotonic i old: forall d m,
+ (pre (inst_smem i d old) ge m) -> (pre d ge m).
+Proof.
+ induction i as [|[y e] i IHi]; cbn; auto.
+ intros d a H; generalize (IHi _ _ H); clear H IHi.
+ unfold smem_set; cbn; intuition.
+Qed.
+
+Lemma bblock_smem_pre_monotonic p: forall d m,
+ (pre (bblock_smem_rec p d) ge m) -> (pre d ge m).
+Proof.
+ induction p as [|i p' IHp']; cbn; eauto.
+ intros d a H; eapply inst_smem_pre_monotonic; eauto.
+Qed.
+
+Local Hint Resolve inst_smem_pre_monotonic bblock_smem_pre_monotonic: core.
+
+Lemma term_eval_exp e (od:smem) m0 old:
+ (forall x, term_eval ge (od x) m0 = Some (old x)) ->
+ forall (d:smem) m1,
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ term_eval ge (exp_term e d od) m0 = exp_eval ge e m1 old.
+Proof.
+ intro H.
+ induction e using exp_mut with
+ (P0:=fun l => forall (d:smem) m1, (forall x, term_eval ge (d x) m0 = Some (m1 x)) -> list_term_eval ge (list_exp_term l d od) m0 = list_exp_eval ge l m1 old);
+ cbn; auto.
+ - intros; erewrite IHe; eauto.
+ - intros. erewrite IHe, IHe0; eauto.
+Qed.
+
+Lemma inst_smem_abort i m0 x old: forall (d:smem),
+ pre (inst_smem i d old) ge m0 ->
+ term_eval ge (d x) m0 = None ->
+ term_eval ge (inst_smem i d old x) m0 = None.
+Proof.
+ induction i as [|[y e] i IHi]; cbn; auto.
+ intros d VALID H; erewrite IHi; eauto. clear IHi.
+ unfold smem_set; cbn; destruct (R.eq_dec y x); auto.
+ subst;
+ generalize (inst_smem_pre_monotonic _ _ _ _ VALID); clear VALID.
+ unfold smem_set; cbn. intuition congruence.
+Qed.
+
+Lemma block_smem_rec_abort p m0 x: forall d,
+ pre (bblock_smem_rec p d) ge m0 ->
+ term_eval ge (d x) m0 = None ->
+ term_eval ge (bblock_smem_rec p d x) m0 = None.
+Proof.
+ induction p; cbn; auto.
+ intros d VALID H; erewrite IHp; eauto. clear IHp.
+ eapply inst_smem_abort; eauto.
+Qed.
+
+Lemma inst_smem_Some_correct1 i m0 old (od:smem):
+ (forall x, term_eval ge (od x) m0 = Some (old x)) ->
+ forall (m1 m2: mem) (d: smem),
+ inst_run ge i m1 old = Some m2 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ forall x, term_eval ge (inst_smem i d od x) m0 = Some (m2 x).
+Proof.
+ intro X; induction i as [|[x e] i IHi]; cbn; intros m1 m2 d H.
+ - inversion_clear H; eauto.
+ - intros H0 x0.
+ destruct (exp_eval ge e m1 old) eqn:Heqov; try congruence.
+ refine (IHi _ _ _ _ _ _); eauto.
+ clear x0; intros x0.
+ unfold assign, smem_set; cbn. destruct (R.eq_dec x x0); auto.
+ subst; erewrite term_eval_exp; eauto.
+Qed.
+
+Lemma bblocks_smem_rec_Some_correct1 p m0: forall (m1 m2: mem) (d: smem),
+ run ge p m1 = Some m2 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ forall x, term_eval ge (bblock_smem_rec p d x) m0 = Some (m2 x).
+Proof.
+ Local Hint Resolve inst_smem_Some_correct1: core.
+ induction p as [ | i p]; cbn; intros m1 m2 d H.
+ - inversion_clear H; eauto.
+ - intros H0 x0.
+ destruct (inst_run ge i m1 m1) eqn: Heqov.
+ + refine (IHp _ _ _ _ _ _); eauto.
+ + inversion H.
+Qed.
+
+Lemma bblock_smem_Some_correct1 p m0 m1:
+ run ge p m0 = Some m1
+ -> forall x, term_eval ge (bblock_smem p x) m0 = Some (m1 x).
+Proof.
+ intros; eapply bblocks_smem_rec_Some_correct1; eauto.
+Qed.
+
+Lemma inst_smem_None_correct i m0 old (od: smem):
+ (forall x, term_eval ge (od x) m0 = Some (old x)) ->
+ forall m1 d, pre (inst_smem i d od) ge m0 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ inst_run ge i m1 old = None -> exists x, term_eval ge (inst_smem i d od x) m0 = None.
+Proof.
+ intro X; induction i as [|[x e] i IHi]; cbn; intros m1 d.
+ - discriminate.
+ - intros VALID H0.
+ destruct (exp_eval ge e m1 old) eqn: Heqov.
+ + refine (IHi _ _ _ _); eauto.
+ intros x0; unfold assign, smem_set; cbn. destruct (R.eq_dec x x0); auto.
+ subst; erewrite term_eval_exp; eauto.
+ + intuition.
+ constructor 1 with (x:=x); cbn.
+ apply inst_smem_abort; auto.
+ rewrite set_spec_eq.
+ erewrite term_eval_exp; eauto.
+Qed.
+
+Lemma inst_smem_Some_correct2 i m0 old (od: smem):
+ (forall x, term_eval ge (od x) m0 = Some (old x)) ->
+ forall (m1 m2: mem) d,
+ pre (inst_smem i d od) ge m0 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ (forall x, term_eval ge (inst_smem i d od x) m0 = Some (m2 x)) ->
+ res_eq (Some m2) (inst_run ge i m1 old).
+Proof.
+ intro X.
+ induction i as [|[x e] i IHi]; cbn; intros m1 m2 d VALID H0.
+ - intros H; eapply ex_intro; intuition eauto.
+ generalize (H0 x); rewrite H.
+ congruence.
+ - intros H.
+ destruct (exp_eval ge e m1 old) eqn: Heqov.
+ + refine (IHi _ _ _ _ _ _); eauto.
+ intros x0; unfold assign, smem_set; cbn; destruct (R.eq_dec x x0); auto.
+ subst; erewrite term_eval_exp; eauto.
+ + generalize (H x).
+ rewrite inst_smem_abort; discriminate || auto.
+ rewrite set_spec_eq.
+ erewrite term_eval_exp; eauto.
+Qed.
+
+Lemma bblocks_smem_rec_Some_correct2 p m0: forall (m1 m2: mem) d,
+ pre (bblock_smem_rec p d) ge m0 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ (forall x, term_eval ge (bblock_smem_rec p d x) m0 = Some (m2 x)) ->
+ res_eq (Some m2) (run ge p m1).
+Proof.
+ induction p as [|i p]; cbn; intros m1 m2 d VALID H0.
+ - intros H; eapply ex_intro; intuition eauto.
+ generalize (H0 x); rewrite H.
+ congruence.
+ - intros H.
+ destruct (inst_run ge i m1 m1) eqn: Heqom.
+ + refine (IHp _ _ _ _ _ _); eauto.
+ + assert (X: exists x, term_eval ge (inst_smem i d d x) m0 = None).
+ { eapply inst_smem_None_correct; eauto. }
+ destruct X as [x H1].
+ generalize (H x).
+ erewrite block_smem_rec_abort; eauto.
+ congruence.
+Qed.
+
+Lemma bblock_smem_Some_correct2 p m0 m1:
+ pre (bblock_smem p) ge m0 ->
+ (forall x, term_eval ge (bblock_smem p x) m0 = Some (m1 x))
+ -> res_eq (Some m1) (run ge p m0).
+Proof.
+ intros; eapply bblocks_smem_rec_Some_correct2; eauto.
+Qed.
+
+Lemma inst_valid i m0 old (od:smem):
+ (forall x, term_eval ge (od x) m0 = Some (old x)) ->
+ forall (m1 m2: mem) (d: smem),
+ pre d ge m0 ->
+ inst_run ge i m1 old = Some m2 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ pre (inst_smem i d od) ge m0.
+Proof.
+ induction i as [|[x e] i IHi]; cbn; auto.
+ intros Hold m1 m2 d VALID0 H Hm1.
+ destruct (exp_eval ge e m1 old) eqn: Heq; cbn; try congruence.
+ eapply IHi; eauto.
+ + unfold smem_set in * |- *; cbn.
+ rewrite Hm1; intuition congruence.
+ + intros x0. unfold assign, smem_set; cbn; destruct (R.eq_dec x x0); auto.
+ subst; erewrite term_eval_exp; eauto.
+Qed.
+
+
+Lemma block_smem_rec_valid p m0: forall (m1 m2: mem) (d:smem),
+ pre d ge m0 ->
+ run ge p m1 = Some m2 ->
+ (forall x, term_eval ge (d x) m0 = Some (m1 x)) ->
+ pre (bblock_smem_rec p d) ge m0.
+Proof.
+ Local Hint Resolve inst_valid: core.
+ induction p as [ | i p]; cbn; intros m1 d H; auto.
+ intros H0 H1.
+ destruct (inst_run ge i m1 m1) eqn: Heqov; eauto.
+ congruence.
+Qed.
+
+Lemma bblock_smem_valid p m0 m1:
+ run ge p m0 = Some m1 ->
+ pre (bblock_smem p) ge m0.
+Proof.
+ intros; eapply block_smem_rec_valid; eauto.
+ unfold smem_empty; cbn. auto.
+Qed.
+
+Definition smem_valid ge d m := pre d ge m /\ forall x, term_eval ge (d x) m <> None.
+
+Definition smem_simu (d1 d2: smem): Prop :=
+ (forall m, smem_valid ge d1 m -> smem_valid ge d2 m)
+ /\ (forall m0 x, smem_valid ge d1 m0 ->
+ term_eval ge (d1 x) m0 = term_eval ge (d2 x) m0).
+
+
+Theorem bblock_smem_simu p1 p2:
+ smem_simu (bblock_smem p1) (bblock_smem p2) ->
+ bblock_simu ge p1 p2.
+Proof.
+ Local Hint Resolve bblock_smem_valid bblock_smem_Some_correct1: core.
+ intros (INCL & EQUIV) m DONTFAIL; unfold smem_valid in * |-.
+ destruct (run ge p1 m) as [m1|] eqn: RUN1; cbn; try congruence.
+ assert (X: forall x, term_eval ge (bblock_smem p1 x) m = Some (m1 x)); eauto.
+ eapply bblock_smem_Some_correct2; eauto.
+ + destruct (INCL m); intuition eauto.
+ congruence.
+ + intro x; erewrite <- EQUIV; intuition eauto.
+ congruence.
+Qed.
+
+Lemma smem_valid_set_decompose_1 d t x m:
+ smem_valid ge (smem_set d x t) m -> smem_valid ge d m.
+Proof.
+ unfold smem_valid; intros ((PRE1 & PRE2) & VALID); split.
+ + intuition.
+ + intros x0 H. case (R.eq_dec x x0).
+ * intuition congruence.
+ * intros DIFF; eapply VALID. erewrite set_spec_diff; eauto.
+Qed.
+
+Lemma smem_valid_set_decompose_2 d t x m:
+ smem_valid ge (smem_set d x t) m -> term_eval ge t m <> None.
+Proof.
+ unfold smem_valid; intros ((PRE1 & PRE2) & VALID) H.
+ generalize (VALID x); rewrite set_spec_eq.
+ tauto.
+Qed.
+
+Lemma smem_valid_set_proof d x t m:
+ smem_valid ge d m -> term_eval ge t m <> None -> smem_valid ge (smem_set d x t) m.
+Proof.
+ unfold smem_valid; intros (PRE & VALID) PREt. split.
+ + split; auto.
+ + intros x0; unfold smem_set; cbn; case (R.eq_dec x x0); intros; subst; auto.
+Qed.
+
+
+End SIMU_THEORY.
+
+(** REMARK: this theorem reformulates the lemma above in a more abstract way (but relies on functional_extensionality axiom).
+*)
+Definition smem_correct ge (d: smem) (m: mem) (om: option mem): Prop:=
+ forall m', om=Some m' <-> (d.(pre) ge m /\ forall x, term_eval ge (d x) m = Some (m' x)).
+
+Lemma bblock_smem_correct ge p m: smem_correct ge (bblock_smem p) m (run ge p m).
+Proof.
+ unfold smem_correct; cbn; intros m'; split.
+ + intros; split.
+ * eapply bblock_smem_valid; eauto.
+ * eapply bblock_smem_Some_correct1; eauto.
+ + intros (H1 & H2).
+ destruct (bblock_smem_Some_correct2 ge p m m') as (m2 & X & Y); eauto.
+ rewrite X. f_equal.
+ apply FunctionalExtensionality.functional_extensionality; auto.
+Qed.
+
+End SimuTheory.
diff --git a/scheduling/postpass_lib/ForwardSimulationBlock.v b/scheduling/postpass_lib/ForwardSimulationBlock.v
new file mode 100644
index 00000000..cc6ecdd8
--- /dev/null
+++ b/scheduling/postpass_lib/ForwardSimulationBlock.v
@@ -0,0 +1,387 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(***
+
+Auxiliary lemmas on starN and forward_simulation
+in order to prove the forward simulation of Mach -> Machblock.
+
+***)
+
+Require Import Relations.
+Require Import Wellfounded.
+Require Import Coqlib.
+Require Import Events.
+Require Import Globalenvs.
+Require Import Smallstep.
+Require Import Lia.
+
+Local Open Scope nat_scope.
+
+
+(** Auxiliary lemma on starN *)
+Section starN_lemma.
+
+Variable L: semantics.
+
+Local Hint Resolve starN_refl starN_step Eapp_assoc: core.
+
+Lemma starN_split n s t s':
+ starN (step L) (globalenv L) n s t s' ->
+ forall m k, n=m+k ->
+ exists (t1 t2:trace) s0, starN (step L) (globalenv L) m s t1 s0 /\ starN (step L) (globalenv L) k s0 t2 s' /\ t=t1**t2.
+Proof.
+ induction 1; cbn.
+ + intros m k H; assert (X: m=0); try lia.
+ assert (X0: k=0); try lia.
+ subst; repeat (eapply ex_intro); intuition eauto.
+ + intros m; destruct m as [| m']; cbn.
+ - intros k H2; subst; repeat (eapply ex_intro); intuition eauto.
+ - intros k H2. inversion H2.
+ exploit (IHstarN m' k); eauto. intro.
+ destruct H3 as (t5 & t6 & s0 & H5 & H6 & H7).
+ repeat (eapply ex_intro).
+ instantiate (1 := t6); instantiate (1 := t1 ** t5); instantiate (1 := s0).
+ intuition eauto. subst. auto.
+Qed.
+
+Lemma starN_tailstep n s t1 s':
+ starN (step L) (globalenv L) n s t1 s' ->
+ forall (t t2:trace) s'',
+ Step L s' t2 s'' -> t = t1 ** t2 -> starN (step L) (globalenv L) (S n) s t s''.
+Proof.
+ induction 1; cbn.
+ + intros t t1 s0; autorewrite with trace_rewrite.
+ intros; subst; eapply starN_step; eauto.
+ autorewrite with trace_rewrite; auto.
+ + intros. eapply starN_step; eauto.
+ intros; subst; autorewrite with trace_rewrite; auto.
+Qed.
+
+End starN_lemma.
+
+
+
+(** General scheme from a "match_states" relation *)
+
+Section ForwardSimuBlock_REL.
+
+Variable L1 L2: semantics.
+
+
+(** Hypothèses de la preuve *)
+
+Variable dist_end_block: state L1 -> nat.
+
+Hypothesis simu_mid_block:
+ forall s1 t s1', Step L1 s1 t s1' -> (dist_end_block s1)<>0 -> t = E0 /\ dist_end_block s1=S (dist_end_block s1').
+
+Hypothesis public_preserved:
+ forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id.
+
+Variable match_states: state L1 -> state L2 -> Prop.
+
+Hypothesis match_initial_states:
+ forall s1, initial_state L1 s1 -> exists s2, match_states s1 s2 /\ initial_state L2 s2.
+
+Hypothesis match_final_states:
+ forall s1 s2 r, final_state L1 s1 r -> match_states s1 s2 -> final_state L2 s2 r.
+
+Hypothesis final_states_end_block:
+ forall s1 t s1' r, Step L1 s1 t s1' -> final_state L1 s1' r -> dist_end_block s1 = 0.
+
+Hypothesis simu_end_block:
+ forall s1 t s1' s2, starN (step L1) (globalenv L1) (S (dist_end_block s1)) s1 t s1' -> match_states s1 s2 -> exists s2', Step L2 s2 t s2' /\ match_states s1' s2'.
+
+
+(** Introduction d'une sémantique par bloc sur L1 appelée "memoL1" *)
+
+Local Hint Resolve starN_refl starN_step: core.
+
+Definition follows_in_block (head current: state L1): Prop :=
+ dist_end_block head >= dist_end_block current
+ /\ starN (step L1) (globalenv L1) (minus (dist_end_block head) (dist_end_block current)) head E0 current.
+
+Lemma follows_in_block_step (head previous next: state L1):
+ forall t, follows_in_block head previous -> Step L1 previous t next -> (dist_end_block previous)<>0 -> follows_in_block head next.
+Proof.
+ intros t [H1 H2] H3 H4.
+ destruct (simu_mid_block _ _ _ H3 H4) as [H5 H6]; subst.
+ constructor 1.
+ + lia.
+ + replace (dist_end_block head - dist_end_block next) with (S (dist_end_block head - dist_end_block previous)).
+ - eapply starN_tailstep; eauto.
+ - lia.
+Qed.
+
+Lemma follows_in_block_init (head current: state L1):
+ forall t, Step L1 head t current -> (dist_end_block head)<>0 -> follows_in_block head current.
+Proof.
+ intros t H3 H4.
+ destruct (simu_mid_block _ _ _ H3 H4) as [H5 H6]; subst.
+ constructor 1.
+ + lia.
+ + replace (dist_end_block head - dist_end_block current) with 1.
+ - eapply starN_tailstep; eauto.
+ - lia.
+Qed.
+
+
+Record memostate := {
+ real: state L1;
+ memorized: option (state L1);
+ memo_star: forall head, memorized = Some head -> follows_in_block head real;
+ memo_final: forall r, final_state L1 real r -> memorized = None
+}.
+
+Definition head (s: memostate): state L1 :=
+ match memorized s with
+ | None => real s
+ | Some s' => s'
+ end.
+
+Lemma head_followed (s: memostate): follows_in_block (head s) (real s).
+Proof.
+ destruct s as [rs ms Hs]. cbn.
+ destruct ms as [ms|]; unfold head; cbn; auto.
+ constructor 1.
+ lia.
+ replace (dist_end_block rs - dist_end_block rs)%nat with O.
+ + apply starN_refl; auto.
+ + lia.
+Qed.
+
+Inductive is_well_memorized (s s': memostate): Prop :=
+ | StartBloc:
+ dist_end_block (real s) <> O ->
+ memorized s = None ->
+ memorized s' = Some (real s) ->
+ is_well_memorized s s'
+ | MidBloc:
+ dist_end_block (real s) <> O ->
+ memorized s <> None ->
+ memorized s' = memorized s ->
+ is_well_memorized s s'
+ | ExitBloc:
+ dist_end_block (real s) = O ->
+ memorized s' = None ->
+ is_well_memorized s s'.
+
+Local Hint Resolve StartBloc MidBloc ExitBloc: core.
+
+Definition memoL1 := {|
+ state := memostate;
+ genvtype := genvtype L1;
+ step := fun ge s t s' =>
+ step L1 ge (real s) t (real s')
+ /\ is_well_memorized s s' ;
+ initial_state := fun s => initial_state L1 (real s) /\ memorized s = None;
+ final_state := fun s r => final_state L1 (real s) r;
+ globalenv:= globalenv L1;
+ symbolenv:= symbolenv L1
+|}.
+
+
+(** Preuve des 2 forward simulations: L1 -> memoL1 et memoL1 -> L2 *)
+
+Lemma discr_dist_end s:
+ {dist_end_block s = O} + {dist_end_block s <> O}.
+Proof.
+ destruct (dist_end_block s); cbn; intuition.
+Qed.
+
+Lemma memo_simulation_step:
+ forall s1 t s1', Step L1 s1 t s1' ->
+ forall s2, s1 = (real s2) -> exists s2', Step memoL1 s2 t s2' /\ s1' = (real s2').
+Proof.
+ intros s1 t s1' H1 [rs2 ms2 Hmoi] H2. cbn in H2; subst.
+ destruct (discr_dist_end rs2) as [H3 | H3].
+ + refine (ex_intro _ {|real:=s1'; memorized:=None |} _); cbn.
+ intuition.
+ + destruct ms2 as [s|].
+ - refine (ex_intro _ {|real:=s1'; memorized:=Some s |} _); cbn.
+ intuition.
+ - refine (ex_intro _ {|real:=s1'; memorized:=Some rs2 |} _); cbn.
+ intuition.
+ Unshelve.
+ * intros; discriminate.
+ * intros; auto.
+ * intros head X; injection X; clear X; intros; subst.
+ eapply follows_in_block_step; eauto.
+ * intros r X; erewrite final_states_end_block in H3; intuition eauto.
+ * intros head X; injection X; clear X; intros; subst.
+ eapply follows_in_block_init; eauto.
+ * intros r X; erewrite final_states_end_block in H3; intuition eauto.
+Qed.
+
+Lemma forward_memo_simulation_1: forward_simulation L1 memoL1.
+Proof.
+ apply forward_simulation_step with (match_states:=fun s1 s2 => s1 = (real s2)); auto.
+ + intros s1 H; eapply ex_intro with (x:={|real:=s1; memorized:=None |}); cbn.
+ intuition.
+ + intros; subst; auto.
+ + intros; exploit memo_simulation_step; eauto.
+ Unshelve.
+ * intros; discriminate.
+ * auto.
+Qed.
+
+Lemma forward_memo_simulation_2: forward_simulation memoL1 L2.
+Proof.
+ unfold memoL1; cbn.
+ apply forward_simulation_opt with (measure:=fun s => dist_end_block (real s)) (match_states:=fun s1 s2 => match_states (head s1) s2); cbn; auto.
+ + intros s1 [H0 H1]; destruct (match_initial_states (real s1) H0).
+ unfold head; rewrite H1.
+ intuition eauto.
+ + intros s1 s2 r X H0; unfold head in X.
+ erewrite memo_final in X; eauto.
+ + intros s1 t s1' [H1 H2] s2 H; subst.
+ destruct H2 as [ H0 H2 H3 | H0 H2 H3 | H0 H2].
+ - (* StartBloc *)
+ constructor 2. destruct (simu_mid_block (real s1) t (real s1')) as [H5 H4]; auto.
+ unfold head in * |- *. rewrite H2 in H. rewrite H3. rewrite H4. intuition.
+ - (* MidBloc *)
+ constructor 2. destruct (simu_mid_block (real s1) t (real s1')) as [H5 H4]; auto.
+ unfold head in * |- *. rewrite H3. rewrite H4. intuition.
+ destruct (memorized s1); cbn; auto. tauto.
+ - (* EndBloc *)
+ constructor 1.
+ destruct (simu_end_block (head s1) t (real s1') s2) as (s2' & H3 & H4); auto.
+ * destruct (head_followed s1) as [H4 H3].
+ replace (dist_end_block (head s1) - dist_end_block (real s1)) with (dist_end_block (head s1)) in H3; try lia.
+ eapply starN_tailstep; eauto.
+ * unfold head; rewrite H2; cbn. intuition eauto.
+Qed.
+
+Lemma forward_simulation_block_rel: forward_simulation L1 L2.
+Proof.
+ eapply compose_forward_simulations.
+ eapply forward_memo_simulation_1.
+ apply forward_memo_simulation_2.
+Qed.
+
+
+End ForwardSimuBlock_REL.
+
+
+
+(* An instance of the previous scheme, when there is a translation from L1 states to L2 states
+
+Here, we do not require that the sequence of S2 states does exactly match the sequence of L1 states by trans_state.
+This is because the exact matching is broken in Machblock on "goto" instruction (due to the find_label).
+
+However, the Machblock state after a goto remains "equivalent" to the trans_state of the Mach state in the sense of "equiv_on_next_step" below...
+
+*)
+
+
+Section ForwardSimuBlock_TRANS.
+
+Variable L1 L2: semantics.
+
+Variable trans_state: state L1 -> state L2.
+
+Definition equiv_on_next_step (P Q: Prop) s2_a s2_b: Prop :=
+ (P -> (forall t s', Step L2 s2_a t s' <-> Step L2 s2_b t s')) /\ (Q -> (forall r, (final_state L2 s2_a r) <-> (final_state L2 s2_b r))).
+
+Definition match_states s1 s2: Prop :=
+ equiv_on_next_step (exists t s1', Step L1 s1 t s1') (exists r, final_state L1 s1 r) s2 (trans_state s1).
+
+Lemma match_states_trans_state s1: match_states s1 (trans_state s1).
+Proof.
+ unfold match_states, equiv_on_next_step. intuition.
+Qed.
+
+Variable dist_end_block: state L1 -> nat.
+
+Hypothesis simu_mid_block:
+ forall s1 t s1', Step L1 s1 t s1' -> (dist_end_block s1)<>0 -> t = E0 /\ dist_end_block s1=S (dist_end_block s1').
+
+Hypothesis public_preserved:
+ forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id.
+
+Hypothesis match_initial_states:
+ forall s1, initial_state L1 s1 -> exists s2, match_states s1 s2 /\ initial_state L2 s2.
+
+Hypothesis match_final_states:
+ forall s1 r, final_state L1 s1 r -> final_state L2 (trans_state s1) r.
+
+Hypothesis final_states_end_block:
+ forall s1 t s1' r, Step L1 s1 t s1' -> final_state L1 s1' r -> dist_end_block s1 = 0.
+
+Hypothesis simu_end_block:
+ forall s1 t s1', starN (step L1) (globalenv L1) (S (dist_end_block s1)) s1 t s1' -> exists s2', Step L2 (trans_state s1) t s2' /\ match_states s1' s2'.
+
+Lemma forward_simulation_block_trans: forward_simulation L1 L2.
+Proof.
+ eapply forward_simulation_block_rel with (dist_end_block:=dist_end_block) (match_states:=match_states); try tauto.
+ + (* final_states *) intros s1 s2 r H1 [H2 H3]. rewrite H3; eauto.
+ + (* simu_end_block *)
+ intros s1 t s1' s2 H1 [H2a H2b]. exploit simu_end_block; eauto.
+ intros (s2' & H3 & H4); econstructor 1; intuition eauto.
+ rewrite H2a; auto.
+ inversion_clear H1. eauto.
+Qed.
+
+End ForwardSimuBlock_TRANS.
+
+
+(* another version with a relation [trans_state_R] instead of a function [trans_state] *)
+Section ForwardSimuBlock_TRANS_R.
+
+Variable L1 L2: semantics.
+
+Variable trans_state_R: state L1 -> state L2 -> Prop.
+
+Definition match_states_R s1 s2: Prop :=
+ exists s2', trans_state_R s1 s2' /\ equiv_on_next_step _ (exists t s1', Step L1 s1 t s1') (exists r, final_state L1 s1 r) s2 s2'.
+
+Lemma match_states_trans_state_R s1 s2: trans_state_R s1 s2 -> match_states_R s1 s2.
+Proof.
+ unfold match_states, equiv_on_next_step. firstorder.
+Qed.
+
+Variable dist_end_block: state L1 -> nat.
+
+Hypothesis simu_mid_block:
+ forall s1 t s1', Step L1 s1 t s1' -> (dist_end_block s1)<>0 -> t = E0 /\ dist_end_block s1=S (dist_end_block s1').
+
+Hypothesis public_preserved:
+ forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id.
+
+Hypothesis match_initial_states:
+ forall s1, initial_state L1 s1 -> exists s2, match_states_R s1 s2 /\ initial_state L2 s2.
+
+Hypothesis match_final_states:
+ forall s1 s2 r, final_state L1 s1 r -> trans_state_R s1 s2 -> final_state L2 s2 r.
+
+Hypothesis final_states_end_block:
+ forall s1 t s1' r, Step L1 s1 t s1' -> final_state L1 s1' r -> dist_end_block s1 = 0.
+
+Hypothesis simu_end_block:
+ forall s1 t s1' s2, starN (step L1) (globalenv L1) (S (dist_end_block s1)) s1 t s1' -> trans_state_R s1 s2 -> exists s2', Step L2 s2 t s2' /\ match_states_R s1' s2'.
+
+Lemma forward_simulation_block_trans_R: forward_simulation L1 L2.
+Proof.
+ eapply forward_simulation_block_rel with (dist_end_block:=dist_end_block) (match_states:=match_states_R); try tauto.
+ + (* final_states *) intros s1 s2 r H1 (s2' & H2 & H3 & H4). rewrite H4; eauto.
+ + (* simu_end_block *)
+ intros s1 t s1' s2 H1 (s2' & H2 & H2a & H2b). exploit simu_end_block; eauto.
+ intros (x & Hx & (y & H3 & H4 & H5)). repeat (econstructor; eauto).
+ rewrite H2a; eauto.
+ inversion_clear H1. eauto.
+Qed.
+
+End ForwardSimuBlock_TRANS_R.
+
diff --git a/scheduling/postpass_lib/Machblock.v b/scheduling/postpass_lib/Machblock.v
new file mode 100644
index 00000000..b588cca8
--- /dev/null
+++ b/scheduling/postpass_lib/Machblock.v
@@ -0,0 +1,388 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(** Abstract syntax and semantics of a Mach variant, structured with basic-blocks. *)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Values.
+Require Import Memory.
+Require Import Globalenvs.
+Require Import Events.
+Require Import Smallstep.
+Require Import Op.
+Require Import Locations.
+Require Import Conventions.
+Require Stacklayout.
+Require Import Mach.
+Require Import Linking.
+Require Import Lia.
+
+(** * Abstract Syntax *)
+
+(** ** basic instructions (ie no control-flow) *)
+Inductive basic_inst: Type :=
+ | MBgetstack: ptrofs -> typ -> mreg -> basic_inst
+ | MBsetstack: mreg -> ptrofs -> typ -> basic_inst
+ | MBgetparam: ptrofs -> typ -> mreg -> basic_inst
+ | MBop: operation -> list mreg -> mreg -> basic_inst
+ | MBload: trapping_mode -> memory_chunk -> addressing -> list mreg -> mreg -> basic_inst
+ | MBstore: memory_chunk -> addressing -> list mreg -> mreg -> basic_inst
+ .
+
+Definition bblock_body := list basic_inst.
+
+(** ** control flow instructions *)
+Inductive control_flow_inst: Type :=
+ | MBcall: signature -> mreg + ident -> control_flow_inst
+ | MBtailcall: signature -> mreg + ident -> control_flow_inst
+ | MBbuiltin: external_function -> list (builtin_arg mreg) -> builtin_res mreg -> control_flow_inst
+ | MBgoto: label -> control_flow_inst
+ | MBcond: condition -> list mreg -> label -> control_flow_inst
+ | MBjumptable: mreg -> list label -> control_flow_inst
+ | MBreturn: control_flow_inst
+ .
+
+(** ** basic block *)
+Record bblock := mk_bblock {
+ header: list label;
+ body: bblock_body;
+ exit: option control_flow_inst
+}.
+
+Lemma bblock_eq:
+ forall b1 b2,
+ header b1 = header b2 ->
+ body b1 = body b2 ->
+ exit b1 = exit b2 ->
+ b1 = b2.
+Proof.
+ intros. destruct b1. destruct b2.
+ cbn in *. subst. auto.
+Qed.
+
+Definition length_opt {A} (o: option A) : nat :=
+ match o with
+ | Some o => 1
+ | None => 0
+ end.
+
+Definition size (b:bblock): nat := (length (header b))+(length (body b))+(length_opt (exit b)).
+
+Lemma size_null b:
+ size b = 0%nat ->
+ header b = nil /\ body b = nil /\ exit b = None.
+Proof.
+ destruct b as [h b e]. cbn. unfold size. cbn.
+ intros H.
+ assert (length h = 0%nat) as Hh; [ lia |].
+ assert (length b = 0%nat) as Hb; [ lia |].
+ assert (length_opt e = 0%nat) as He; [ lia|].
+ repeat split.
+ destruct h; try (cbn in Hh; discriminate); auto.
+ destruct b; try (cbn in Hb; discriminate); auto.
+ destruct e; try (cbn in He; discriminate); auto.
+Qed.
+
+(** ** programs *)
+
+Definition code := list bblock.
+
+Record function: Type := mkfunction
+ { fn_sig: signature;
+ fn_code: code;
+ fn_stacksize: Z;
+ fn_link_ofs: ptrofs;
+ fn_retaddr_ofs: ptrofs }.
+
+Definition fundef := AST.fundef function.
+
+Definition program := AST.program fundef unit.
+
+Definition genv := Genv.t fundef unit.
+
+(** * Operational (blockstep) semantics ***)
+
+Lemma in_dec (lbl: label) (l: list label): { List.In lbl l } + { ~(List.In lbl l) }.
+Proof.
+ apply List.in_dec.
+ apply Pos.eq_dec.
+Qed.
+
+Definition is_label (lbl: label) (bb: bblock) : bool :=
+ if in_dec lbl (header bb) then true else false.
+
+Lemma is_label_correct_true lbl bb:
+ List.In lbl (header bb) <-> is_label lbl bb = true.
+Proof.
+ unfold is_label; destruct (in_dec lbl (header bb)); cbn; intuition.
+Qed.
+
+Lemma is_label_correct_false lbl bb:
+ ~(List.In lbl (header bb)) <-> is_label lbl bb = false.
+Proof.
+ unfold is_label; destruct (in_dec lbl (header bb)); cbn; intuition.
+Qed.
+
+
+Local Open Scope nat_scope.
+
+Fixpoint find_label (lbl: label) (c: code) {struct c} : option code :=
+ match c with
+ | nil => None
+ | bb1 :: bbl => if is_label lbl bb1 then Some c else find_label lbl bbl
+ end.
+
+Section RELSEM.
+
+Variable rao:function -> code -> ptrofs -> Prop.
+Variable ge:genv.
+
+Definition find_function_ptr
+ (ge: genv) (ros: mreg + ident) (rs: regset) : option block :=
+ match ros with
+ | inl r =>
+ match rs r with
+ | Vptr b ofs => if Ptrofs.eq ofs Ptrofs.zero then Some b else None
+ | _ => None
+ end
+ | inr symb =>
+ Genv.find_symbol ge symb
+ end.
+
+(** ** Machblock execution states. *)
+
+Inductive stackframe: Type :=
+ | Stackframe:
+ forall (f: block) (**r pointer to calling function *)
+ (sp: val) (**r stack pointer in calling function *)
+ (retaddr: val) (**r Asm return address in calling function *)
+ (c: code), (**r program point in calling function *)
+ stackframe.
+
+Inductive state: Type :=
+ | State:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: block) (**r pointer to current function *)
+ (sp: val) (**r stack pointer *)
+ (c: code) (**r current program point *)
+ (rs: regset) (**r register state *)
+ (m: mem), (**r memory state *)
+ state
+ | Callstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: block) (**r pointer to function to call *)
+ (rs: regset) (**r register state *)
+ (m: mem), (**r memory state *)
+ state
+ | Returnstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (rs: regset) (**r register state *)
+ (m: mem), (**r memory state *)
+ state.
+
+Definition parent_sp (s: list stackframe) : val :=
+ match s with
+ | nil => Vnullptr
+ | Stackframe f sp ra c :: s' => sp
+ end.
+
+Definition parent_ra (s: list stackframe) : val :=
+ match s with
+ | nil => Vnullptr
+ | Stackframe f sp ra c :: s' => ra
+ end.
+
+Inductive basic_step (s: list stackframe) (fb: block) (sp: val) (rs: regset) (m:mem): basic_inst -> regset -> mem -> Prop :=
+ | exec_MBgetstack:
+ forall ofs ty dst v,
+ load_stack m sp ty ofs = Some v ->
+ basic_step s fb sp rs m (MBgetstack ofs ty dst) (rs#dst <- v) m
+ | exec_MBsetstack:
+ forall src ofs ty m' rs',
+ store_stack m sp ty ofs (rs src) = Some m' ->
+ rs' = undef_regs (destroyed_by_setstack ty) rs ->
+ basic_step s fb sp rs m (MBsetstack src ofs ty) rs' m'
+ | exec_MBgetparam:
+ forall ofs ty dst v rs' f,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ load_stack m sp Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m (parent_sp s) ty ofs = Some v ->
+ rs' = (rs # temp_for_parent_frame <- Vundef # dst <- v) ->
+ basic_step s fb sp rs m (MBgetparam ofs ty dst) rs' m
+ | exec_MBop:
+ forall op args v rs' res,
+ eval_operation ge sp op rs##args m = Some v ->
+ rs' = ((undef_regs (destroyed_by_op op) rs)#res <- v) ->
+ basic_step s fb sp rs m (MBop op args res) rs' m
+ | exec_MBload:
+ forall addr args a v rs' trap chunk dst,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- v) ->
+ basic_step s fb sp rs m (MBload trap chunk addr args dst) rs' m
+ | exec_MBload_notrap1:
+ forall addr args rs' chunk dst,
+ eval_addressing ge sp addr rs##args = None ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- Vundef) ->
+ basic_step s fb sp rs m (MBload NOTRAP chunk addr args dst) rs' m
+ | exec_MBload_notrap2:
+ forall addr args a rs' chunk dst,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- Vundef) ->
+ basic_step s fb sp rs m (MBload NOTRAP chunk addr args dst) rs' m
+ | exec_MBstore:
+ forall chunk addr args src m' a rs',
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.storev chunk m a (rs src) = Some m' ->
+ rs' = undef_regs (destroyed_by_store chunk addr) rs ->
+ basic_step s fb sp rs m (MBstore chunk addr args src) rs' m'
+ .
+
+
+Inductive body_step (s: list stackframe) (f: block) (sp: val): bblock_body -> regset -> mem -> regset -> mem -> Prop :=
+ | exec_nil_body:
+ forall rs m,
+ body_step s f sp nil rs m rs m
+ | exec_cons_body:
+ forall rs m bi p rs' m' rs'' m'',
+ basic_step s f sp rs m bi rs' m' ->
+ body_step s f sp p rs' m' rs'' m'' ->
+ body_step s f sp (bi::p) rs m rs'' m''
+ .
+
+Inductive cfi_step: control_flow_inst -> state -> trace -> state -> Prop :=
+ | exec_MBcall:
+ forall s fb sp sig ros c b rs m f f' ra,
+ find_function_ptr ge ros rs = Some f' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ rao f c ra ->
+ cfi_step (MBcall sig ros) (State s fb sp (b::c) rs m)
+ E0 (Callstate (Stackframe fb sp (Vptr fb ra) c :: s)
+ f' rs m)
+ | exec_MBtailcall:
+ forall s fb stk soff sig ros c rs m f f' m',
+ find_function_ptr ge ros rs = Some f' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ cfi_step (MBtailcall sig ros) (State s fb (Vptr stk soff) c rs m)
+ E0 (Callstate s f' rs m')
+ | exec_MBbuiltin:
+ forall s f sp rs m ef args res b c vargs t vres rs' m',
+ eval_builtin_args ge rs sp m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = set_res res vres (undef_regs (destroyed_by_builtin ef) rs) ->
+ cfi_step (MBbuiltin ef args res) (State s f sp (b :: c) rs m)
+ t (State s f sp c rs' m')
+ | exec_MBgoto:
+ forall s fb f sp lbl c rs m c',
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ find_label lbl f.(fn_code) = Some c' ->
+ cfi_step (MBgoto lbl) (State s fb sp c rs m)
+ E0 (State s fb sp c' rs m)
+ | exec_MBcond_true:
+ forall s fb f sp cond args lbl c rs m c' rs',
+ eval_condition cond rs##args m = Some true ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ find_label lbl f.(fn_code) = Some c' ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
+ cfi_step (MBcond cond args lbl) (State s fb sp c rs m)
+ E0 (State s fb sp c' rs' m)
+ | exec_MBcond_false:
+ forall s f sp cond args lbl b c rs m rs',
+ eval_condition cond rs##args m = Some false ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
+ cfi_step (MBcond cond args lbl) (State s f sp (b :: c) rs m)
+ E0 (State s f sp c rs' m)
+ | exec_MBjumptable:
+ forall s fb f sp arg tbl c rs m n lbl c' rs',
+ rs arg = Vint n ->
+ list_nth_z tbl (Int.unsigned n) = Some lbl ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ find_label lbl f.(fn_code) = Some c' ->
+ rs' = undef_regs destroyed_by_jumptable rs ->
+ cfi_step (MBjumptable arg tbl) (State s fb sp c rs m)
+ E0 (State s fb sp c' rs' m)
+ | exec_MBreturn:
+ forall s fb stk soff c rs m f m',
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ cfi_step MBreturn (State s fb (Vptr stk soff) c rs m)
+ E0 (Returnstate s rs m')
+ .
+
+Inductive exit_step: option control_flow_inst -> state -> trace -> state -> Prop :=
+ | exec_Some_exit:
+ forall ctl s t s',
+ cfi_step ctl s t s' ->
+ exit_step (Some ctl) s t s'
+ | exec_None_exit:
+ forall stk f sp b lb rs m,
+ exit_step None (State stk f sp (b::lb) rs m) E0 (State stk f sp lb rs m)
+ .
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_bblock:
+ forall sf f sp bb c rs m rs' m' t s',
+ body_step sf f sp (body bb) rs m rs' m' ->
+ exit_step (exit bb) (State sf f sp (bb::c) rs' m') t s' ->
+ step (State sf f sp (bb::c) rs m) t s'
+ | exec_function_internal:
+ forall s fb rs m f m1 m2 m3 stk rs',
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) ->
+ let sp := Vptr stk Ptrofs.zero in
+ store_stack m1 sp Tptr f.(fn_link_ofs) (parent_sp s) = Some m2 ->
+ store_stack m2 sp Tptr f.(fn_retaddr_ofs) (parent_ra s) = Some m3 ->
+ rs' = undef_regs destroyed_at_function_entry rs ->
+ step (Callstate s fb rs m)
+ E0 (State s fb sp f.(fn_code) rs' m3)
+ | exec_function_external:
+ forall s fb rs m t rs' ef args res m',
+ Genv.find_funct_ptr ge fb = Some (External ef) ->
+ extcall_arguments rs m (parent_sp s) (ef_sig ef) args ->
+ external_call ef ge args m t res m' ->
+ rs' = set_pair (loc_result (ef_sig ef)) res (undef_caller_save_regs rs) ->
+ step (Callstate s fb rs m)
+ t (Returnstate s rs' m')
+ | exec_return:
+ forall s f sp ra c rs m,
+ step (Returnstate (Stackframe f sp ra c :: s) rs m)
+ E0 (State s f sp c rs m)
+ .
+
+End RELSEM.
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall fb m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ Genv.find_symbol ge p.(prog_main) = Some fb ->
+ initial_state p (Callstate nil fb (Regmap.init Vundef) m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall rs m r retcode,
+ loc_result signature_main = One r ->
+ rs r = Vint retcode ->
+ final_state (Returnstate nil rs m) retcode.
+
+Definition semantics (rao: function -> code -> ptrofs -> Prop) (p: program) :=
+ Semantics (step rao) (initial_state p) final_state (Genv.globalenv p).
diff --git a/scheduling/postpass_lib/Machblockgen.v b/scheduling/postpass_lib/Machblockgen.v
new file mode 100644
index 00000000..5a2f2a61
--- /dev/null
+++ b/scheduling/postpass_lib/Machblockgen.v
@@ -0,0 +1,221 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Values.
+Require Import Memory.
+Require Import Globalenvs.
+Require Import Events.
+Require Import Smallstep.
+Require Import Op.
+Require Import Locations.
+Require Import Conventions.
+Require Stacklayout.
+Require Import Mach.
+Require Import Linking.
+Require Import Machblock.
+
+(** * Tail-recursive (greedy) translation from Mach code to Machblock code *)
+
+Inductive Machblock_inst: Type :=
+| MB_label (lbl: label)
+| MB_basic (bi: basic_inst)
+| MB_cfi (cfi: control_flow_inst).
+
+Definition trans_inst (i:Mach.instruction) : Machblock_inst :=
+ match i with
+ | Mcall sig ros => MB_cfi (MBcall sig ros)
+ | Mtailcall sig ros => MB_cfi (MBtailcall sig ros)
+ | Mbuiltin ef args res => MB_cfi (MBbuiltin ef args res)
+ | Mgoto lbl => MB_cfi (MBgoto lbl)
+ | Mcond cond args lbl => MB_cfi (MBcond cond args lbl)
+ | Mjumptable arg tbl => MB_cfi (MBjumptable arg tbl)
+ | Mreturn => MB_cfi (MBreturn)
+ | Mgetstack ofs ty dst => MB_basic (MBgetstack ofs ty dst)
+ | Msetstack src ofs ty => MB_basic (MBsetstack src ofs ty)
+ | Mgetparam ofs ty dst => MB_basic (MBgetparam ofs ty dst)
+ | Mop op args res => MB_basic (MBop op args res)
+ | Mload trap chunk addr args dst=> MB_basic (MBload trap chunk addr args dst)
+ | Mstore chunk addr args src => MB_basic (MBstore chunk addr args src)
+ | Mlabel l => MB_label l
+ end.
+
+Definition empty_bblock:={| header := nil; body := nil; exit := None |}.
+Extraction Inline empty_bblock.
+
+Definition add_label l bb:={| header := l::(header bb); body := (body bb); exit := (exit bb) |}.
+Extraction Inline add_label.
+
+Definition add_basic bi bb :={| header := nil; body := bi::(body bb); exit := (exit bb) |}.
+Extraction Inline add_basic.
+
+Definition cfi_bblock cfi:={| header := nil; body := nil; exit := Some cfi |}.
+Extraction Inline cfi_bblock.
+
+Definition add_to_new_bblock (i:Machblock_inst) : bblock :=
+ match i with
+ | MB_label l => add_label l empty_bblock
+ | MB_basic i => add_basic i empty_bblock
+ | MB_cfi i => cfi_bblock i
+ end.
+
+(** Adding an instruction to the beginning of a bblock list by
+
+- either adding the instruction to the head of the list,
+
+- or creating a new bblock with the instruction
+*)
+Definition add_to_code (i:Machblock_inst) (bl:code) : code :=
+ match bl with
+ | bh::bl0 => match i with
+ | MB_label l => add_label l bh::bl0
+ | MB_cfi i0 => cfi_bblock i0::bl
+ | MB_basic i0 => match header bh with
+ |_::_ => add_basic i0 empty_bblock::bl
+ | nil => add_basic i0 bh::bl0
+ end
+ end
+ | _ => add_to_new_bblock i::nil
+ end.
+
+Fixpoint trans_code_rev (c: Mach.code) (bl:code) : code :=
+ match c with
+ | nil => bl
+ | i::c0 =>
+ trans_code_rev c0 (add_to_code (trans_inst i) bl)
+ end.
+
+Function trans_code (c: Mach.code) : code :=
+ trans_code_rev (List.rev_append c nil) nil.
+
+Definition transf_function (f: Mach.function) : function :=
+ {| fn_sig:=Mach.fn_sig f;
+ fn_code:=trans_code (Mach.fn_code f);
+ fn_stacksize := Mach.fn_stacksize f;
+ fn_link_ofs := Mach.fn_link_ofs f;
+ fn_retaddr_ofs := Mach.fn_retaddr_ofs f
+ |}.
+
+Definition transf_fundef (f: Mach.fundef) : fundef :=
+ transf_fundef transf_function f.
+
+Definition transf_program (src: Mach.program) : program :=
+ transform_program transf_fundef src.
+
+
+(** * Abstracting trans_code with a simpler inductive relation *)
+
+Inductive is_end_block: Machblock_inst -> code -> Prop :=
+ | End_empty mbi: is_end_block mbi nil
+ | End_basic bi bh bl: header bh <> nil -> is_end_block (MB_basic bi) (bh::bl)
+ | End_cfi cfi bl: bl <> nil -> is_end_block (MB_cfi cfi) bl.
+
+Local Hint Resolve End_empty End_basic End_cfi: core.
+
+Inductive is_trans_code: Mach.code -> code -> Prop :=
+ | Tr_nil: is_trans_code nil nil
+ | Tr_end_block i c bl:
+ is_trans_code c bl ->
+ is_end_block (trans_inst i) bl ->
+ is_trans_code (i::c) (add_to_new_bblock (trans_inst i)::bl)
+ | Tr_add_label i l bh c bl:
+ is_trans_code c (bh::bl) ->
+ i = Mlabel l ->
+ is_trans_code (i::c) (add_label l bh::bl)
+ | Tr_add_basic i bi bh c bl:
+ is_trans_code c (bh::bl) ->
+ trans_inst i = MB_basic bi ->
+ header bh = nil ->
+ is_trans_code (i::c) (add_basic bi bh::bl).
+
+Local Hint Resolve Tr_nil Tr_end_block: core.
+
+Lemma add_to_code_is_trans_code i c bl:
+ is_trans_code c bl ->
+ is_trans_code (i::c) (add_to_code (trans_inst i) bl).
+Proof.
+ destruct bl as [|bh0 bl]; cbn.
+ - intro H. inversion H. subst. eauto.
+ - remember (trans_inst i) as ti.
+ destruct ti as [l|bi|cfi].
+ + intros; eapply Tr_add_label; eauto. destruct i; cbn in * |- *; congruence.
+ + intros. remember (header bh0) as hbh0. destruct hbh0 as [|b].
+ * eapply Tr_add_basic; eauto.
+ * replace (add_basic bi empty_bblock) with (add_to_new_bblock (MB_basic bi)); auto.
+ rewrite Heqti; eapply Tr_end_block; eauto.
+ rewrite <- Heqti. eapply End_basic. congruence.
+ + intros.
+ replace (cfi_bblock cfi) with (add_to_new_bblock (MB_cfi cfi)); auto.
+ rewrite Heqti. eapply Tr_end_block; eauto.
+ rewrite <- Heqti. eapply End_cfi. congruence.
+Qed.
+
+Local Hint Resolve add_to_code_is_trans_code: core.
+
+Lemma trans_code_is_trans_code_rev c1: forall c2 mbi,
+ is_trans_code c2 mbi ->
+ is_trans_code (rev_append c1 c2) (trans_code_rev c1 mbi).
+Proof.
+ induction c1 as [| i c1]; cbn; auto.
+Qed.
+
+Lemma trans_code_is_trans_code c: is_trans_code c (trans_code c).
+Proof.
+ unfold trans_code.
+ rewrite <- rev_alt.
+ rewrite <- (rev_involutive c) at 1.
+ rewrite rev_alt at 1.
+ apply trans_code_is_trans_code_rev; auto.
+Qed.
+
+Lemma add_to_code_is_trans_code_inv i c bl:
+ is_trans_code (i::c) bl -> exists bl0, is_trans_code c bl0 /\ bl = add_to_code (trans_inst i) bl0.
+Proof.
+ intro H; inversion H as [|H0 H1 bl0| | H0 bi bh H1 bl0]; clear H; subst; (repeat econstructor); eauto.
+ + (* case Tr_end_block *) inversion H3; subst; cbn; auto.
+ * destruct (header bh); congruence.
+ * destruct bl0; cbn; congruence.
+ + (* case Tr_add_basic *) rewrite H3. cbn. destruct (header bh); congruence.
+Qed.
+
+Lemma trans_code_is_trans_code_rev_inv c1: forall c2 mbi,
+ is_trans_code (rev_append c1 c2) mbi ->
+ exists mbi0, is_trans_code c2 mbi0 /\ mbi=trans_code_rev c1 mbi0.
+Proof.
+ induction c1 as [| i c1]; cbn; eauto.
+ intros; exploit IHc1; eauto.
+ intros (mbi0 & H1 & H2); subst.
+ exploit add_to_code_is_trans_code_inv; eauto.
+ intros. destruct H0 as [mbi1 [H2 H3]].
+ exists mbi1. split; congruence.
+Qed.
+
+Local Hint Resolve trans_code_is_trans_code: core.
+
+Theorem is_trans_code_inv c bl: is_trans_code c bl <-> bl=(trans_code c).
+Proof.
+ constructor; intros; subst; auto.
+ unfold trans_code.
+ exploit (trans_code_is_trans_code_rev_inv (rev_append c nil) nil bl); eauto.
+ * rewrite <- rev_alt.
+ rewrite <- rev_alt.
+ rewrite (rev_involutive c).
+ apply H.
+ * intros.
+ destruct H0 as [mbi [H0 H1]].
+ inversion H0. subst. reflexivity.
+Qed.
diff --git a/scheduling/postpass_lib/Machblockgenproof.v b/scheduling/postpass_lib/Machblockgenproof.v
new file mode 100644
index 00000000..1d6c6e18
--- /dev/null
+++ b/scheduling/postpass_lib/Machblockgenproof.v
@@ -0,0 +1,810 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Values.
+Require Import Memory.
+Require Import Globalenvs.
+Require Import Events.
+Require Import Smallstep.
+Require Import Op.
+Require Import Locations.
+Require Import Conventions.
+Require Stacklayout.
+Require Import Mach.
+Require Import Linking.
+Require Import Machblock.
+Require Import Machblockgen.
+Require Import ForwardSimulationBlock.
+Require Import Lia.
+
+Ltac subst_is_trans_code H :=
+ rewrite is_trans_code_inv in H;
+ rewrite <- H in * |- *;
+ rewrite <- is_trans_code_inv in H.
+
+Definition inv_trans_rao (rao: function -> code -> ptrofs -> Prop) (f: Mach.function) (c: Mach.code) :=
+ rao (transf_function f) (trans_code c).
+
+Definition match_prog (p: Mach.program) (tp: Machblock.program) :=
+ match_program (fun _ f tf => tf = transf_fundef f) eq p tp.
+
+Lemma transf_program_match: forall p tp, transf_program p = tp -> match_prog p tp.
+Proof.
+ intros. rewrite <- H. eapply match_transform_program; eauto.
+Qed.
+
+Definition trans_stackframe (msf: Mach.stackframe) : stackframe :=
+ match msf with
+ | Mach.Stackframe f sp retaddr c => Stackframe f sp retaddr (trans_code c)
+ end.
+
+Fixpoint trans_stack (mst: list Mach.stackframe) : list stackframe :=
+ match mst with
+ | nil => nil
+ | msf :: mst0 => (trans_stackframe msf) :: (trans_stack mst0)
+ end.
+
+Definition trans_state (ms: Mach.state): state :=
+ match ms with
+ | Mach.State s f sp c rs m => State (trans_stack s) f sp (trans_code c) rs m
+ | Mach.Callstate s f rs m => Callstate (trans_stack s) f rs m
+ | Mach.Returnstate s rs m => Returnstate (trans_stack s) rs m
+ end.
+
+Section PRESERVATION.
+
+Local Open Scope nat_scope.
+
+Variable prog: Mach.program.
+Variable tprog: Machblock.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+
+Variable rao: function -> code -> ptrofs -> Prop.
+
+Definition match_states: Mach.state -> state -> Prop
+ := ForwardSimulationBlock.match_states (Mach.semantics (inv_trans_rao rao) prog) (Machblock.semantics rao tprog) trans_state.
+
+Lemma match_states_trans_state s1: match_states s1 (trans_state s1).
+Proof.
+ apply match_states_trans_state.
+Qed.
+
+Local Hint Resolve match_states_trans_state: core.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+Lemma init_mem_preserved:
+ forall m,
+ Genv.init_mem prog = Some m ->
+ Genv.init_mem tprog = Some m.
+Proof (Genv.init_mem_transf TRANSF).
+
+Lemma prog_main_preserved:
+ prog_main tprog = prog_main prog.
+Proof (match_program_main TRANSF).
+
+Lemma functions_translated:
+ forall b f,
+ Genv.find_funct_ptr ge b = Some f ->
+ exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = tf.
+Proof.
+ intros.
+ exploit (Genv.find_funct_ptr_match TRANSF); eauto. intro.
+ destruct H0 as (cunit & tf & A & B & C).
+ eapply ex_intro. intuition; eauto. subst. eapply A.
+Qed.
+
+Lemma find_function_ptr_same:
+ forall s rs,
+ Mach.find_function_ptr ge s rs = find_function_ptr tge s rs.
+Proof.
+ intros. unfold Mach.find_function_ptr. unfold find_function_ptr.
+ destruct s; auto.
+ rewrite symbols_preserved; auto.
+Qed.
+
+Lemma find_funct_ptr_same:
+ forall f f0,
+ Genv.find_funct_ptr ge f = Some (Internal f0) ->
+ Genv.find_funct_ptr tge f = Some (Internal (transf_function f0)).
+Proof.
+ intros. exploit (Genv.find_funct_ptr_transf TRANSF); eauto.
+Qed.
+
+Lemma find_funct_ptr_same_external:
+ forall f f0,
+ Genv.find_funct_ptr ge f = Some (External f0) ->
+ Genv.find_funct_ptr tge f = Some (External f0).
+Proof.
+ intros. exploit (Genv.find_funct_ptr_transf TRANSF); eauto.
+Qed.
+
+Lemma parent_sp_preserved:
+ forall s,
+ Mach.parent_sp s = parent_sp (trans_stack s).
+Proof.
+ unfold parent_sp. unfold Mach.parent_sp. destruct s; cbn; auto.
+ unfold trans_stackframe. destruct s; cbn; auto.
+Qed.
+
+Lemma parent_ra_preserved:
+ forall s,
+ Mach.parent_ra s = parent_ra (trans_stack s).
+Proof.
+ unfold parent_ra. unfold Mach.parent_ra. destruct s; cbn; auto.
+ unfold trans_stackframe. destruct s; cbn; auto.
+Qed.
+
+Lemma external_call_preserved:
+ forall ef args m t res m',
+ external_call ef ge args m t res m' ->
+ external_call ef tge args m t res m'.
+Proof.
+ intros. eapply external_call_symbols_preserved; eauto.
+ apply senv_preserved.
+Qed.
+
+Lemma Mach_find_label_split l i c c':
+ Mach.find_label l (i :: c) = Some c' ->
+ (i=Mlabel l /\ c' = c) \/ (i <> Mlabel l /\ Mach.find_label l c = Some c').
+Proof.
+ intros H.
+ destruct i; try (constructor 2; split; auto; discriminate ).
+ destruct (peq l0 l) as [P|P].
+ - constructor. subst l0; split; auto.
+ revert H. unfold Mach.find_label. cbn. rewrite peq_true.
+ intros H; injection H; auto.
+ - constructor 2. split.
+ + intro F. injection F. intros. contradict P; auto.
+ + revert H. unfold Mach.find_label. cbn. rewrite peq_false; auto.
+Qed.
+
+Lemma find_label_is_end_block_not_label i l c bl:
+ is_end_block (trans_inst i) bl ->
+ is_trans_code c bl ->
+ i <> Mlabel l -> find_label l (add_to_new_bblock (trans_inst i) :: bl) = find_label l bl.
+Proof.
+ intros H H0 H1.
+ unfold find_label.
+ remember (is_label l _) as b.
+ replace b with false; auto.
+ subst; unfold is_label.
+ destruct i; cbn in * |- *; try (destruct (in_dec l nil); intuition).
+ inversion H.
+ destruct (in_dec l (l0::nil)) as [H6|H6]; auto.
+ cbn in H6; intuition try congruence.
+Qed.
+
+Lemma find_label_at_begin l bh bl:
+ In l (header bh)
+ -> find_label l (bh :: bl) = Some (bh::bl).
+Proof.
+ unfold find_label; rewrite is_label_correct_true; intro H; rewrite H; cbn; auto.
+Qed.
+
+Lemma find_label_add_label_diff l bh bl:
+ ~(In l (header bh)) ->
+ find_label l (bh::bl) = find_label l bl.
+Proof.
+ unfold find_label; rewrite is_label_correct_false; intro H; rewrite H; cbn; auto.
+Qed.
+
+Definition concat (h: list label) (c: code): code :=
+ match c with
+ | nil => {| header := h; body := nil; exit := None |}::nil
+ | b::c' => {| header := h ++ (header b); body := body b; exit := exit b |}::c'
+ end.
+
+Lemma find_label_transcode_preserved:
+ forall l c c',
+ Mach.find_label l c = Some c' ->
+ exists h, In l h /\ find_label l (trans_code c) = Some (concat h (trans_code c')).
+Proof.
+ intros l c. remember (trans_code _) as bl.
+ rewrite <- is_trans_code_inv in * |-.
+ induction Heqbl.
+ + (* Tr_nil *)
+ intros; exists (l::nil); cbn in * |- *; intuition.
+ discriminate.
+ + (* Tr_end_block *)
+ intros.
+ exploit Mach_find_label_split; eauto.
+ clear H0; destruct 1 as [(H0&H2)|(H0&H2)].
+ - subst. rewrite find_label_at_begin; cbn; auto.
+ inversion H as [mbi H1 H2| | ].
+ subst.
+ inversion Heqbl.
+ subst.
+ exists (l :: nil); cbn; eauto.
+ - exploit IHHeqbl; eauto.
+ destruct 1 as (h & H3 & H4).
+ exists h.
+ split; auto.
+ erewrite find_label_is_end_block_not_label;eauto.
+ + (* Tr_add_label *)
+ intros.
+ exploit Mach_find_label_split; eauto.
+ clear H0; destruct 1 as [(H0&H2)|(H0&H2)].
+ - subst.
+ inversion H0 as [H1].
+ clear H0.
+ erewrite find_label_at_begin; cbn; eauto.
+ subst_is_trans_code Heqbl.
+ exists (l :: nil); cbn; eauto.
+ - subst; assert (H: l0 <> l); try congruence; clear H0.
+ exploit IHHeqbl; eauto.
+ clear IHHeqbl Heqbl.
+ intros (h & H3 & H4).
+ cbn; unfold is_label, add_label; cbn.
+ destruct (in_dec l (l0::header bh)) as [H5|H5]; cbn in H5.
+ * destruct H5; try congruence.
+ exists (l0::h); cbn; intuition.
+ rewrite find_label_at_begin in H4; auto.
+ apply f_equal. inversion H4 as [H5]. clear H4.
+ destruct (trans_code c'); cbn in * |- *;
+ inversion H5; subst; cbn; auto.
+ * exists h. intuition.
+ erewrite <- find_label_add_label_diff; eauto.
+ + (* Tr_add_basic *)
+ intros.
+ exploit Mach_find_label_split; eauto.
+ destruct 1 as [(H2&H3)|(H2&H3)].
+ rewrite H2 in H. unfold trans_inst in H. congruence.
+ exploit IHHeqbl; eauto.
+ clear IHHeqbl Heqbl.
+ intros (h & H4 & H5).
+ rewrite find_label_add_label_diff; auto.
+ rewrite find_label_add_label_diff in H5; eauto.
+ rewrite H0; auto.
+Qed.
+
+Lemma find_label_preserved:
+ forall l f c,
+ Mach.find_label l (Mach.fn_code f) = Some c ->
+ exists h, In l h /\ find_label l (fn_code (transf_function f)) = Some (concat h (trans_code c)).
+Proof.
+ intros. replace (fn_code (transf_function f)) with (trans_code (Mach.fn_code f)); eauto.
+ apply find_label_transcode_preserved; auto.
+Qed.
+
+Lemma mem_free_preserved:
+ forall m stk f,
+ Mem.free m stk 0 (Mach.fn_stacksize f) = Mem.free m stk 0 (fn_stacksize (transf_function f)).
+Proof.
+ intros. auto.
+Qed.
+
+Local Hint Resolve symbols_preserved senv_preserved init_mem_preserved prog_main_preserved functions_translated
+ parent_sp_preserved: core.
+
+
+Definition dist_end_block_code (c: Mach.code) :=
+ match trans_code c with
+ | nil => 0
+ | bh::_ => (size bh-1)%nat
+ end.
+
+Definition dist_end_block (s: Mach.state): nat :=
+ match s with
+ | Mach.State _ _ _ c _ _ => dist_end_block_code c
+ | _ => 0
+ end.
+
+Local Hint Resolve exec_nil_body exec_cons_body: core.
+Local Hint Resolve exec_MBgetstack exec_MBsetstack exec_MBgetparam exec_MBop exec_MBload exec_MBstore: core.
+
+Lemma size_add_label l bh: size (add_label l bh) = size bh + 1.
+Proof.
+ unfold add_label, size; cbn; lia.
+Qed.
+
+Lemma size_add_basic bi bh: header bh = nil -> size (add_basic bi bh) = size bh + 1.
+Proof.
+ intro H. unfold add_basic, size; rewrite H; cbn. lia.
+Qed.
+
+
+Lemma size_add_to_newblock i: size (add_to_new_bblock i) = 1.
+Proof.
+ destruct i; auto.
+Qed.
+
+
+Lemma dist_end_block_code_simu_mid_block i c:
+ dist_end_block_code (i::c) <> 0 ->
+ (dist_end_block_code (i::c) = Datatypes.S (dist_end_block_code c)).
+Proof.
+ unfold dist_end_block_code.
+ remember (trans_code (i::c)) as bl.
+ rewrite <- is_trans_code_inv in Heqbl.
+ inversion Heqbl as [|bl0 H| |]; subst; clear Heqbl.
+ - rewrite size_add_to_newblock; lia.
+ - rewrite size_add_label;
+ subst_is_trans_code H.
+ lia.
+ - rewrite size_add_basic; auto.
+ subst_is_trans_code H.
+ lia.
+Qed.
+
+Local Hint Resolve dist_end_block_code_simu_mid_block: core.
+
+
+Lemma size_nonzero c b bl:
+ is_trans_code c (b :: bl) -> size b <> 0.
+Proof.
+ intros H; inversion H; subst.
+ - rewrite size_add_to_newblock; lia.
+ - rewrite size_add_label; lia.
+ - rewrite size_add_basic; auto; lia.
+Qed.
+
+Inductive is_header: list label -> Mach.code -> Mach.code -> Prop :=
+ | header_empty : is_header nil nil nil
+ | header_not_label i c: (forall l, i <> Mlabel l) -> is_header nil (i::c) (i::c)
+ | header_is_label l h c c0: is_header h c c0 -> is_header (l::h) ((Mlabel l)::c) c0
+ .
+
+Inductive is_body: list basic_inst -> Mach.code -> Mach.code -> Prop :=
+ | body_empty : is_body nil nil nil
+ | body_not_bi i c: (forall bi, (trans_inst i) <> (MB_basic bi)) -> is_body nil (i::c) (i::c)
+ | body_is_bi i lbi c0 c1 bi: (trans_inst i) = MB_basic bi -> is_body lbi c0 c1 -> is_body (bi::lbi) (i::c0) c1
+ .
+
+Inductive is_exit: option control_flow_inst -> Mach.code -> Mach.code -> Prop :=
+ | exit_empty: is_exit None nil nil
+ | exit_not_cfi i c: (forall cfi, (trans_inst i) <> MB_cfi cfi) -> is_exit None (i::c) (i::c)
+ | exit_is_cfi i c cfi: (trans_inst i) = MB_cfi cfi -> is_exit (Some cfi) (i::c) c
+ .
+
+Lemma Mlabel_is_not_basic i:
+ forall bi, trans_inst i = MB_basic bi -> forall l, i <> Mlabel l.
+Proof.
+intros.
+unfold trans_inst in H.
+destruct i; congruence.
+Qed.
+
+Lemma Mlabel_is_not_cfi i:
+ forall cfi, trans_inst i = MB_cfi cfi -> forall l, i <> Mlabel l.
+Proof.
+intros.
+unfold trans_inst in H.
+destruct i; congruence.
+Qed.
+
+Lemma MBbasic_is_not_cfi i:
+ forall cfi, trans_inst i = MB_cfi cfi -> forall bi, trans_inst i <> MB_basic bi.
+Proof.
+intros.
+unfold trans_inst in H.
+unfold trans_inst.
+destruct i; congruence.
+Qed.
+
+
+Local Hint Resolve Mlabel_is_not_cfi: core.
+Local Hint Resolve MBbasic_is_not_cfi: core.
+
+Lemma add_to_new_block_is_label i:
+ header (add_to_new_bblock (trans_inst i)) <> nil -> exists l, i = Mlabel l.
+Proof.
+ intros.
+ unfold add_to_new_bblock in H.
+ destruct (trans_inst i) eqn : H1.
+ + exists lbl.
+ unfold trans_inst in H1.
+ destruct i; congruence.
+ + unfold add_basic in H; cbn in H; congruence.
+ + unfold cfi_bblock in H; cbn in H; congruence.
+Qed.
+
+Local Hint Resolve Mlabel_is_not_basic: core.
+
+Lemma trans_code_decompose c: forall b bl,
+ is_trans_code c (b::bl) ->
+ exists c0 c1 c2, is_header (header b) c c0 /\ is_body (body b) c0 c1 /\ is_exit (exit b) c1 c2 /\ is_trans_code c2 bl.
+Proof.
+ induction c as [|i c].
+ { (* nil => absurd *) intros b bl H; inversion H. }
+ intros b bl H; remember (trans_inst i) as ti.
+ destruct ti as [lbl|bi|cfi];
+ inversion H as [|d0 d1 d2 H0 H1| |]; subst;
+ try (rewrite <- Heqti in * |- *); cbn in * |- *;
+ try congruence.
+ + (* label at end block *)
+ inversion H1; subst. inversion H0; subst.
+ assert (X:i=Mlabel lbl). { destruct i; cbn in Heqti; congruence. }
+ subst. repeat econstructor; eauto.
+ + (* label at mid block *)
+ exploit IHc; eauto.
+ intros (c0 & c1 & c2 & H1 & H2 & H3 & H4).
+ repeat econstructor; eauto.
+ + (* basic at end block *)
+ inversion H1; subst.
+ lapply (Mlabel_is_not_basic i bi); auto.
+ intro H2.
+ - inversion H0; subst.
+ assert (X:(trans_inst i) = MB_basic bi ). { repeat econstructor; congruence. }
+ repeat econstructor; congruence.
+ - exists (i::c), c, c.
+ repeat econstructor; eauto; inversion H0; subst; repeat econstructor; cbn; try congruence.
+ * exploit (add_to_new_block_is_label i0); eauto.
+ intros (l & H8); subst; cbn; congruence.
+ * exploit H3; eauto.
+ * exploit (add_to_new_block_is_label i0); eauto.
+ intros (l & H8); subst; cbn; congruence.
+ + (* basic at mid block *)
+ inversion H1; subst.
+ exploit IHc; eauto.
+ intros (c0 & c1 & c2 & H3 & H4 & H5 & H6).
+ exists (i::c0), c1, c2.
+ repeat econstructor; eauto.
+ rewrite H2 in H3.
+ inversion H3; econstructor; eauto.
+ + (* cfi at end block *)
+ inversion H1; subst;
+ repeat econstructor; eauto.
+Qed.
+
+
+Lemma step_simu_header st f sp rs m s c h c' t:
+ is_header h c c' ->
+ starN (Mach.step (inv_trans_rao rao)) (Genv.globalenv prog) (length h) (Mach.State st f sp c rs m) t s ->
+ s = Mach.State st f sp c' rs m /\ t = E0.
+Proof.
+ induction 1; cbn; intros hs; try (inversion hs; tauto).
+ inversion hs as [|n1 s1 t1 t2 s2 t3 s3 H1]. inversion H1. subst. auto.
+Qed.
+
+
+
+Lemma step_simu_basic_step (i: Mach.instruction) (bi: basic_inst) (c: Mach.code) s f sp rs m (t:trace) (s':Mach.state):
+ trans_inst i = MB_basic bi ->
+ Mach.step (inv_trans_rao rao) ge (Mach.State s f sp (i::c) rs m) t s' ->
+ exists rs' m', s'=Mach.State s f sp c rs' m' /\ t=E0 /\ basic_step tge (trans_stack s) f sp rs m bi rs' m'.
+Proof.
+ destruct i; cbn in * |-;
+ (discriminate
+ || (intro H; inversion_clear H; intro X; inversion_clear X; eapply ex_intro; eapply ex_intro; intuition eauto)).
+ - eapply exec_MBgetparam; eauto. exploit (functions_translated); eauto. intro.
+ destruct H3 as (tf & A & B). subst. eapply A.
+ all: cbn; rewrite <- parent_sp_preserved; auto.
+ - eapply exec_MBop; eauto. rewrite <- H. destruct o; cbn; auto. destruct (rs ## l); cbn; auto.
+ unfold Genv.symbol_address; rewrite symbols_preserved; auto.
+ - eapply exec_MBload; eauto; rewrite <- H; destruct a; cbn; auto; destruct (rs ## l); cbn; auto;
+ unfold Genv.symbol_address; rewrite symbols_preserved; auto.
+ - eapply exec_MBload_notrap1; eauto; rewrite <- H; destruct a; cbn; auto; destruct (rs ## l); cbn; auto;
+ unfold Genv.symbol_address; rewrite symbols_preserved; auto.
+ - eapply exec_MBload_notrap2; eauto; rewrite <- H; destruct a; cbn; auto; destruct (rs ## l); cbn; auto;
+ unfold Genv.symbol_address; rewrite symbols_preserved; auto.
+ - eapply exec_MBstore; eauto; rewrite <- H; destruct a; cbn; auto; destruct (rs ## l); cbn; auto;
+ unfold Genv.symbol_address; rewrite symbols_preserved; auto.
+Qed.
+
+
+Lemma star_step_simu_body_step s f sp c bdy c':
+ is_body bdy c c' -> forall rs m t s',
+ starN (Mach.step (inv_trans_rao rao)) ge (length bdy) (Mach.State s f sp c rs m) t s' ->
+ exists rs' m', s'=Mach.State s f sp c' rs' m' /\ t=E0 /\ body_step tge (trans_stack s) f sp bdy rs m rs' m'.
+Proof.
+ induction 1; cbn.
+ + intros. inversion H. exists rs. exists m. auto.
+ + intros. inversion H0. exists rs. exists m. auto.
+ + intros. inversion H1; subst.
+ exploit (step_simu_basic_step ); eauto.
+ destruct 1 as [ rs1 [ m1 Hs]].
+ destruct Hs as [Hs1 [Hs2 Hs3]].
+ destruct (IHis_body rs1 m1 t2 s') as [rs2 Hb]. rewrite <- Hs1; eauto.
+ destruct Hb as [m2 [Hb1 [Hb2 Hb3]]].
+ exists rs2, m2.
+ rewrite Hs2, Hb2; eauto.
+ Qed.
+
+Local Hint Resolve exec_MBcall exec_MBtailcall exec_MBbuiltin exec_MBgoto exec_MBcond_true exec_MBcond_false exec_MBjumptable exec_MBreturn exec_Some_exit exec_None_exit: core.
+Local Hint Resolve eval_builtin_args_preserved external_call_symbols_preserved find_funct_ptr_same: core.
+
+
+Lemma match_states_concat_trans_code st f sp c rs m h:
+ match_states (Mach.State st f sp c rs m) (State (trans_stack st) f sp (concat h (trans_code c)) rs m).
+Proof.
+ intros; constructor 1; cbn.
+ + intros (t0 & s1' & H0) t s'.
+ remember (trans_code _) as bl.
+ destruct bl as [|bh bl].
+ { rewrite <- is_trans_code_inv in Heqbl; inversion Heqbl; inversion H0; congruence. }
+ clear H0.
+ cbn; constructor 1;
+ intros X; inversion X as [d1 d2 d3 d4 d5 d6 d7 rs' m' d10 d11 X1 X2| | | ]; subst; cbn in * |- *;
+ eapply exec_bblock; eauto; cbn;
+ inversion X2 as [cfi d1 d2 d3 H1|]; subst; eauto;
+ inversion H1; subst; eauto.
+ + intros H r; constructor 1; intro X; inversion X.
+Qed.
+
+Lemma step_simu_cfi_step (i: Mach.instruction) (cfi: control_flow_inst) (c: Mach.code) (blc:code) stk f sp rs m (t:trace) (s':Mach.state) b:
+ trans_inst i = MB_cfi cfi ->
+ is_trans_code c blc ->
+ Mach.step (inv_trans_rao rao) ge (Mach.State stk f sp (i::c) rs m) t s' ->
+ exists s2, cfi_step rao tge cfi (State (trans_stack stk) f sp (b::blc) rs m) t s2 /\ match_states s' s2.
+Proof.
+ destruct i; cbn in * |-;
+ (intro H; intro Htc;apply is_trans_code_inv in Htc;rewrite Htc;inversion_clear H;intro X; inversion_clear X).
+ * eapply ex_intro.
+ intuition auto.
+ eapply exec_MBcall;eauto.
+ rewrite <-H; exploit (find_function_ptr_same); eauto.
+ * eapply ex_intro.
+ intuition auto.
+ eapply exec_MBtailcall;eauto.
+ - rewrite <-H; exploit (find_function_ptr_same); eauto.
+ - cbn; rewrite <- parent_sp_preserved; auto.
+ - cbn; rewrite <- parent_ra_preserved; auto.
+ * eapply ex_intro.
+ intuition auto.
+ eapply exec_MBbuiltin ;eauto.
+ * exploit find_label_transcode_preserved; eauto.
+ intros (x & X1 & X2).
+ eapply ex_intro; constructor 1; [ idtac | eapply match_states_concat_trans_code ]; eauto.
+ * exploit find_label_transcode_preserved; eauto.
+ intros (x & X1 & X2).
+ eapply ex_intro; constructor 1; [ idtac | eapply match_states_concat_trans_code ]; eauto.
+ * eapply ex_intro; constructor 1; [ idtac | eapply match_states_trans_state ]; eauto.
+ eapply exec_MBcond_false; eauto.
+ * exploit find_label_transcode_preserved; eauto. intros (h & X1 & X2).
+ eapply ex_intro; constructor 1; [ idtac | eapply match_states_concat_trans_code ]; eauto.
+ * eapply ex_intro; constructor 1; [ idtac | eapply match_states_trans_state ]; eauto.
+ eapply exec_MBreturn; eauto.
+ rewrite parent_sp_preserved in H0; subst; auto.
+ rewrite parent_ra_preserved in H1; subst; auto.
+Qed.
+
+Lemma step_simu_exit_step stk f sp rs m t s1 e c c' b blc:
+ is_exit e c c' -> is_trans_code c' blc ->
+ starN (Mach.step (inv_trans_rao rao)) (Genv.globalenv prog) (length_opt e) (Mach.State stk f sp c rs m) t s1 ->
+ exists s2, exit_step rao tge e (State (trans_stack stk) f sp (b::blc) rs m) t s2 /\ match_states s1 s2.
+Proof.
+ destruct 1.
+ - (* None *)
+ intros H0 H1. inversion H1. exists (State (trans_stack stk) f sp blc rs m).
+ split; eauto.
+ apply is_trans_code_inv in H0.
+ rewrite H0.
+ apply match_states_trans_state.
+ - (* None *)
+ intros H0 H1. inversion H1. exists (State (trans_stack stk) f sp blc rs m).
+ split; eauto.
+ apply is_trans_code_inv in H0.
+ rewrite H0.
+ apply match_states_trans_state.
+ - (* Some *)
+ intros H0 H1.
+ inversion H1; subst.
+ exploit (step_simu_cfi_step); eauto.
+ intros [s2 [Hcfi1 Hcfi3]].
+ inversion H4. subst; cbn.
+ autorewrite with trace_rewrite.
+ exists s2.
+ split;eauto.
+Qed.
+
+Lemma simu_end_block:
+ forall s1 t s1',
+ starN (Mach.step (inv_trans_rao rao)) ge (Datatypes.S (dist_end_block s1)) s1 t s1' ->
+ exists s2', step rao tge (trans_state s1) t s2' /\ match_states s1' s2'.
+Proof.
+ destruct s1; cbn.
+ + (* State *)
+ remember (trans_code _) as tc.
+ rewrite <- is_trans_code_inv in Heqtc.
+ intros t s1 H.
+ destruct tc as [|b bl].
+ { (* nil => absurd *)
+ inversion Heqtc. subst.
+ unfold dist_end_block_code; cbn.
+ inversion_clear H;
+ inversion_clear H0.
+ }
+ assert (X: Datatypes.S (dist_end_block_code c) = (size b)).
+ {
+ unfold dist_end_block_code.
+ subst_is_trans_code Heqtc.
+ lapply (size_nonzero c b bl); auto.
+ lia.
+ }
+ rewrite X in H; unfold size in H.
+ (* decomposition of starN in 3 parts: header + body + exit *)
+ destruct (starN_split (Mach.semantics (inv_trans_rao rao) prog) _ _ _ _ H _ _ refl_equal) as (t3&t4&s1'&H0&H3&H4).
+ subst t; clear X H.
+ destruct (starN_split (Mach.semantics (inv_trans_rao rao) prog) _ _ _ _ H0 _ _ refl_equal) as (t1&t2&s1''&H&H1&H2).
+ subst t3; clear H0.
+ exploit trans_code_decompose; eauto. clear Heqtc.
+ intros (c0&c1&c2&Hc0&Hc1&Hc2&Heqtc).
+ (* header steps *)
+ exploit step_simu_header; eauto.
+ clear H; intros [X1 X2]; subst.
+ (* body steps *)
+ exploit (star_step_simu_body_step); eauto.
+ clear H1; intros (rs'&m'&H0&H1&H2). subst.
+ autorewrite with trace_rewrite.
+ (* exit step *)
+ exploit step_simu_exit_step; eauto.
+ clear H3; intros (s2' & H3 & H4).
+ eapply ex_intro; intuition eauto.
+ eapply exec_bblock; eauto.
+ + (* Callstate *)
+ intros t s1' H; inversion_clear H.
+ eapply ex_intro; constructor 1; eauto.
+ inversion H1; subst; clear H1.
+ inversion_clear H0; cbn.
+ - (* function_internal*)
+ replace (trans_code (Mach.fn_code f0)) with (fn_code (transf_function f0)); eauto.
+ eapply exec_function_internal; eauto.
+ rewrite <- parent_sp_preserved; eauto.
+ rewrite <- parent_ra_preserved; eauto.
+ - (* function_external *)
+ autorewrite with trace_rewrite.
+ eapply exec_function_external; eauto.
+ apply find_funct_ptr_same_external; auto.
+ rewrite <- parent_sp_preserved; eauto.
+ + (* Returnstate *)
+ intros t s1' H; inversion_clear H.
+ eapply ex_intro; constructor 1; eauto.
+ inversion H1; subst; clear H1.
+ inversion_clear H0; cbn.
+ eapply exec_return.
+Qed.
+
+
+Lemma cfi_dist_end_block i c:
+(exists cfi, trans_inst i = MB_cfi cfi) ->
+dist_end_block_code (i :: c) = 0.
+Proof.
+ unfold dist_end_block_code.
+ intro H. destruct H as [cfi H].
+ destruct i;cbn in H;try(congruence); (
+ remember (trans_code _) as bl;
+ rewrite <- is_trans_code_inv in Heqbl;
+ inversion Heqbl; subst; cbn in * |- *; try (congruence)).
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (Mach.semantics (inv_trans_rao rao) prog) (Machblock.semantics rao tprog).
+Proof.
+ apply forward_simulation_block_trans with (dist_end_block := dist_end_block) (trans_state := trans_state).
+(* simu_mid_block *)
+ - intros s1 t s1' H1 H2.
+ destruct H1; cbn in * |- *; lia || (intuition auto);
+ destruct H2; eapply cfi_dist_end_block; cbn; eauto.
+(* public_preserved *)
+ - apply senv_preserved.
+(* match_initial_states *)
+ - intros. cbn.
+ eapply ex_intro; constructor 1.
+ eapply match_states_trans_state.
+ destruct H. split.
+ apply init_mem_preserved; auto.
+ rewrite prog_main_preserved. rewrite <- H0. apply symbols_preserved.
+(* match_final_states *)
+ - intros. cbn. destruct H. split with (r := r); auto.
+(* final_states_end_block *)
+ - intros. cbn in H0.
+ inversion H0.
+ inversion H; cbn; auto.
+ all: try (subst; discriminate).
+ apply cfi_dist_end_block; exists MBreturn; eauto.
+(* simu_end_block *)
+ - apply simu_end_block.
+Qed.
+
+End PRESERVATION.
+
+(** Auxiliary lemmas used to prove existence of a Mach return adress from a Machblock return address. *)
+
+
+
+Lemma is_trans_code_monotonic i c b l:
+ is_trans_code c (b::l) ->
+ exists l' b', is_trans_code (i::c) (l' ++ (b'::l)).
+Proof.
+ intro H; destruct c as [|i' c]. { inversion H. }
+ remember (trans_inst i) as ti.
+ destruct ti as [lbl|bi|cfi].
+ - (*i=lbl *) replace (i ) with (Mlabel lbl). 2: ( destruct i; cbn in * |- *; try congruence ).
+ exists nil; cbn; eexists. eapply Tr_add_label; eauto.
+ - (*i=basic*)
+ destruct i'.
+ 10: { exists (add_to_new_bblock (MB_basic bi)::nil). exists b.
+ replace ((add_to_new_bblock (MB_basic bi) :: nil) ++ (b::l)) with ((add_to_new_bblock (MB_basic bi) :: (b::l)));eauto.
+ rewrite Heqti.
+ eapply Tr_end_block; eauto.
+ rewrite <-Heqti.
+ eapply End_basic. inversion H; try(cbn; congruence).
+ cbn in H5; congruence. }
+ all: try(exists nil; cbn; eexists; eapply Tr_add_basic; eauto; inversion H; try(eauto || congruence)).
+ - (*i=cfi*)
+ destruct i; try(cbn in Heqti; congruence).
+ all: exists (add_to_new_bblock (MB_cfi cfi)::nil); exists b;
+ replace ((add_to_new_bblock (MB_cfi cfi) :: nil) ++ (b::l)) with ((add_to_new_bblock (MB_cfi cfi) :: (b::l)));eauto;
+ rewrite Heqti;
+ eapply Tr_end_block; eauto;
+ rewrite <-Heqti;
+ eapply End_cfi; congruence.
+Qed.
+
+Lemma trans_code_monotonic i c b l:
+ (b::l) = trans_code c ->
+ exists l' b', trans_code (i::c) = (l' ++ (b'::l)).
+Proof.
+ intro H; rewrite <- is_trans_code_inv in H.
+ destruct (is_trans_code_monotonic i c b l H) as (l' & b' & H0).
+ subst_is_trans_code H0.
+ eauto.
+Qed.
+
+Lemma Mach_Machblock_tail sg ros c c1 c2: c1=(Mcall sg ros :: c) -> is_tail c1 c2 ->
+ exists b, is_tail (b :: trans_code c) (trans_code c2).
+Proof.
+ intros H; induction 1.
+ - intros; subst.
+ remember (trans_code (Mcall _ _::c)) as tc2.
+ rewrite <- is_trans_code_inv in Heqtc2.
+ inversion Heqtc2; cbn in * |- *; subst; try congruence.
+ subst_is_trans_code H1.
+ eapply ex_intro; eauto with coqlib.
+ - intros; exploit IHis_tail; eauto. clear IHis_tail.
+ intros (b & Hb). inversion Hb; clear Hb.
+ * exploit (trans_code_monotonic i c2); eauto.
+ intros (l' & b' & Hl'); rewrite Hl'.
+ exists b'; cbn; eauto with coqlib.
+ * exploit (trans_code_monotonic i c2); eauto.
+ intros (l' & b' & Hl'); rewrite Hl'.
+ cbn; eapply ex_intro.
+ eapply is_tail_trans; eauto with coqlib.
+Qed.
+
+Section Mach_Return_Address.
+
+Variable return_address_offset: function -> code -> ptrofs -> Prop.
+
+Hypothesis ra_exists: forall (b: bblock) (f: function) (c : list bblock),
+ is_tail (b :: c) (fn_code f) -> exists ra : ptrofs, return_address_offset f c ra.
+
+Definition Mach_return_address_offset (f: Mach.function) (c: Mach.code) (ofs: ptrofs) : Prop :=
+ return_address_offset (transf_function f) (trans_code c) ofs.
+
+Lemma Mach_return_address_exists:
+ forall f sg ros c, is_tail (Mcall sg ros :: c) f.(Mach.fn_code) ->
+ exists ra, Mach_return_address_offset f c ra.
+Proof.
+ intros.
+ exploit Mach_Machblock_tail; eauto.
+ destruct 1.
+ eapply ra_exists; eauto.
+Qed.
+
+End Mach_Return_Address.